From 05b86d20e85b63e9b19c973dbb6a0c01514aca85 Mon Sep 17 00:00:00 2001 From: jt0202 Date: Sun, 10 Nov 2024 13:35:51 +0100 Subject: [PATCH 01/83] Model function for insertManyList --- src/Std/Data/DHashMap/Internal/Defs.lean | 8 +++++ .../DHashMap/Internal/List/Associative.lean | 21 +++++++++-- src/Std/Data/DHashMap/Internal/Model.lean | 14 ++++++++ src/Std/Data/DHashMap/Internal/WF.lean | 35 +++++++++++++++++++ 4 files changed, 76 insertions(+), 2 deletions(-) diff --git a/src/Std/Data/DHashMap/Internal/Defs.lean b/src/Std/Data/DHashMap/Internal/Defs.lean index e4a03ce53ee8..aac8dd26d1d2 100644 --- a/src/Std/Data/DHashMap/Internal/Defs.lean +++ b/src/Std/Data/DHashMap/Internal/Defs.lean @@ -352,6 +352,14 @@ where r := ⟨r.1.insert a b, fun _ h hm => h (r.2 _ h hm)⟩ return r +/-- Internal implementation detail of the hash map -/ +def insertManyList [BEq α] [Hashable α] + (m : Raw₀ α β) (l : List ((a : α) × β a)) : Raw₀ α β := Id.run do +let mut r := m +for ⟨a,b⟩ in l do + r:= r.insert a b +return r + section variable {β : Type v} diff --git a/src/Std/Data/DHashMap/Internal/List/Associative.lean b/src/Std/Data/DHashMap/Internal/List/Associative.lean index 97bbaf230ace..5f8b9551da39 100644 --- a/src/Std/Data/DHashMap/Internal/List/Associative.lean +++ b/src/Std/Data/DHashMap/Internal/List/Associative.lean @@ -849,10 +849,10 @@ theorem isEmpty_eraseKey [BEq α] {l : List ((a : α) × β a)} {k : α} : theorem keys_eq_map (l : List ((a : α) × β a)) : keys l = l.map (·.1) := by induction l using assoc_induction <;> simp_all -theorem length_keys_eq_length (l : List ((a : α) × β a)) : (keys l).length = l.length := by +theorem length_keys_eq_length (l : List ((a : α) × β a)) : (keys l).length = l.length := by induction l using assoc_induction <;> simp_all -theorem isEmpty_keys_eq_isEmpty (l : List ((a : α) × β a)) : (keys l).isEmpty = l.isEmpty := by +theorem isEmpty_keys_eq_isEmpty (l : List ((a : α) × β a)) : (keys l).isEmpty = l.isEmpty := by induction l using assoc_induction <;> simp_all theorem containsKey_eq_keys_contains [BEq α] [PartialEquivBEq α] {l : List ((a : α) × β a)} @@ -1827,4 +1827,21 @@ theorem eraseKey_append_of_containsKey_right_eq_false [BEq α] {l l' : List ((a · rw [cond_false, cond_false, ih, List.cons_append] · rw [cond_true, cond_true] +/-- Internal implementation detail of the hash map -/ +def insertMany [BEq α] (l toInsert: List ((a : α) × β a)) : List ((a : α) × β a) := + match toInsert with + | .nil => l + | .cons ⟨k, v⟩ toInsert => insertMany (insertEntry k v l) toInsert + +theorem insertMany_perm_of_perm_first [BEq α] [EquivBEq α] (l1 l2 toInsert: List ((a : α) × β a)) (h: Perm l1 l2) (distinct: DistinctKeys l1): Perm (insertMany l1 toInsert) (insertMany l2 toInsert) := by + induction toInsert generalizing l1 l2 with + | nil => simp[insertMany, h] + | cons hd tl ih => + simp[insertMany] + apply ih + apply insertEntry_of_perm + exact distinct + exact h + apply DistinctKeys.insertEntry distinct + end List diff --git a/src/Std/Data/DHashMap/Internal/Model.lean b/src/Std/Data/DHashMap/Internal/Model.lean index 444082ed805d..fbb0d971da0c 100644 --- a/src/Std/Data/DHashMap/Internal/Model.lean +++ b/src/Std/Data/DHashMap/Internal/Model.lean @@ -322,6 +322,12 @@ def mapₘ (m : Raw₀ α β) (f : (a : α) → β a → δ a) : Raw₀ α δ := def filterₘ (m : Raw₀ α β) (f : (a : α) → β a → Bool) : Raw₀ α β := ⟨withComputedSize (updateAllBuckets m.1.buckets fun l => l.filter f), by simpa using m.2⟩ +/-- Internal implementation detail of the hash map -/ +def insertManyListₘ [BEq α] [Hashable α](m : Raw₀ α β) (l: List ((a : α) × β a)): Raw₀ α β := + match l with + | .nil => m + | .cons hd tl => insertManyListₘ (m.insert hd.1 hd.2) tl + section variable {β : Type v} @@ -455,6 +461,14 @@ theorem map_eq_mapₘ (m : Raw₀ α β) (f : (a : α) → β a → δ a) : theorem filter_eq_filterₘ (m : Raw₀ α β) (f : (a : α) → β a → Bool) : m.filter f = m.filterₘ f := rfl +theorem insertManyList_eq_insertManyListₘ [BEq α] [Hashable α](m : Raw₀ α β) (l: List ((a : α) × β a)): insertManyList m l = insertManyListₘ m l := by + simp [insertManyList, Id.run] + induction l generalizing m with + | nil => simp[insertManyListₘ] + | cons hd tl ih => + simp [insertManyListₘ] + rw [ih] + section variable {β : Type v} diff --git a/src/Std/Data/DHashMap/Internal/WF.lean b/src/Std/Data/DHashMap/Internal/WF.lean index eb9740b82e20..9e47c3f76b0b 100644 --- a/src/Std/Data/DHashMap/Internal/WF.lean +++ b/src/Std/Data/DHashMap/Internal/WF.lean @@ -706,6 +706,40 @@ theorem wfImp_filter [BEq α] [Hashable α] [EquivBEq α] [LawfulHashable α] {m rw [filter_eq_filterₘ] exact wfImp_filterₘ h +/-! # `insertManyListₘ` -/ +theorem wfImp_insertManyListₘ [BEq α] [Hashable α] [EquivBEq α][LawfulHashable α] {m : Raw₀ α β} {l: List ((a : α) × β a)} (h : Raw.WFImp m.1): Raw.WFImp (insertManyListₘ m l).1 := by + induction l generalizing m with + | nil => + simp[insertManyListₘ] + exact h + | cons hd tl ih => + simp[insertManyListₘ] + apply ih + apply wfImp_insert h + +theorem toListModel_insertManyListₘ [BEq α] [Hashable α] [EquivBEq α][LawfulHashable α] {m : Raw₀ α β} {l: List ((a : α) × β a)} (h : Raw.WFImp m.1): + Perm (toListModel (insertManyListₘ m l).1.buckets) (List.insertMany (toListModel m.1.buckets) l) := by +induction l generalizing m with +| nil => + simp[insertManyListₘ, List.insertMany] +| cons hd tl ih => + simp[insertManyListₘ, List.insertMany] + apply Perm.trans + apply ih (wfImp_insert h) + apply List.insertMany_perm_of_perm_first + apply toListModel_insert h + apply (wfImp_insert h).distinct + +/-! # `insertManyList` -/ +theorem wfImp_insertManyList [BEq α] [Hashable α] [EquivBEq α][LawfulHashable α] {m : Raw₀ α β} {l: List ((a : α) × β a)} (h : Raw.WFImp m.1): Raw.WFImp (insertManyList m l).1 := by + rw [insertManyList_eq_insertManyListₘ] + apply wfImp_insertManyListₘ h + +theorem toListModel_insertManyList [BEq α] [Hashable α] [EquivBEq α][LawfulHashable α] {m : Raw₀ α β} {l: List ((a : α) × β a)} (h : Raw.WFImp m.1): + Perm (toListModel (insertManyList m l).1.buckets) (List.insertMany (toListModel m.1.buckets) l) := by + rw [insertManyList_eq_insertManyListₘ] + apply toListModel_insertManyListₘ h + end Raw₀ namespace Raw @@ -727,6 +761,7 @@ theorem WF.out [BEq α] [Hashable α] [i₁ : EquivBEq α] [i₂ : LawfulHashabl end Raw namespace Raw₀ +/-! # `insertMany` -/ theorem wfImp_insertMany [BEq α] [Hashable α] [EquivBEq α] [LawfulHashable α] {ρ : Type w} [ForIn Id ρ ((a : α) × β a)] {m : Raw₀ α β} {l : ρ} (h : Raw.WFImp m.1) : From ea17bbb0edb61ae67e882a5b0ce1de306078b478 Mon Sep 17 00:00:00 2001 From: jt0202 Date: Sun, 10 Nov 2024 14:52:07 +0100 Subject: [PATCH 02/83] Verified insertManyList_contains for Raw0 --- src/Std/Data/DHashMap/Internal/List/Associative.lean | 8 ++++++++ src/Std/Data/DHashMap/Internal/RawLemmas.lean | 11 ++++++++++- 2 files changed, 18 insertions(+), 1 deletion(-) diff --git a/src/Std/Data/DHashMap/Internal/List/Associative.lean b/src/Std/Data/DHashMap/Internal/List/Associative.lean index 5f8b9551da39..f77abeccfcac 100644 --- a/src/Std/Data/DHashMap/Internal/List/Associative.lean +++ b/src/Std/Data/DHashMap/Internal/List/Associative.lean @@ -1844,4 +1844,12 @@ theorem insertMany_perm_of_perm_first [BEq α] [EquivBEq α] (l1 l2 toInsert: Li exact h apply DistinctKeys.insertEntry distinct +theorem insertMany_containsKey [BEq α] [PartialEquivBEq α] (l toInsert: List ((a : α) × β a)) (k: α): containsKey k (List.insertMany l toInsert) = true ↔ containsKey k l= true ∨ containsKey k toInsert = true := by + induction toInsert generalizing l with + | nil => simp[insertMany] + | cons hd tl ih => + simp[insertMany] + rw [ih, containsKey_insertEntry, containsKey_cons, Bool.or_eq_true_iff, Bool.or_eq_true_iff, or_comm (a:=containsKey k l), or_assoc, or_assoc, or_comm (a:=containsKey k l)] + + end List diff --git a/src/Std/Data/DHashMap/Internal/RawLemmas.lean b/src/Std/Data/DHashMap/Internal/RawLemmas.lean index 837b6c76545c..31607fdd2e48 100644 --- a/src/Std/Data/DHashMap/Internal/RawLemmas.lean +++ b/src/Std/Data/DHashMap/Internal/RawLemmas.lean @@ -832,13 +832,22 @@ theorem contains_keys [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {k : α} : @[simp] theorem mem_keys [LawfulBEq α] [LawfulHashable α] (h : m.1.WF) {k : α} : k ∈ m.1.keys ↔ m.contains k := by - simp_to_model + simp_to_model rw [List.containsKey_eq_keys_contains] theorem distinct_keys [EquivBEq α] [LawfulHashable α] (h : m.1.WF) : m.1.keys.Pairwise (fun a b => (a == b) = false) := by simp_to_model using (Raw.WF.out h).distinct.distinct +theorem insertManyList_contains [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l: List ((a:α) × (β a))} {k: α}: + (m.insertManyList l).contains k ↔ m.contains k ∨ List.containsKey k l := by + rw [contains_eq_containsKey,contains_eq_containsKey, containsKey_of_perm (l':= List.insertMany (toListModel m.1.buckets) l)] + apply insertMany_containsKey + apply toListModel_insertManyList (Raw.WF.out h) + apply (Raw.WF.out h) + apply wfImp_insertManyList (Raw.WF.out h) + + end Raw₀ end Std.DHashMap.Internal From ed56c5e70e32b99e2f255f9320776afd4398f7f4 Mon Sep 17 00:00:00 2001 From: jt0202 Date: Sun, 10 Nov 2024 17:35:35 +0100 Subject: [PATCH 03/83] verified InsertList together with insert/size --- src/Std/Data/DHashMap/Internal/Defs.lean | 10 +-- src/Std/Data/DHashMap/Internal/Model.lean | 12 ++-- src/Std/Data/DHashMap/Internal/RawLemmas.lean | 70 +++++++++++++++++-- src/Std/Data/DHashMap/Internal/WF.lean | 32 ++++----- 4 files changed, 92 insertions(+), 32 deletions(-) diff --git a/src/Std/Data/DHashMap/Internal/Defs.lean b/src/Std/Data/DHashMap/Internal/Defs.lean index aac8dd26d1d2..51b602b026fc 100644 --- a/src/Std/Data/DHashMap/Internal/Defs.lean +++ b/src/Std/Data/DHashMap/Internal/Defs.lean @@ -353,12 +353,12 @@ where return r /-- Internal implementation detail of the hash map -/ -def insertManyList [BEq α] [Hashable α] +def insertList [BEq α] [Hashable α] (m : Raw₀ α β) (l : List ((a : α) × β a)) : Raw₀ α β := Id.run do -let mut r := m -for ⟨a,b⟩ in l do - r:= r.insert a b -return r + let mut r := m + for ⟨a,b⟩ in l do + r:= r.insert a b + return r section diff --git a/src/Std/Data/DHashMap/Internal/Model.lean b/src/Std/Data/DHashMap/Internal/Model.lean index fbb0d971da0c..c7f2b38264f0 100644 --- a/src/Std/Data/DHashMap/Internal/Model.lean +++ b/src/Std/Data/DHashMap/Internal/Model.lean @@ -323,10 +323,10 @@ def filterₘ (m : Raw₀ α β) (f : (a : α) → β a → Bool) : Raw₀ α β ⟨withComputedSize (updateAllBuckets m.1.buckets fun l => l.filter f), by simpa using m.2⟩ /-- Internal implementation detail of the hash map -/ -def insertManyListₘ [BEq α] [Hashable α](m : Raw₀ α β) (l: List ((a : α) × β a)): Raw₀ α β := +def insertListₘ [BEq α] [Hashable α](m : Raw₀ α β) (l: List ((a : α) × β a)): Raw₀ α β := match l with | .nil => m - | .cons hd tl => insertManyListₘ (m.insert hd.1 hd.2) tl + | .cons hd tl => insertListₘ (m.insert hd.1 hd.2) tl section @@ -461,12 +461,12 @@ theorem map_eq_mapₘ (m : Raw₀ α β) (f : (a : α) → β a → δ a) : theorem filter_eq_filterₘ (m : Raw₀ α β) (f : (a : α) → β a → Bool) : m.filter f = m.filterₘ f := rfl -theorem insertManyList_eq_insertManyListₘ [BEq α] [Hashable α](m : Raw₀ α β) (l: List ((a : α) × β a)): insertManyList m l = insertManyListₘ m l := by - simp [insertManyList, Id.run] +theorem insertList_eq_insertListₘ [BEq α] [Hashable α](m : Raw₀ α β) (l: List ((a : α) × β a)): insertList m l = insertListₘ m l := by + simp [insertList, Id.run] induction l generalizing m with - | nil => simp[insertManyListₘ] + | nil => simp[insertListₘ] | cons hd tl ih => - simp [insertManyListₘ] + simp [insertListₘ] rw [ih] section diff --git a/src/Std/Data/DHashMap/Internal/RawLemmas.lean b/src/Std/Data/DHashMap/Internal/RawLemmas.lean index 31607fdd2e48..5b20c928eebc 100644 --- a/src/Std/Data/DHashMap/Internal/RawLemmas.lean +++ b/src/Std/Data/DHashMap/Internal/RawLemmas.lean @@ -839,14 +839,74 @@ theorem distinct_keys [EquivBEq α] [LawfulHashable α] (h : m.1.WF) : m.1.keys.Pairwise (fun a b => (a == b) = false) := by simp_to_model using (Raw.WF.out h).distinct.distinct -theorem insertManyList_contains [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l: List ((a:α) × (β a))} {k: α}: - (m.insertManyList l).contains k ↔ m.contains k ∨ List.containsKey k l := by +@[simp] +theorem insertList_nil: m.insertList [] = m := by + simp[insertList, Id.run] + +@[simp] +theorem insertList_singleton {k: α} {v: β k}: m.insertList [⟨k,v⟩] = m.insert k v := by + simp[insertList, Id.run] + +theorem insert_insertList {l: List ((a:α) × (β a))} {k: α} {v: β k}: (m.insert k v).insertList l = m.insertList (⟨k,v⟩::l) := by + cases l with + | nil => simp[insertList] + | cons hd tl => simp[insertList] + +theorem insertList_insertList {l1 l2: List ((a:α) × (β a))}: + (m.insertList l1).insertList l2 = m.insertList (l1++l2) := by + simp[insertList_eq_insertListₘ] + induction l1 generalizing m with + | nil => simp[insertListₘ] + | cons hd tl ih => + simp[insertListₘ] + apply ih + +theorem insertList_insert {l: List ((a:α) × (β a))} {k: α} {v: β k}: + (m.insertList l).insert k v = m.insertList (l ++ [⟨k,v⟩]) := by + simp [insertList_eq_insertListₘ] + induction l generalizing m with + | nil => simp[insertListₘ] + | cons hd tl ih => + simp[insertListₘ] + rw [ih] + +theorem insertList_contains [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l: List ((a:α) × (β a))} {k: α}: + (m.insertList l).contains k ↔ m.contains k ∨ List.containsKey k l := by rw [contains_eq_containsKey,contains_eq_containsKey, containsKey_of_perm (l':= List.insertMany (toListModel m.1.buckets) l)] apply insertMany_containsKey - apply toListModel_insertManyList (Raw.WF.out h) + apply toListModel_insertList (Raw.WF.out h) apply (Raw.WF.out h) - apply wfImp_insertManyList (Raw.WF.out h) - + apply wfImp_insertList (Raw.WF.out h) + +theorem insertList_size [LawfulBEq α][LawfulHashable α]{l: List ((a:α) × (β a))} {distinct: DistinctKeys l} {distinct2: ∀ (a:α), ¬ (m.contains a = true ∧ List.containsKey a l = true)} (h: m.1.WF): (m.insertList l).1.size = m.1.size + l.length := by + simp[insertList_eq_insertListₘ] + induction l generalizing m with + | nil => simp[insertListₘ] + | cons hd tl ih => + simp[insertListₘ] + rw [ih, size_insert] + · split + · rename_i h + specialize distinct2 hd.1 + simp[h, containsKey] at distinct2 + · rw [Nat.add_assoc, Nat.add_comm 1 _] + · exact h + · apply DistinctKeys.tail distinct + · intro a + rw [contains_insert _ h] + simp + specialize distinct2 a + rw [containsKey_cons] at distinct2 + simp at distinct2 + intro h' + cases h' with + | inl h' => + rw [← h'] + apply DistinctKeys.containsKey_eq_false distinct + | inr h' => + specialize distinct2 h' + apply And.right distinct2 + · apply Raw.WF.insert₀ h end Raw₀ diff --git a/src/Std/Data/DHashMap/Internal/WF.lean b/src/Std/Data/DHashMap/Internal/WF.lean index 9e47c3f76b0b..ef29a2c010f9 100644 --- a/src/Std/Data/DHashMap/Internal/WF.lean +++ b/src/Std/Data/DHashMap/Internal/WF.lean @@ -706,39 +706,39 @@ theorem wfImp_filter [BEq α] [Hashable α] [EquivBEq α] [LawfulHashable α] {m rw [filter_eq_filterₘ] exact wfImp_filterₘ h -/-! # `insertManyListₘ` -/ -theorem wfImp_insertManyListₘ [BEq α] [Hashable α] [EquivBEq α][LawfulHashable α] {m : Raw₀ α β} {l: List ((a : α) × β a)} (h : Raw.WFImp m.1): Raw.WFImp (insertManyListₘ m l).1 := by +/-! # `insertListₘ` -/ +theorem wfImp_insertListₘ [BEq α] [Hashable α] [EquivBEq α][LawfulHashable α] {m : Raw₀ α β} {l: List ((a : α) × β a)} (h : Raw.WFImp m.1): Raw.WFImp (insertListₘ m l).1 := by induction l generalizing m with | nil => - simp[insertManyListₘ] + simp[insertListₘ] exact h | cons hd tl ih => - simp[insertManyListₘ] + simp[insertListₘ] apply ih apply wfImp_insert h -theorem toListModel_insertManyListₘ [BEq α] [Hashable α] [EquivBEq α][LawfulHashable α] {m : Raw₀ α β} {l: List ((a : α) × β a)} (h : Raw.WFImp m.1): - Perm (toListModel (insertManyListₘ m l).1.buckets) (List.insertMany (toListModel m.1.buckets) l) := by +theorem toListModel_insertListₘ [BEq α] [Hashable α] [EquivBEq α][LawfulHashable α] {m : Raw₀ α β} {l: List ((a : α) × β a)} (h : Raw.WFImp m.1): + Perm (toListModel (insertListₘ m l).1.buckets) (List.insertMany (toListModel m.1.buckets) l) := by induction l generalizing m with | nil => - simp[insertManyListₘ, List.insertMany] + simp[insertListₘ, List.insertMany] | cons hd tl ih => - simp[insertManyListₘ, List.insertMany] + simp[insertListₘ, List.insertMany] apply Perm.trans apply ih (wfImp_insert h) apply List.insertMany_perm_of_perm_first apply toListModel_insert h apply (wfImp_insert h).distinct -/-! # `insertManyList` -/ -theorem wfImp_insertManyList [BEq α] [Hashable α] [EquivBEq α][LawfulHashable α] {m : Raw₀ α β} {l: List ((a : α) × β a)} (h : Raw.WFImp m.1): Raw.WFImp (insertManyList m l).1 := by - rw [insertManyList_eq_insertManyListₘ] - apply wfImp_insertManyListₘ h +/-! # `insertList` -/ +theorem wfImp_insertList [BEq α] [Hashable α] [EquivBEq α][LawfulHashable α] {m : Raw₀ α β} {l: List ((a : α) × β a)} (h : Raw.WFImp m.1): Raw.WFImp (insertList m l).1 := by + rw [insertList_eq_insertListₘ] + apply wfImp_insertListₘ h -theorem toListModel_insertManyList [BEq α] [Hashable α] [EquivBEq α][LawfulHashable α] {m : Raw₀ α β} {l: List ((a : α) × β a)} (h : Raw.WFImp m.1): - Perm (toListModel (insertManyList m l).1.buckets) (List.insertMany (toListModel m.1.buckets) l) := by - rw [insertManyList_eq_insertManyListₘ] - apply toListModel_insertManyListₘ h +theorem toListModel_insertList [BEq α] [Hashable α] [EquivBEq α][LawfulHashable α] {m : Raw₀ α β} {l: List ((a : α) × β a)} (h : Raw.WFImp m.1): + Perm (toListModel (insertList m l).1.buckets) (List.insertMany (toListModel m.1.buckets) l) := by + rw [insertList_eq_insertListₘ] + apply toListModel_insertListₘ h end Raw₀ From db8a6427f749510b884364af322084d7a7ad8071 Mon Sep 17 00:00:00 2001 From: jt0202 Date: Sun, 10 Nov 2024 20:49:03 +0100 Subject: [PATCH 04/83] Renamed List.insertMany to List.insertList and added Perm result for Raw0.insertList --- .../DHashMap/Internal/List/Associative.lean | 83 +++++++++++++++++-- src/Std/Data/DHashMap/Internal/RawLemmas.lean | 4 +- src/Std/Data/DHashMap/Internal/WF.lean | 10 +-- 3 files changed, 82 insertions(+), 15 deletions(-) diff --git a/src/Std/Data/DHashMap/Internal/List/Associative.lean b/src/Std/Data/DHashMap/Internal/List/Associative.lean index f77abeccfcac..ce9dcd77b0a4 100644 --- a/src/Std/Data/DHashMap/Internal/List/Associative.lean +++ b/src/Std/Data/DHashMap/Internal/List/Associative.lean @@ -257,6 +257,26 @@ theorem containsKey_eq_isSome_getEntry? [BEq α] {l : List ((a : α) × β a)} { · simp [getEntry?_cons_of_false h, h, ih] · simp [getEntry?_cons_of_true h, h] +theorem containsKey_of_eq [BEq α] [PartialEquivBEq α]{l : List ((a : α) × β a)} {a b: α} (eq: a == b): containsKey a l ↔ containsKey b l := by + induction l with + | nil=> simp + | cons hd tl ih => + simp[containsKey] + by_cases hd_a: hd.fst == a + · have hd_b: hd.fst == b := by + apply PartialEquivBEq.trans hd_a eq + simp [hd_a, hd_b] + · have hd_b: (hd.fst == b) = false := by + false_or_by_contra + rename_i h + simp at h + have hd_a': hd.fst == a := by + apply PartialEquivBEq.trans h (PartialEquivBEq.symm eq) + contradiction + simp [hd_a, hd_b] + rw [Bool.eq_iff_iff] + apply ih + theorem isEmpty_eq_false_of_containsKey [BEq α] {l : List ((a : α) × β a)} {a : α} (h : containsKey a l = true) : l.isEmpty = false := by cases l <;> simp_all @@ -979,6 +999,13 @@ theorem insertEntry_of_containsKey_eq_false [BEq α] {l : List ((a : α) × β a (h : containsKey k l = false) : insertEntry k v l = ⟨k, v⟩ :: l := by simp [insertEntry, h] +theorem insertEntry_of_containsKey_eq_false_perm [BEq α] {l : List ((a : α) × β a)} {k : α} {v : β k} + (h : containsKey k l = false) : Perm (insertEntry k v l) (⟨k,v⟩::l) := by + cases l with + | nil => simp + | cons hd tl => simp[insertEntry, h] + + theorem DistinctKeys.insertEntry [BEq α] [PartialEquivBEq α] {l : List ((a : α) × β a)} {k : α} {v : β k} (h : DistinctKeys l) : DistinctKeys (insertEntry k v l) := by cases h' : containsKey k l @@ -1828,28 +1855,68 @@ theorem eraseKey_append_of_containsKey_right_eq_false [BEq α] {l l' : List ((a · rw [cond_true, cond_true] /-- Internal implementation detail of the hash map -/ -def insertMany [BEq α] (l toInsert: List ((a : α) × β a)) : List ((a : α) × β a) := +def insertList [BEq α] (l toInsert: List ((a : α) × β a)) : List ((a : α) × β a) := match toInsert with | .nil => l - | .cons ⟨k, v⟩ toInsert => insertMany (insertEntry k v l) toInsert + | .cons ⟨k, v⟩ toInsert => insertList (insertEntry k v l) toInsert -theorem insertMany_perm_of_perm_first [BEq α] [EquivBEq α] (l1 l2 toInsert: List ((a : α) × β a)) (h: Perm l1 l2) (distinct: DistinctKeys l1): Perm (insertMany l1 toInsert) (insertMany l2 toInsert) := by +theorem insertList_perm_of_perm_first [BEq α] [EquivBEq α] (l1 l2 toInsert: List ((a : α) × β a)) (h: Perm l1 l2) (distinct: DistinctKeys l1): Perm (insertList l1 toInsert) (insertList l2 toInsert) := by induction toInsert generalizing l1 l2 with - | nil => simp[insertMany, h] + | nil => simp[insertList, h] | cons hd tl ih => - simp[insertMany] + simp[insertList] apply ih apply insertEntry_of_perm exact distinct exact h apply DistinctKeys.insertEntry distinct -theorem insertMany_containsKey [BEq α] [PartialEquivBEq α] (l toInsert: List ((a : α) × β a)) (k: α): containsKey k (List.insertMany l toInsert) = true ↔ containsKey k l= true ∨ containsKey k toInsert = true := by +theorem insertList_containsKey [BEq α] [PartialEquivBEq α] (l toInsert: List ((a : α) × β a)) (k: α): containsKey k (List.insertList l toInsert) = true ↔ containsKey k l= true ∨ containsKey k toInsert = true := by induction toInsert generalizing l with - | nil => simp[insertMany] + | nil => simp[insertList] | cons hd tl ih => - simp[insertMany] + simp[insertList] rw [ih, containsKey_insertEntry, containsKey_cons, Bool.or_eq_true_iff, Bool.or_eq_true_iff, or_comm (a:=containsKey k l), or_assoc, or_assoc, or_comm (a:=containsKey k l)] +theorem insertList_perm [BEq α] [ReflBEq α] [PartialEquivBEq α] (l toInsert: List ((a : α) × β a)) (distinct_l: DistinctKeys l) (distinct_toInsert: DistinctKeys toInsert) (distinct_both: ∀ (a:α), ¬ containsKey a l ∨ ¬ containsKey a toInsert): + Perm (insertList l toInsert) (l++toInsert) := by + induction toInsert generalizing l with + | nil => simp[insertList] + | cons hd tl ih => + simp[insertList] + specialize ih (insertEntry hd.fst hd.snd l) + apply Perm.trans + · apply ih + · simp[insertEntry] + specialize distinct_both hd.1 + simp[containsKey] at distinct_both + simp [distinct_both] + exact distinct_l + · apply DistinctKeys.tail distinct_toInsert + · intro a + specialize distinct_both a + simp[containsKey] at distinct_both + simp[containsKey] + by_cases hd_a: hd.fst == a + · rw [distinctKeys_cons_iff] at distinct_toInsert + have contains_a_tl: containsKey a tl = false := by + false_or_by_contra + rename_i h + simp at h + rw [containsKey_of_eq (PartialEquivBEq.symm hd_a)] at h + simp [h] at distinct_toInsert + simp [contains_a_tl] + · simp[hd_a] + simp at hd_a + simp [hd_a] at distinct_both + apply distinct_both + · simp[insertEntry] + specialize distinct_both hd.1 + simp[containsKey] at distinct_both + simp [distinct_both] + have h_hd: hd = ⟨hd.fst, hd.snd⟩:= rfl + rw [← h_hd] + apply Perm.symm + apply List.perm_middle end List diff --git a/src/Std/Data/DHashMap/Internal/RawLemmas.lean b/src/Std/Data/DHashMap/Internal/RawLemmas.lean index 5b20c928eebc..806b31f0feca 100644 --- a/src/Std/Data/DHashMap/Internal/RawLemmas.lean +++ b/src/Std/Data/DHashMap/Internal/RawLemmas.lean @@ -872,8 +872,8 @@ theorem insertList_insert {l: List ((a:α) × (β a))} {k: α} {v: β k}: theorem insertList_contains [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l: List ((a:α) × (β a))} {k: α}: (m.insertList l).contains k ↔ m.contains k ∨ List.containsKey k l := by - rw [contains_eq_containsKey,contains_eq_containsKey, containsKey_of_perm (l':= List.insertMany (toListModel m.1.buckets) l)] - apply insertMany_containsKey + rw [contains_eq_containsKey,contains_eq_containsKey, containsKey_of_perm (l':= List.insertList (toListModel m.1.buckets) l)] + apply insertList_containsKey apply toListModel_insertList (Raw.WF.out h) apply (Raw.WF.out h) apply wfImp_insertList (Raw.WF.out h) diff --git a/src/Std/Data/DHashMap/Internal/WF.lean b/src/Std/Data/DHashMap/Internal/WF.lean index ef29a2c010f9..075cd1c47887 100644 --- a/src/Std/Data/DHashMap/Internal/WF.lean +++ b/src/Std/Data/DHashMap/Internal/WF.lean @@ -718,15 +718,15 @@ theorem wfImp_insertListₘ [BEq α] [Hashable α] [EquivBEq α][LawfulHashable apply wfImp_insert h theorem toListModel_insertListₘ [BEq α] [Hashable α] [EquivBEq α][LawfulHashable α] {m : Raw₀ α β} {l: List ((a : α) × β a)} (h : Raw.WFImp m.1): - Perm (toListModel (insertListₘ m l).1.buckets) (List.insertMany (toListModel m.1.buckets) l) := by + Perm (toListModel (insertListₘ m l).1.buckets) (List.insertList (toListModel m.1.buckets) l) := by induction l generalizing m with | nil => - simp[insertListₘ, List.insertMany] + simp[insertListₘ, List.insertList] | cons hd tl ih => - simp[insertListₘ, List.insertMany] + simp[insertListₘ, List.insertList] apply Perm.trans apply ih (wfImp_insert h) - apply List.insertMany_perm_of_perm_first + apply List.insertList_perm_of_perm_first apply toListModel_insert h apply (wfImp_insert h).distinct @@ -736,7 +736,7 @@ theorem wfImp_insertList [BEq α] [Hashable α] [EquivBEq α][LawfulHashable α] apply wfImp_insertListₘ h theorem toListModel_insertList [BEq α] [Hashable α] [EquivBEq α][LawfulHashable α] {m : Raw₀ α β} {l: List ((a : α) × β a)} (h : Raw.WFImp m.1): - Perm (toListModel (insertList m l).1.buckets) (List.insertMany (toListModel m.1.buckets) l) := by + Perm (toListModel (insertList m l).1.buckets) (List.insertList (toListModel m.1.buckets) l) := by rw [insertList_eq_insertListₘ] apply toListModel_insertListₘ h From 2aecd5a5bfc41baf5ff2f3d3ef77dff2bfff8922 Mon Sep 17 00:00:00 2001 From: monsterkrampe Date: Mon, 11 Nov 2024 15:06:19 +0100 Subject: [PATCH 05/83] Use simp_to_model for insertList_contains --- .../DHashMap/Internal/List/Associative.lean | 11 +++- src/Std/Data/DHashMap/Internal/RawLemmas.lean | 58 +++++++------------ src/Std/Data/DHashMap/Internal/WF.lean | 4 +- 3 files changed, 30 insertions(+), 43 deletions(-) diff --git a/src/Std/Data/DHashMap/Internal/List/Associative.lean b/src/Std/Data/DHashMap/Internal/List/Associative.lean index ce9dcd77b0a4..f3fae142ebf5 100644 --- a/src/Std/Data/DHashMap/Internal/List/Associative.lean +++ b/src/Std/Data/DHashMap/Internal/List/Associative.lean @@ -1871,12 +1871,17 @@ theorem insertList_perm_of_perm_first [BEq α] [EquivBEq α] (l1 l2 toInsert: Li exact h apply DistinctKeys.insertEntry distinct -theorem insertList_containsKey [BEq α] [PartialEquivBEq α] (l toInsert: List ((a : α) × β a)) (k: α): containsKey k (List.insertList l toInsert) = true ↔ containsKey k l= true ∨ containsKey k toInsert = true := by +theorem insertList_containsKey [BEq α] [PartialEquivBEq α] (l toInsert : List ((a : α) × β a)) (k: α): containsKey k (List.insertList l toInsert) ↔ containsKey k l ∨ (toInsert.map Sigma.fst).contains k := by induction toInsert generalizing l with | nil => simp[insertList] | cons hd tl ih => - simp[insertList] - rw [ih, containsKey_insertEntry, containsKey_cons, Bool.or_eq_true_iff, Bool.or_eq_true_iff, or_comm (a:=containsKey k l), or_assoc, or_assoc, or_comm (a:=containsKey k l)] + unfold insertList + rw [ih] + rw [containsKey_insertEntry] + simp + rw [BEq.comm] + conv => left; left; rw [or_comm] + rw [or_assoc] theorem insertList_perm [BEq α] [ReflBEq α] [PartialEquivBEq α] (l toInsert: List ((a : α) × β a)) (distinct_l: DistinctKeys l) (distinct_toInsert: DistinctKeys toInsert) (distinct_both: ∀ (a:α), ¬ containsKey a l ∨ ¬ containsKey a toInsert): Perm (insertList l toInsert) (l++toInsert) := by diff --git a/src/Std/Data/DHashMap/Internal/RawLemmas.lean b/src/Std/Data/DHashMap/Internal/RawLemmas.lean index 806b31f0feca..c622754bc548 100644 --- a/src/Std/Data/DHashMap/Internal/RawLemmas.lean +++ b/src/Std/Data/DHashMap/Internal/RawLemmas.lean @@ -70,7 +70,7 @@ variable [BEq α] [Hashable α] /-- Internal implementation detail of the hash map -/ scoped macro "wf_trivial" : tactic => `(tactic| repeat (first - | apply Raw₀.wfImp_insert | apply Raw₀.wfImp_insertIfNew | apply Raw₀.wfImp_erase + | apply Raw₀.wfImp_insert | apply Raw₀.wfImp_insertIfNew | apply Raw₀.wfImp_insertList | apply Raw₀.wfImp_erase | apply Raw.WF.out | assumption | apply Raw₀.wfImp_empty | apply Raw.WFImp.distinct | apply Raw.WF.empty₀)) @@ -90,7 +90,7 @@ private def queryNames : Array Name := ``Raw.pairwise_keys_iff_pairwise_keys] private def modifyNames : Array Name := - #[``toListModel_insert, ``toListModel_erase, ``toListModel_insertIfNew] + #[``toListModel_insert, ``toListModel_erase, ``toListModel_insertIfNew, ``toListModel_insertList] private def congrNames : MacroM (Array (TSyntax `term)) := do return #[← `(_root_.List.Perm.isEmpty_eq), ← `(containsKey_of_perm), @@ -840,14 +840,15 @@ theorem distinct_keys [EquivBEq α] [LawfulHashable α] (h : m.1.WF) : simp_to_model using (Raw.WF.out h).distinct.distinct @[simp] -theorem insertList_nil: m.insertList [] = m := by +theorem insertList_nil : m.insertList [] = m := by simp[insertList, Id.run] @[simp] theorem insertList_singleton {k: α} {v: β k}: m.insertList [⟨k,v⟩] = m.insert k v := by simp[insertList, Id.run] -theorem insert_insertList {l: List ((a:α) × (β a))} {k: α} {v: β k}: (m.insert k v).insertList l = m.insertList (⟨k,v⟩::l) := by +@[simp] +theorem insertList_cons {l: List ((a:α) × (β a))} {k: α} {v: β k}: m.insertList (⟨k,v⟩::l) = (m.insert k v).insertList l := by cases l with | nil => simp[insertList] | cons hd tl => simp[insertList] @@ -871,43 +872,24 @@ theorem insertList_insert {l: List ((a:α) × (β a))} {k: α} {v: β k}: rw [ih] theorem insertList_contains [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l: List ((a:α) × (β a))} {k: α}: - (m.insertList l).contains k ↔ m.contains k ∨ List.containsKey k l := by - rw [contains_eq_containsKey,contains_eq_containsKey, containsKey_of_perm (l':= List.insertList (toListModel m.1.buckets) l)] - apply insertList_containsKey - apply toListModel_insertList (Raw.WF.out h) - apply (Raw.WF.out h) - apply wfImp_insertList (Raw.WF.out h) + (m.insertList l).contains k ↔ m.contains k ∨ (l.map Sigma.fst).contains k := by + simp_to_model using List.insertList_containsKey theorem insertList_size [LawfulBEq α][LawfulHashable α]{l: List ((a:α) × (β a))} {distinct: DistinctKeys l} {distinct2: ∀ (a:α), ¬ (m.contains a = true ∧ List.containsKey a l = true)} (h: m.1.WF): (m.insertList l).1.size = m.1.size + l.length := by - simp[insertList_eq_insertListₘ] - induction l generalizing m with - | nil => simp[insertListₘ] - | cons hd tl ih => - simp[insertListₘ] - rw [ih, size_insert] - · split - · rename_i h - specialize distinct2 hd.1 - simp[h, containsKey] at distinct2 - · rw [Nat.add_assoc, Nat.add_comm 1 _] - · exact h - · apply DistinctKeys.tail distinct - · intro a - rw [contains_insert _ h] + simp_to_model + rw [← List.length_append] + apply List.Perm.length_eq + apply List.insertList_perm + . apply (Raw.WF.out h).distinct + . exact distinct + . simp at distinct2 + intro a + cases eq : containsKey a (toListModel m.val.buckets) with + | false => simp + | true => simp - specialize distinct2 a - rw [containsKey_cons] at distinct2 - simp at distinct2 - intro h' - cases h' with - | inl h' => - rw [← h'] - apply DistinctKeys.containsKey_eq_false distinct - | inr h' => - specialize distinct2 h' - apply And.right distinct2 - · apply Raw.WF.insert₀ h - + apply distinct2 + simp_to_model end Raw₀ end Std.DHashMap.Internal diff --git a/src/Std/Data/DHashMap/Internal/WF.lean b/src/Std/Data/DHashMap/Internal/WF.lean index 075cd1c47887..a20634c76906 100644 --- a/src/Std/Data/DHashMap/Internal/WF.lean +++ b/src/Std/Data/DHashMap/Internal/WF.lean @@ -707,7 +707,7 @@ theorem wfImp_filter [BEq α] [Hashable α] [EquivBEq α] [LawfulHashable α] {m exact wfImp_filterₘ h /-! # `insertListₘ` -/ -theorem wfImp_insertListₘ [BEq α] [Hashable α] [EquivBEq α][LawfulHashable α] {m : Raw₀ α β} {l: List ((a : α) × β a)} (h : Raw.WFImp m.1): Raw.WFImp (insertListₘ m l).1 := by +theorem wfImp_insertListₘ [BEq α] [Hashable α] [EquivBEq α][LawfulHashable α] {m : Raw₀ α β} {l: List ((a : α) × β a)} (h : Raw.WFImp m.1): Raw.WFImp (m.insertListₘ l).1 := by induction l generalizing m with | nil => simp[insertListₘ] @@ -731,7 +731,7 @@ induction l generalizing m with apply (wfImp_insert h).distinct /-! # `insertList` -/ -theorem wfImp_insertList [BEq α] [Hashable α] [EquivBEq α][LawfulHashable α] {m : Raw₀ α β} {l: List ((a : α) × β a)} (h : Raw.WFImp m.1): Raw.WFImp (insertList m l).1 := by +theorem wfImp_insertList [BEq α] [Hashable α] [EquivBEq α][LawfulHashable α] {m : Raw₀ α β} {l: List ((a : α) × β a)} (h : Raw.WFImp m.1): Raw.WFImp (m.insertList l).1 := by rw [insertList_eq_insertListₘ] apply wfImp_insertListₘ h From 3845a2e137b7b34fff16b9b5352dfb35ee25a21c Mon Sep 17 00:00:00 2001 From: monsterkrampe Date: Wed, 13 Nov 2024 16:29:46 +0100 Subject: [PATCH 06/83] Add get_insertList result to internal hashmap --- .../DHashMap/Internal/List/Associative.lean | 76 ++++++++++++++++++- src/Std/Data/DHashMap/Internal/RawLemmas.lean | 16 +++- 2 files changed, 88 insertions(+), 4 deletions(-) diff --git a/src/Std/Data/DHashMap/Internal/List/Associative.lean b/src/Std/Data/DHashMap/Internal/List/Associative.lean index f3fae142ebf5..1bc97caebf10 100644 --- a/src/Std/Data/DHashMap/Internal/List/Associative.lean +++ b/src/Std/Data/DHashMap/Internal/List/Associative.lean @@ -16,6 +16,48 @@ the contents of this file. File contents: Verification of associative lists -/ +-- TODO: this does not belong here; maybe create a new datatype for such dependant associative lists? +def List.get_by_key {α : Type u} {β : α -> Type v} [BEq α] [LawfulBEq α] (l: List ((a:α) × (β a))) (k: α) (h : (l.map Sigma.fst).contains k) : β k := + match l with + | .nil => by simp at h + | .cons ⟨k', v⟩ l => if h' : k == k' + then + have : k = k' := LawfulBEq.eq_of_beq h' + cast (by rw [this]) v + else List.get_by_key l k (by unfold List.map at h; rw [List.contains_cons] at h; rw [Bool.or_eq_true_iff] at h; cases h with | inl hl => simp at hl; rw [hl] at h'; simp at h' | inr _ => assumption) + +theorem List.get_by_key_append_elem {α : Type u} {β : α -> Type v} [BEq α] [LawfulBEq α] (l: List ((a:α) × (β a))) (elem: ((a:α) × (β a))) (k: α) (h : (l.map Sigma.fst).contains k) : ∀ h', (l ++ [elem]).get_by_key k h' = l.get_by_key k h := by + intro h' + induction l with + | nil => simp at h + | cons elem' l ih => + simp [get_by_key] + split + . rfl + . apply ih + +theorem List.get_by_key_append_elem' {α : Type u} {β : α -> Type v} [BEq α] [LawfulBEq α] (l: List ((a:α) × (β a))) (elem: ((a:α) × (β a))) (k: α) (h : ¬(l.map Sigma.fst).contains k) : ∀ h', (l ++ [elem]).get_by_key k h' = cast (by simp at h'; simp at h; cases h' with + | inl h' => cases h' with | intro a h' => specialize h a h'.left h'.right; simp at h + | inr h' => rw [h']) elem.snd := +by + intro h' + induction l with + | nil => simp at h'; simp [get_by_key, h'] + | cons elem' l ih => + simp [get_by_key] + simp at h + simp at h' + cases h' with + | inl h' => rw [h'] at h; simp at h + | inr h' => + cases h' with + | inl h' => cases h' with | intro a h' => have hr := h.right; specialize hr a h'.left h'.right; simp at hr + | inr h' => + simp [h.left] + apply ih + simp + apply h.right + set_option linter.missingDocs true set_option autoImplicit false @@ -1871,7 +1913,7 @@ theorem insertList_perm_of_perm_first [BEq α] [EquivBEq α] (l1 l2 toInsert: Li exact h apply DistinctKeys.insertEntry distinct -theorem insertList_containsKey [BEq α] [PartialEquivBEq α] (l toInsert : List ((a : α) × β a)) (k: α): containsKey k (List.insertList l toInsert) ↔ containsKey k l ∨ (toInsert.map Sigma.fst).contains k := by +theorem containsKey_insertList [BEq α] [PartialEquivBEq α] (l toInsert : List ((a : α) × β a)) (k: α): containsKey k (List.insertList l toInsert) ↔ containsKey k l ∨ (toInsert.map Sigma.fst).contains k := by induction toInsert generalizing l with | nil => simp[insertList] | cons hd tl ih => @@ -1883,6 +1925,38 @@ theorem insertList_containsKey [BEq α] [PartialEquivBEq α] (l toInsert : List conv => left; left; rw [or_comm] rw [or_assoc] +theorem containsKey_of_containsKey_insertList [BEq α] [PartialEquivBEq α] + {l toInsert : List ((a : α) × β a)} {k : α} (h₁ : containsKey k (insertList l toInsert)) + (h₂ : (toInsert.map Sigma.fst).contains k = false) : containsKey k l := by + rw [containsKey_insertList, h₂] at h₁; simp at h₁; exact h₁ + +theorem getValueCast_insertList [BEq α] [LawfulBEq α] + {l toInsert : List ((a : α) × β a)} {k : α} + {h} : getValueCast k (insertList l toInsert) h = + if h' : (toInsert.map Sigma.fst).contains k then List.get_by_key toInsert.reverse k (by simp; simp at h'; exact h') + else getValueCast k l (containsKey_of_containsKey_insertList h (Bool.eq_false_iff.2 h')) := by + induction toInsert generalizing l with + | nil => simp [insertList] + | cons elem toInsert ih => + simp [insertList] + rw [ih] + split + case isTrue eq => + split + case isTrue eq2 => rw [List.get_by_key_append_elem] + case isFalse eq2 => rw [List.contains_cons] at eq2; rw [eq] at eq2; simp at eq2 + case isFalse eq => + rw [getValueCast_insertEntry] + split + case isTrue eq2 => + split + case isTrue eq3 => rw [List.get_by_key_append_elem']; simp; simp at eq; exact eq + case isFalse eq3 => simp at eq3; simp at eq2; rw [eq2] at eq3; simp at eq3 + case isFalse eq2 => + split + case isTrue eq3 => simp at eq3; simp at eq; cases eq3 with | inl hl => rw [hl] at eq2; simp at eq2 | inr hr => cases hr with | intro a hr => specialize eq a hr.left hr.right; simp at eq + case isFalse eq3 => rfl + theorem insertList_perm [BEq α] [ReflBEq α] [PartialEquivBEq α] (l toInsert: List ((a : α) × β a)) (distinct_l: DistinctKeys l) (distinct_toInsert: DistinctKeys toInsert) (distinct_both: ∀ (a:α), ¬ containsKey a l ∨ ¬ containsKey a toInsert): Perm (insertList l toInsert) (l++toInsert) := by induction toInsert generalizing l with diff --git a/src/Std/Data/DHashMap/Internal/RawLemmas.lean b/src/Std/Data/DHashMap/Internal/RawLemmas.lean index c622754bc548..a19acfb56f44 100644 --- a/src/Std/Data/DHashMap/Internal/RawLemmas.lean +++ b/src/Std/Data/DHashMap/Internal/RawLemmas.lean @@ -871,11 +871,21 @@ theorem insertList_insert {l: List ((a:α) × (β a))} {k: α} {v: β k}: simp[insertListₘ] rw [ih] -theorem insertList_contains [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l: List ((a:α) × (β a))} {k: α}: +theorem contains_insertList [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l: List ((a:α) × (β a))} {k: α}: (m.insertList l).contains k ↔ m.contains k ∨ (l.map Sigma.fst).contains k := by - simp_to_model using List.insertList_containsKey + simp_to_model using List.containsKey_insertList -theorem insertList_size [LawfulBEq α][LawfulHashable α]{l: List ((a:α) × (β a))} {distinct: DistinctKeys l} {distinct2: ∀ (a:α), ¬ (m.contains a = true ∧ List.containsKey a l = true)} (h: m.1.WF): (m.insertList l).1.size = m.1.size + l.length := by +theorem contains_of_contains_insertList [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l: List ((a:α) × (β a))} {k: α} : + (m.insertList l).contains k → (l.map Sigma.fst).contains k = false → m.contains k := by + simp_to_model using List.containsKey_of_containsKey_insertList + +theorem get_insertList [LawfulBEq α] [LawfulHashable α] (h : m.1.WF) {l: List ((a:α) × (β a))} {k: α} {h₁} : + get (m.insertList l) k h₁ = + if h₂ : (l.map Sigma.fst).contains k then List.get_by_key l.reverse k (by simp; simp at h₂; exact h₂) + else get m k (contains_of_contains_insertList _ h h₁ (Bool.eq_false_iff.2 h₂)) := by + simp_to_model using List.getValueCast_insertList + +theorem size_insertList [LawfulBEq α][LawfulHashable α]{l: List ((a:α) × (β a))} {distinct: DistinctKeys l} {distinct2: ∀ (a:α), ¬ (m.contains a = true ∧ List.containsKey a l = true)} (h: m.1.WF): (m.insertList l).1.size = m.1.size + l.length := by simp_to_model rw [← List.length_append] apply List.Perm.length_eq From 42d72e37f0f6a6b06d9a89b326d974c439d18b56 Mon Sep 17 00:00:00 2001 From: monsterkrampe Date: Fri, 15 Nov 2024 09:10:13 +0100 Subject: [PATCH 07/83] Try to show insertMany_eq_insertList --- src/Std/Data/DHashMap/Internal/Defs.lean | 10 ++++++++++ 1 file changed, 10 insertions(+) diff --git a/src/Std/Data/DHashMap/Internal/Defs.lean b/src/Std/Data/DHashMap/Internal/Defs.lean index 51b602b026fc..69f5ba8e6d7e 100644 --- a/src/Std/Data/DHashMap/Internal/Defs.lean +++ b/src/Std/Data/DHashMap/Internal/Defs.lean @@ -360,6 +360,16 @@ def insertList [BEq α] [Hashable α] r:= r.insert a b return r +theorem insertMany_eq_insertList [BEq α] [Hashable α] (m : Raw₀ α β) (l : List ((a : α) × β a)) : (m.insertMany l).val = m.insertList l := by + simp [insertMany, Id.run] + induction l generalizing m with + | nil => simp [insertList, Id.run] + | cons elem l ih => + simp [insertList] + refine Eq.trans (ih (m.insert elem.fst elem.snd)) ?_ + rw [ih (m.insert elem.fst elem.snd)] + sorry + section variable {β : Type v} From a1d965d1ba49ba495dc28a3e76b85ae0d68553e6 Mon Sep 17 00:00:00 2001 From: monsterkrampe Date: Fri, 15 Nov 2024 09:27:19 +0100 Subject: [PATCH 08/83] Add remarks from meeting --- src/Std/Data/DHashMap/Internal/List/Associative.lean | 1 + src/Std/Data/DHashMap/Internal/RawLemmas.lean | 3 ++- 2 files changed, 3 insertions(+), 1 deletion(-) diff --git a/src/Std/Data/DHashMap/Internal/List/Associative.lean b/src/Std/Data/DHashMap/Internal/List/Associative.lean index 1bc97caebf10..2aa666799fa5 100644 --- a/src/Std/Data/DHashMap/Internal/List/Associative.lean +++ b/src/Std/Data/DHashMap/Internal/List/Associative.lean @@ -299,6 +299,7 @@ theorem containsKey_eq_isSome_getEntry? [BEq α] {l : List ((a : α) × β a)} { · simp [getEntry?_cons_of_false h, h, ih] · simp [getEntry?_cons_of_true h, h] +-- TODO: not needed; containsKey_congr instead theorem containsKey_of_eq [BEq α] [PartialEquivBEq α]{l : List ((a : α) × β a)} {a b: α} (eq: a == b): containsKey a l ↔ containsKey b l := by induction l with | nil=> simp diff --git a/src/Std/Data/DHashMap/Internal/RawLemmas.lean b/src/Std/Data/DHashMap/Internal/RawLemmas.lean index a19acfb56f44..33dfdfa3b162 100644 --- a/src/Std/Data/DHashMap/Internal/RawLemmas.lean +++ b/src/Std/Data/DHashMap/Internal/RawLemmas.lean @@ -885,7 +885,8 @@ theorem get_insertList [LawfulBEq α] [LawfulHashable α] (h : m.1.WF) {l: List else get m k (contains_of_contains_insertList _ h h₁ (Bool.eq_false_iff.2 h₂)) := by simp_to_model using List.getValueCast_insertList -theorem size_insertList [LawfulBEq α][LawfulHashable α]{l: List ((a:α) × (β a))} {distinct: DistinctKeys l} {distinct2: ∀ (a:α), ¬ (m.contains a = true ∧ List.containsKey a l = true)} (h: m.1.WF): (m.insertList l).1.size = m.1.size + l.length := by +-- try to revert extra conditions to use simp_to_model_on_them +theorem size_insertList [EquivBEq α] [LawfulHashable α] {l: List ((a:α) × (β a))} {distinct: DistinctKeys l} {distinct2: ∀ (a:α), ¬ (m.contains a = true ∧ List.containsKey a l = true)} (h: m.1.WF): (m.insertList l).1.size = m.1.size + l.length := by simp_to_model rw [← List.length_append] apply List.Perm.length_eq From 83b4e10ef8f279fa6a8828151ae6ac385b52e8a8 Mon Sep 17 00:00:00 2001 From: jt0202 Date: Sat, 16 Nov 2024 14:04:54 +0100 Subject: [PATCH 09/83] Added equality proof between insertMany & insertList by @TwoFX --- src/Std/Data/DHashMap/Internal/Defs.lean | 10 ------ src/Std/Data/DHashMap/Internal/RawLemmas.lean | 31 ++++++++++++++++--- 2 files changed, 27 insertions(+), 14 deletions(-) diff --git a/src/Std/Data/DHashMap/Internal/Defs.lean b/src/Std/Data/DHashMap/Internal/Defs.lean index 69f5ba8e6d7e..51b602b026fc 100644 --- a/src/Std/Data/DHashMap/Internal/Defs.lean +++ b/src/Std/Data/DHashMap/Internal/Defs.lean @@ -360,16 +360,6 @@ def insertList [BEq α] [Hashable α] r:= r.insert a b return r -theorem insertMany_eq_insertList [BEq α] [Hashable α] (m : Raw₀ α β) (l : List ((a : α) × β a)) : (m.insertMany l).val = m.insertList l := by - simp [insertMany, Id.run] - induction l generalizing m with - | nil => simp [insertList, Id.run] - | cons elem l ih => - simp [insertList] - refine Eq.trans (ih (m.insert elem.fst elem.snd)) ?_ - rw [ih (m.insert elem.fst elem.snd)] - sorry - section variable {β : Type v} diff --git a/src/Std/Data/DHashMap/Internal/RawLemmas.lean b/src/Std/Data/DHashMap/Internal/RawLemmas.lean index 33dfdfa3b162..d626af79bd9c 100644 --- a/src/Std/Data/DHashMap/Internal/RawLemmas.lean +++ b/src/Std/Data/DHashMap/Internal/RawLemmas.lean @@ -839,6 +839,29 @@ theorem distinct_keys [EquivBEq α] [LawfulHashable α] (h : m.1.WF) : m.1.keys.Pairwise (fun a b => (a == b) = false) := by simp_to_model using (Raw.WF.out h).distinct.distinct +theorem insertList_eq_foldl + (m : Raw₀ α β) (l : List ((a : α) × β a)) : + insertList m l = l.foldl (init := m) fun m' p => m'.insert p.1 p.2 := by + simp [insertList, Id.run] + +theorem insertMany_val (m : Raw₀ α β) (l : List ((a : α) × β a)) : + (insertMany m l).val = l.foldl (init := m) fun m' p => m'.insert p.1 p.2 := by + simp only [insertMany, Id.run, Id.pure_eq, Id.bind_eq, List.forIn_yield_eq_foldl] + suffices ∀ (t : { m' // ∀ (P : Raw₀ α β → Prop), + (∀ {m'' : Raw₀ α β} {a : α} {b : β a}, P m'' → P (m''.insert a b)) → P m → P m' }), + (List.foldl (fun m' p => ⟨m'.val.insert p.1 p.2, fun P h₁ h₂ => h₁ (m'.2 _ h₁ h₂)⟩) t l).val = + List.foldl (fun m' p => m'.insert p.fst p.snd) t.val l from this _ + intro t + induction l generalizing m with + | nil => simp + | cons h t ih => + simp + rw [ih] + +theorem insertMany_eq_insertList + (m : Raw₀ α β) (l : List ((a : α) × β a)) : (insertMany m l).val = insertList m l := by + rw [insertList_eq_foldl, insertMany_val] + @[simp] theorem insertList_nil : m.insertList [] = m := by simp[insertList, Id.run] @@ -875,7 +898,7 @@ theorem contains_insertList [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l: L (m.insertList l).contains k ↔ m.contains k ∨ (l.map Sigma.fst).contains k := by simp_to_model using List.containsKey_insertList -theorem contains_of_contains_insertList [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l: List ((a:α) × (β a))} {k: α} : +theorem contains_of_contains_insertList [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l: List ((a:α) × (β a))} {k: α} : (m.insertList l).contains k → (l.map Sigma.fst).contains k = false → m.contains k := by simp_to_model using List.containsKey_of_containsKey_insertList @@ -893,11 +916,11 @@ theorem size_insertList [EquivBEq α] [LawfulHashable α] {l: List ((a:α) × ( apply List.insertList_perm . apply (Raw.WF.out h).distinct . exact distinct - . simp at distinct2 + . simp at distinct2 intro a - cases eq : containsKey a (toListModel m.val.buckets) with + cases eq : containsKey a (toListModel m.val.buckets) with | false => simp - | true => + | true => simp apply distinct2 simp_to_model From ed9f63a3c891adddd28c765edd83785478b9357a Mon Sep 17 00:00:00 2001 From: jt0202 Date: Sat, 16 Nov 2024 14:36:21 +0100 Subject: [PATCH 10/83] Simplified size_insertList --- .../DHashMap/Internal/List/Associative.lean | 50 +++++++++---------- src/Std/Data/DHashMap/Internal/RawLemmas.lean | 20 +++----- 2 files changed, 33 insertions(+), 37 deletions(-) diff --git a/src/Std/Data/DHashMap/Internal/List/Associative.lean b/src/Std/Data/DHashMap/Internal/List/Associative.lean index 2aa666799fa5..2b955e517358 100644 --- a/src/Std/Data/DHashMap/Internal/List/Associative.lean +++ b/src/Std/Data/DHashMap/Internal/List/Associative.lean @@ -18,17 +18,17 @@ File contents: Verification of associative lists -- TODO: this does not belong here; maybe create a new datatype for such dependant associative lists? def List.get_by_key {α : Type u} {β : α -> Type v} [BEq α] [LawfulBEq α] (l: List ((a:α) × (β a))) (k: α) (h : (l.map Sigma.fst).contains k) : β k := - match l with + match l with | .nil => by simp at h - | .cons ⟨k', v⟩ l => if h' : k == k' - then + | .cons ⟨k', v⟩ l => if h' : k == k' + then have : k = k' := LawfulBEq.eq_of_beq h' - cast (by rw [this]) v + cast (by rw [this]) v else List.get_by_key l k (by unfold List.map at h; rw [List.contains_cons] at h; rw [Bool.or_eq_true_iff] at h; cases h with | inl hl => simp at hl; rw [hl] at h'; simp at h' | inr _ => assumption) -theorem List.get_by_key_append_elem {α : Type u} {β : α -> Type v} [BEq α] [LawfulBEq α] (l: List ((a:α) × (β a))) (elem: ((a:α) × (β a))) (k: α) (h : (l.map Sigma.fst).contains k) : ∀ h', (l ++ [elem]).get_by_key k h' = l.get_by_key k h := by +theorem List.get_by_key_append_elem {α : Type u} {β : α -> Type v} [BEq α] [LawfulBEq α] (l: List ((a:α) × (β a))) (elem: ((a:α) × (β a))) (k: α) (h : (l.map Sigma.fst).contains k) : ∀ h', (l ++ [elem]).get_by_key k h' = l.get_by_key k h := by intro h' - induction l with + induction l with | nil => simp at h | cons elem' l ih => simp [get_by_key] @@ -36,21 +36,21 @@ theorem List.get_by_key_append_elem {α : Type u} {β : α -> Type v} [BEq α] [ . rfl . apply ih -theorem List.get_by_key_append_elem' {α : Type u} {β : α -> Type v} [BEq α] [LawfulBEq α] (l: List ((a:α) × (β a))) (elem: ((a:α) × (β a))) (k: α) (h : ¬(l.map Sigma.fst).contains k) : ∀ h', (l ++ [elem]).get_by_key k h' = cast (by simp at h'; simp at h; cases h' with - | inl h' => cases h' with | intro a h' => specialize h a h'.left h'.right; simp at h - | inr h' => rw [h']) elem.snd := -by +theorem List.get_by_key_append_elem' {α : Type u} {β : α -> Type v} [BEq α] [LawfulBEq α] (l: List ((a:α) × (β a))) (elem: ((a:α) × (β a))) (k: α) (h : ¬(l.map Sigma.fst).contains k) : ∀ h', (l ++ [elem]).get_by_key k h' = cast (by simp at h'; simp at h; cases h' with + | inl h' => cases h' with | intro a h' => specialize h a h'.left h'.right; simp at h + | inr h' => rw [h']) elem.snd := +by intro h' - induction l with + induction l with | nil => simp at h'; simp [get_by_key, h'] | cons elem' l ih => simp [get_by_key] simp at h simp at h' - cases h' with + cases h' with | inl h' => rw [h'] at h; simp at h - | inr h' => - cases h' with + | inr h' => + cases h' with | inl h' => cases h' with | intro a h' => have hr := h.right; specialize hr a h'.left h'.right; simp at hr | inr h' => simp [h.left] @@ -1929,36 +1929,36 @@ theorem containsKey_insertList [BEq α] [PartialEquivBEq α] (l toInsert : List theorem containsKey_of_containsKey_insertList [BEq α] [PartialEquivBEq α] {l toInsert : List ((a : α) × β a)} {k : α} (h₁ : containsKey k (insertList l toInsert)) (h₂ : (toInsert.map Sigma.fst).contains k = false) : containsKey k l := by - rw [containsKey_insertList, h₂] at h₁; simp at h₁; exact h₁ + rw [containsKey_insertList, h₂] at h₁; simp at h₁; exact h₁ -theorem getValueCast_insertList [BEq α] [LawfulBEq α] +theorem getValueCast_insertList [BEq α] [LawfulBEq α] {l toInsert : List ((a : α) × β a)} {k : α} {h} : getValueCast k (insertList l toInsert) h = if h' : (toInsert.map Sigma.fst).contains k then List.get_by_key toInsert.reverse k (by simp; simp at h'; exact h') else getValueCast k l (containsKey_of_containsKey_insertList h (Bool.eq_false_iff.2 h')) := by - induction toInsert generalizing l with + induction toInsert generalizing l with | nil => simp [insertList] - | cons elem toInsert ih => + | cons elem toInsert ih => simp [insertList] rw [ih] split - case isTrue eq => + case isTrue eq => split case isTrue eq2 => rw [List.get_by_key_append_elem] case isFalse eq2 => rw [List.contains_cons] at eq2; rw [eq] at eq2; simp at eq2 - case isFalse eq => + case isFalse eq => rw [getValueCast_insertEntry] split - case isTrue eq2 => - split + case isTrue eq2 => + split case isTrue eq3 => rw [List.get_by_key_append_elem']; simp; simp at eq; exact eq case isFalse eq3 => simp at eq3; simp at eq2; rw [eq2] at eq3; simp at eq3 - case isFalse eq2 => - split + case isFalse eq2 => + split case isTrue eq3 => simp at eq3; simp at eq; cases eq3 with | inl hl => rw [hl] at eq2; simp at eq2 | inr hr => cases hr with | intro a hr => specialize eq a hr.left hr.right; simp at eq case isFalse eq3 => rfl -theorem insertList_perm [BEq α] [ReflBEq α] [PartialEquivBEq α] (l toInsert: List ((a : α) × β a)) (distinct_l: DistinctKeys l) (distinct_toInsert: DistinctKeys toInsert) (distinct_both: ∀ (a:α), ¬ containsKey a l ∨ ¬ containsKey a toInsert): +theorem insertList_perm [BEq α] [ReflBEq α] [PartialEquivBEq α] (l toInsert: List ((a : α) × β a)) (distinct_l: DistinctKeys l) (distinct_toInsert: DistinctKeys toInsert) (distinct_both: ∀ (a:α), ¬ (containsKey a l ∧ containsKey a toInsert)): Perm (insertList l toInsert) (l++toInsert) := by induction toInsert generalizing l with | nil => simp[insertList] diff --git a/src/Std/Data/DHashMap/Internal/RawLemmas.lean b/src/Std/Data/DHashMap/Internal/RawLemmas.lean index d626af79bd9c..1f81b1e94ae9 100644 --- a/src/Std/Data/DHashMap/Internal/RawLemmas.lean +++ b/src/Std/Data/DHashMap/Internal/RawLemmas.lean @@ -863,11 +863,12 @@ theorem insertMany_eq_insertList rw [insertList_eq_foldl, insertMany_val] @[simp] -theorem insertList_nil : m.insertList [] = m := by +theorem insertList_nil: m.insertList [] = m := by simp[insertList, Id.run] @[simp] -theorem insertList_singleton {k: α} {v: β k}: m.insertList [⟨k,v⟩] = m.insert k v := by +theorem insertList_singleton {k: α} {v: β k} (h: m.1.WF): m.insertList [⟨k,v⟩] = m.insert k v := by + simp_to_model simp[insertList, Id.run] @[simp] @@ -908,22 +909,17 @@ theorem get_insertList [LawfulBEq α] [LawfulHashable α] (h : m.1.WF) {l: List else get m k (contains_of_contains_insertList _ h h₁ (Bool.eq_false_iff.2 h₂)) := by simp_to_model using List.getValueCast_insertList --- try to revert extra conditions to use simp_to_model_on_them -theorem size_insertList [EquivBEq α] [LawfulHashable α] {l: List ((a:α) × (β a))} {distinct: DistinctKeys l} {distinct2: ∀ (a:α), ¬ (m.contains a = true ∧ List.containsKey a l = true)} (h: m.1.WF): (m.insertList l).1.size = m.1.size + l.length := by +theorem size_insertList [EquivBEq α] [LawfulHashable α] {l: List ((a:α) × (β a))} {distinct: DistinctKeys l} (h: m.1.WF): (∀ (a:α), ¬ (m.contains a = true ∧ List.containsKey a l = true)) → (m.insertList l).1.size = m.1.size + l.length := by simp_to_model rw [← List.length_append] + intro distinct' apply List.Perm.length_eq apply List.insertList_perm . apply (Raw.WF.out h).distinct . exact distinct - . simp at distinct2 - intro a - cases eq : containsKey a (toListModel m.val.buckets) with - | false => simp - | true => - simp - apply distinct2 - simp_to_model + . apply distinct' + + end Raw₀ end Std.DHashMap.Internal From ceb23c99fe5a609b19a9259232bf11b568b8e686 Mon Sep 17 00:00:00 2001 From: jt0202 Date: Sun, 17 Nov 2024 11:58:05 +0100 Subject: [PATCH 11/83] Added results for isEmpty --- .../DHashMap/Internal/List/Associative.lean | 19 +++++++++++++++++++ src/Std/Data/DHashMap/Internal/RawLemmas.lean | 8 ++++++-- 2 files changed, 25 insertions(+), 2 deletions(-) diff --git a/src/Std/Data/DHashMap/Internal/List/Associative.lean b/src/Std/Data/DHashMap/Internal/List/Associative.lean index 2b955e517358..5b71a014d2da 100644 --- a/src/Std/Data/DHashMap/Internal/List/Associative.lean +++ b/src/Std/Data/DHashMap/Internal/List/Associative.lean @@ -1999,4 +1999,23 @@ theorem insertList_perm [BEq α] [ReflBEq α] [PartialEquivBEq α] (l toInsert: apply Perm.symm apply List.perm_middle +theorem insertList_not_isEmpty_if_start_not_isEmpty [BEq α]{l toInsert: List ((a : α) × β a)} {h:l.isEmpty = false}: + (List.insertList l toInsert).isEmpty = false := by + induction toInsert generalizing l with + | nil => simp[insertList, h] + | cons hd tl ih => + simp[insertList, ih] + + +theorem insertList_isEmpty [BEq α]{l toInsert: List ((a : α) × β a)}: (List.insertList l toInsert).isEmpty ↔ l.isEmpty ∧ toInsert.isEmpty := by + induction toInsert with + | nil => simp[insertList] + | cons hd tl ih => + simp only [insertList, List.isEmpty_cons, Bool.false_eq_true, and_false, + iff_false] + apply ne_true_of_eq_false + apply insertList_not_isEmpty_if_start_not_isEmpty + simp + + end List diff --git a/src/Std/Data/DHashMap/Internal/RawLemmas.lean b/src/Std/Data/DHashMap/Internal/RawLemmas.lean index 1f81b1e94ae9..b62f9184e6fe 100644 --- a/src/Std/Data/DHashMap/Internal/RawLemmas.lean +++ b/src/Std/Data/DHashMap/Internal/RawLemmas.lean @@ -858,6 +858,7 @@ theorem insertMany_val (m : Raw₀ α β) (l : List ((a : α) × β a)) : simp rw [ih] +@[simp] theorem insertMany_eq_insertList (m : Raw₀ α β) (l : List ((a : α) × β a)) : (insertMany m l).val = insertList m l := by rw [insertList_eq_foldl, insertMany_val] @@ -867,8 +868,7 @@ theorem insertList_nil: m.insertList [] = m := by simp[insertList, Id.run] @[simp] -theorem insertList_singleton {k: α} {v: β k} (h: m.1.WF): m.insertList [⟨k,v⟩] = m.insert k v := by - simp_to_model +theorem insertList_singleton [EquivBEq α] [LawfulHashable α] {k: α} {v: β k}: m.insertList [⟨k,v⟩] = m.insert k v := by simp[insertList, Id.run] @[simp] @@ -919,7 +919,11 @@ theorem size_insertList [EquivBEq α] [LawfulHashable α] {l: List ((a:α) × ( . exact distinct . apply distinct' +theorem insertList_notEmpty_if_m_notEmpty [EquivBEq α] [LawfulHashable α] {l: List ((a:α) × (β a))}(h: m.1.WF): (m.1.isEmpty = false) → (m.insertList l).1.isEmpty = false := by + simp_to_model using List.insertList_not_isEmpty_if_start_not_isEmpty +theorem insertList_isEmpty [EquivBEq α] [LawfulHashable α] {l: List ((a:α) × (β a))}(h: m.1.WF): (m.insertList l).1.isEmpty ↔ m.1.isEmpty ∧ l.isEmpty := by + simp_to_model using List.insertList_isEmpty end Raw₀ end Std.DHashMap.Internal From a9b48470806b44a569c0b894f153126d8f3a45a1 Mon Sep 17 00:00:00 2001 From: jt0202 Date: Sun, 17 Nov 2024 23:50:10 +0100 Subject: [PATCH 12/83] Start get? verification --- .../DHashMap/Internal/List/Associative.lean | 151 ++++++++++++++++++ src/Std/Data/DHashMap/Internal/RawLemmas.lean | 16 ++ 2 files changed, 167 insertions(+) diff --git a/src/Std/Data/DHashMap/Internal/List/Associative.lean b/src/Std/Data/DHashMap/Internal/List/Associative.lean index 5b71a014d2da..88751b3a6c47 100644 --- a/src/Std/Data/DHashMap/Internal/List/Associative.lean +++ b/src/Std/Data/DHashMap/Internal/List/Associative.lean @@ -205,6 +205,34 @@ theorem getValueCast?_cons_self [BEq α] [LawfulBEq α] {l : List ((a : α) × getValueCast? k (⟨k, v⟩ :: l) = some v := by rw [getValueCast?_cons_of_true BEq.refl, cast_eq] +theorem getValueCast?_mem [BEq α] [LawfulBEq α] {l : List ((a : α) × β a)} {k k': α} {v : β k} {k_eq: k == k'} {mem: ⟨k,v⟩ ∈ l} {distinct: List.Pairwise (fun a b => ! a.1 == b.1) l}: getValueCast? k' l = some (cast (by congr; apply LawfulBEq.eq_of_beq k_eq) v) := by + induction l with + | nil => simp at mem + | cons hd tl ih => + simp[getValueCast?] + simp at mem + cases mem with + | inl mem => + rw [← mem] + simp [k_eq] + | inr mem => + have hd_k': (hd.fst == k') = false := by + simp at distinct + simp + rcases distinct with ⟨distinct, _⟩ + specialize distinct ⟨k,v⟩ mem + simp at distinct + simp at k_eq + rw [k_eq] at distinct + apply distinct + simp [hd_k'] + apply ih + · exact mem + · exact List.Pairwise.of_cons distinct + + + + theorem getValue?_eq_getValueCast? [BEq α] [LawfulBEq α] {β : Type v} {l : List ((_ : α) × β)} {a : α} : getValue? a l = getValueCast? a l := by induction l using assoc_induction <;> simp_all [getValueCast?_cons, getValue?_cons] @@ -721,6 +749,62 @@ theorem replaceEntry_of_containsKey_eq_false [BEq α] {l : List ((a : α) × β rw [containsKey_cons_eq_false] at h rw [replaceEntry_cons_of_false h.1, ih h.2] +theorem replaceEntry_mem_of_mem [BEq α] [LawfulBEq α] {l : List ((a : α) × β a)} {k: α} {v : β k} {ele: (a: α) × β a} (mem: containsKey k l) {distinct: List.Pairwise (fun a b => ! a.1 == b.1) l}: ele ∈ replaceEntry k v l ↔ (ele ∈ l ∧ ¬ (ele.fst == k)) ∨ (ele = ⟨k,v⟩) := by + induction l with + | nil => simp at mem + | cons hd tl ih => + simp [replaceEntry, cond_eq_if] + split + · simp + rename_i hd_k + simp at hd_k + admit + · simp + rw [ih] + admit + admit + admit + + +theorem replaceEntry_distinct_keys [BEq α] [LawfulBEq α] {l : List ((a : α) × β a)} {k : α} {v : β k}: List.Pairwise (fun a b => ! a.1 == b.1) l → List.Pairwise (fun a b => ! a.1 == b.1) (replaceEntry k v l) := by + intro h + induction l with + | nil => simp + | cons hd tl ih => + simp[replaceEntry, cond_eq_if] + simp at h + rcases h with ⟨hd_pw, tl_pw⟩ + split + · simp + constructor + · rename_i hd_k + intro a a_tl + simp at hd_k + rw [← hd_k] + apply hd_pw a a_tl + · exact tl_pw + · rename_i hd_k + simp at hd_k + simp + constructor + · intro a a_mem + by_cases contains_k: containsKey k tl + · rw [replaceEntry_mem_of_mem contains_k] at a_mem + · cases a_mem with + | inl a_tl => + apply hd_pw a (And.left a_tl) + | inr a_k => + simp[a_k, hd_k] + · simp + exact tl_pw + · simp at contains_k + rw [replaceEntry_of_containsKey_eq_false contains_k] at a_mem + apply hd_pw a a_mem + · simp at ih + apply ih + exact tl_pw + + @[simp] theorem isEmpty_replaceEntry [BEq α] {l : List ((a : α) × β a)} {k : α} {v : β k} : (replaceEntry k v l).isEmpty = l.isEmpty := by @@ -1958,6 +2042,73 @@ theorem getValueCast_insertList [BEq α] [LawfulBEq α] case isTrue eq3 => simp at eq3; simp at eq; cases eq3 with | inl hl => rw [hl] at eq2; simp at eq2 | inr hr => cases hr with | intro a hr => specialize eq a hr.left hr.right; simp at eq case isFalse eq3 => rfl +theorem getValueCast?_insertList_start_mem [BEq α] [LawfulBEq α] (l toInsert: List ((a : α) × β a)) {k k': α} {v: β k} {distinct1: List.Pairwise (fun a b => ! a.1 == b.1) toInsert} {distinct2: List.Pairwise (fun a b => ! a.1 == b.1) l} {k_eq: k == k'} {mem: ⟨k,v⟩ ∈ l}: (∀ (a:α), ¬ (containsKey a l = true ∧ containsKey a toInsert = true)) → getValueCast? k' (insertList l toInsert) = some (cast (by congr; apply LawfulBEq.eq_of_beq k_eq) v) := by + induction toInsert generalizing l with + | nil => + simp[insertList] + apply getValueCast?_mem + · exact k_eq + · exact mem + · exact distinct2 + | cons hd tl ih => + intro h + simp [insertList] + apply ih + · exact List.Pairwise.of_cons distinct1 + · simp only[insertEntry, cond_eq_if] + split + · apply replaceEntry_distinct_keys distinct2 + · simp + constructor + · admit + · admit + · admit + · admit + +theorem getValueCast?_insertList_toInsert_mem [BEq α] [LawfulBEq α] (l toInsert: List ((a : α) × β a)) {k k': α} {k_eq: k == k'} {v: β k} {distinct: List.Pairwise (fun a b => ! a.1 == b.1) toInsert} {distinct2: DistinctKeys l} {mem: ⟨k,v⟩ ∈ toInsert}: (∀ (a:α), ¬ (containsKey a l = true ∧ containsKey a toInsert = true)) → getValueCast? k' (insertList l toInsert) = some (cast (by congr;apply LawfulBEq.eq_of_beq k_eq) v) := by + induction toInsert generalizing l with + | nil => simp at mem + | cons hd tl ih => + intro h + simp[insertList, insertEntry, cond_eq_if] + have contains_hd: ¬ (containsKey hd.fst l = true) := by + specialize h hd.fst + simp[containsKey] at h + simp[h] + simp[contains_hd] + by_cases hd_k': (hd.fst ==k') = true + · apply getValueCast?_insertList_start_mem + · exact List.Pairwise.of_cons distinct + · admit + · exact k_eq + · simp + admit + · admit + · apply ih + · exact List.Pairwise.of_cons distinct + · simp at mem + cases mem with + | inl mem => + simp at hd_k' + exfalso + apply hd_k' + rw [← mem] + simp at k_eq + simp [k_eq] + | inr mem => + simp + admit + · simp at mem + cases mem with + | inl mem => + simp [← mem] at hd_k' + simp at k_eq + contradiction + | inr mem => + exact mem + · admit + + theorem insertList_perm [BEq α] [ReflBEq α] [PartialEquivBEq α] (l toInsert: List ((a : α) × β a)) (distinct_l: DistinctKeys l) (distinct_toInsert: DistinctKeys toInsert) (distinct_both: ∀ (a:α), ¬ (containsKey a l ∧ containsKey a toInsert)): Perm (insertList l toInsert) (l++toInsert) := by induction toInsert generalizing l with diff --git a/src/Std/Data/DHashMap/Internal/RawLemmas.lean b/src/Std/Data/DHashMap/Internal/RawLemmas.lean index b62f9184e6fe..3547e8608d7f 100644 --- a/src/Std/Data/DHashMap/Internal/RawLemmas.lean +++ b/src/Std/Data/DHashMap/Internal/RawLemmas.lean @@ -899,10 +899,26 @@ theorem contains_insertList [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l: L (m.insertList l).contains k ↔ m.contains k ∨ (l.map Sigma.fst).contains k := by simp_to_model using List.containsKey_insertList +theorem contains_insertList_of_contains_list [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l: List ((a:α) × (β a))} {k: α} {v: β k}: ⟨k,v⟩ ∈ l → (m.insertList l).contains k := by + sorry + theorem contains_of_contains_insertList [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l: List ((a:α) × (β a))} {k: α} : (m.insertList l).contains k → (l.map Sigma.fst).contains k = false → m.contains k := by simp_to_model using List.containsKey_of_containsKey_insertList +theorem get?_insertList_list [LawfulBEq α] [LawfulHashable α] {l: List ((a:α) × (β a))} {k k': α} {v: β k} {distinct: List.Pairwise (fun a b => ! a.1 == b.1) l} {k_eq: k == k'} {mem: ⟨k,v⟩ ∈ l} (h: m.1.WF): (∀ (a:α), ¬ (m.contains a = true ∧ List.containsKey a l = true)) → (m.insertList l).get? k' = some (cast (by congr; apply LawfulBEq.eq_of_beq k_eq) v) := by + simp_to_model using getValueCast?_insertList_toInsert_mem + + + +theorem get!_insertList_list [LawfulBEq α] [LawfulHashable α] {l: List ((a:α) × (β a))} {k k': α} {v: β k}[Inhabited (β k)] {distinct: List.Pairwise (fun a b => ! a.1 == b.1) l} {k_eq: k == k'} {mem: ⟨k,v⟩ ∈ l} (h: m.1.WF): (∀ (a:α), ¬ (m.contains a = true ∧ List.containsKey a l = true)) → (m.insertList l).get! k = v := by + simp_to_model + sorry + +theorem get_insertList_list [LawfulBEq α] [LawfulHashable α] {l: List ((a:α) × (β a))} {k k': α} {v: β k}[Inhabited (β k)] {distinct: List.Pairwise (fun a b => ! a.1 == b.1) l} {k_eq: k == k'} {mem: ⟨k,v⟩ ∈ l} (h: m.1.WF): (∀ (a:α), ¬ (m.contains a = true ∧ List.containsKey a l = true)) → (m.insertList l).get k (contains_insertList_of_contains_list _ h mem)= v := by + simp_to_model + sorry + theorem get_insertList [LawfulBEq α] [LawfulHashable α] (h : m.1.WF) {l: List ((a:α) × (β a))} {k: α} {h₁} : get (m.insertList l) k h₁ = if h₂ : (l.map Sigma.fst).contains k then List.get_by_key l.reverse k (by simp; simp at h₂; exact h₂) From da76a1debba60de320f87e2c472131c4b8348e52 Mon Sep 17 00:00:00 2001 From: jt0202 Date: Mon, 18 Nov 2024 22:52:17 +0100 Subject: [PATCH 13/83] WIP: verification getValueCast? --- .../DHashMap/Internal/List/Associative.lean | 366 ++++++++++++------ src/Std/Data/DHashMap/Internal/RawLemmas.lean | 26 +- 2 files changed, 270 insertions(+), 122 deletions(-) diff --git a/src/Std/Data/DHashMap/Internal/List/Associative.lean b/src/Std/Data/DHashMap/Internal/List/Associative.lean index 88751b3a6c47..e885de35b5b8 100644 --- a/src/Std/Data/DHashMap/Internal/List/Associative.lean +++ b/src/Std/Data/DHashMap/Internal/List/Associative.lean @@ -16,47 +16,6 @@ the contents of this file. File contents: Verification of associative lists -/ --- TODO: this does not belong here; maybe create a new datatype for such dependant associative lists? -def List.get_by_key {α : Type u} {β : α -> Type v} [BEq α] [LawfulBEq α] (l: List ((a:α) × (β a))) (k: α) (h : (l.map Sigma.fst).contains k) : β k := - match l with - | .nil => by simp at h - | .cons ⟨k', v⟩ l => if h' : k == k' - then - have : k = k' := LawfulBEq.eq_of_beq h' - cast (by rw [this]) v - else List.get_by_key l k (by unfold List.map at h; rw [List.contains_cons] at h; rw [Bool.or_eq_true_iff] at h; cases h with | inl hl => simp at hl; rw [hl] at h'; simp at h' | inr _ => assumption) - -theorem List.get_by_key_append_elem {α : Type u} {β : α -> Type v} [BEq α] [LawfulBEq α] (l: List ((a:α) × (β a))) (elem: ((a:α) × (β a))) (k: α) (h : (l.map Sigma.fst).contains k) : ∀ h', (l ++ [elem]).get_by_key k h' = l.get_by_key k h := by - intro h' - induction l with - | nil => simp at h - | cons elem' l ih => - simp [get_by_key] - split - . rfl - . apply ih - -theorem List.get_by_key_append_elem' {α : Type u} {β : α -> Type v} [BEq α] [LawfulBEq α] (l: List ((a:α) × (β a))) (elem: ((a:α) × (β a))) (k: α) (h : ¬(l.map Sigma.fst).contains k) : ∀ h', (l ++ [elem]).get_by_key k h' = cast (by simp at h'; simp at h; cases h' with - | inl h' => cases h' with | intro a h' => specialize h a h'.left h'.right; simp at h - | inr h' => rw [h']) elem.snd := -by - intro h' - induction l with - | nil => simp at h'; simp [get_by_key, h'] - | cons elem' l ih => - simp [get_by_key] - simp at h - simp at h' - cases h' with - | inl h' => rw [h'] at h; simp at h - | inr h' => - cases h' with - | inl h' => cases h' with | intro a h' => have hr := h.right; specialize hr a h'.left h'.right; simp at hr - | inr h' => - simp [h.left] - apply ih - simp - apply h.right set_option linter.missingDocs true set_option autoImplicit false @@ -294,6 +253,20 @@ def containsKey [BEq α] (a : α) : List ((a : α) × β a) → Bool @[simp] theorem containsKey_cons [BEq α] {l : List ((a : α) × β a)} {k a : α} {v : β k} : containsKey a (⟨k, v⟩ :: l) = (k == a || containsKey a l) := rfl +theorem containsKey_eq_false_iff [BEq α][PartialEquivBEq α] {l : List ((a : α) × β a)} {a : α}: containsKey a l = false ↔ ∀ (b:((a : α) × β a)), b ∈ l → (a == b.fst) = false := by + induction l with + | nil => simp + | cons hd tl ih => + simp[containsKey, ih] + intro _ + rw [Bool.eq_iff_iff] + constructor + · intro h + apply PartialEquivBEq.symm h + · intro h + apply PartialEquivBEq.symm h + + theorem containsKey_cons_eq_false [BEq α] {l : List ((a : α) × β a)} {k a : α} {v : β k} : (containsKey a (⟨k, v⟩ :: l) = false) ↔ ((k == a) = false) ∧ (containsKey a l = false) := by simp [containsKey_cons, not_or] @@ -758,12 +731,80 @@ theorem replaceEntry_mem_of_mem [BEq α] [LawfulBEq α] {l : List ((a : α) × · simp rename_i hd_k simp at hd_k - admit + constructor + · intro h + cases h with + | inl h => right; exact h + | inr h => + left + constructor + · right + exact h + · rw [← hd_k] + simp at distinct + apply Ne.symm + apply (And.left distinct ele h) + · intro h + cases h with + | inl h => + rcases h with ⟨ele_mem, ele_k⟩ + cases ele_mem with + | inl h => + rw [h] at ele_k + contradiction + | inr h => + right + apply h + | inr h => + left + apply h · simp + rename_i hd_k + simp at hd_k rw [ih] - admit - admit - admit + · constructor + · intro h + cases h with + | inl h => + left + constructor + · left + simp[h] + · simp[h, hd_k] + | inr h => + cases h with + | inl h => + simp at h + left + constructor + · right + apply And.left h + · apply And.right h + | inr h => + right + exact h + · intro h + cases h with + | inl h => + rcases h with ⟨h,ele_k⟩ + cases h with + | inl h => + left + simp[h] + | inr h => + right + left + simp[h, ele_k] + | inr h => + right + right + exact h + · rw [containsKey_cons] at mem + simp[hd_k] at mem + exact mem + · simp at distinct + simp + apply And.right distinct theorem replaceEntry_distinct_keys [BEq α] [LawfulBEq α] {l : List ((a : α) × β a)} {k : α} {v : β k}: List.Pairwise (fun a b => ! a.1 == b.1) l → List.Pairwise (fun a b => ! a.1 == b.1) (replaceEntry k v l) := by @@ -1046,6 +1087,7 @@ theorem DistinctKeys.of_keys_eq [BEq α] {l : List ((a : α) × β a)} {l' : Lis (h : keys l = keys l') : DistinctKeys l → DistinctKeys l' := distinctKeys_of_sublist_keys (h ▸ Sublist.refl _) + theorem containsKey_iff_exists [BEq α] [PartialEquivBEq α] {l : List ((a : α) × β a)} {a : α} : containsKey a l ↔ ∃ a' ∈ keys l, a == a' := by rw [containsKey_eq_keys_contains, List.contains_iff_exists_mem_beq] @@ -1132,6 +1174,61 @@ theorem insertEntry_of_containsKey_eq_false_perm [BEq α] {l : List ((a : α) × | nil => simp | cons hd tl => simp[insertEntry, h] +theorem insertEntry_mem_of_different_keys [BEq α] [LawfulBEq α] {l : List ((a : α) × β a)} {k : α} {v : β k} {a: ((a : α) × β a)} {h: ¬ ⟨k,v⟩ = a} {distinct: List.Pairwise (fun a b => ! (a.1 == b.1) ) l}: ⟨k, v⟩ ∈ insertEntry a.fst a.snd l ↔ ⟨k,v⟩ ∈ l ∧ (k == a.fst) = false := by + induction l with + | nil => + simp + apply h + | cons hd tl ih => + unfold insertEntry + simp only [cond_eq_if] + split + · rename_i contains_a + rw [replaceEntry_mem_of_mem contains_a (distinct:=distinct)] + simp + intro h + contradiction + · simp + constructor + · intro h' + cases h' with + | inl h' => contradiction + | inr h' => + rename_i contains_a + simp at contains_a + rw [containsKey_eq_false_iff] at contains_a + cases h' with + | inl h' => + constructor + · left + exact h' + · simp at contains_a + apply Ne.symm + rw [← h'] at contains_a + simp at contains_a + apply And.left contains_a + | inr h' => + constructor + · right + exact h' + · simp at contains_a + apply Ne.symm + rcases contains_a with ⟨_, contains_a⟩ + specialize contains_a ⟨k,v⟩ + simp at contains_a + apply contains_a h' + · intro h' + rcases h' with ⟨h', _⟩ + cases h' with + | inl h' => + right + left + apply h' + | inr h' => + right + right + apply h' + theorem DistinctKeys.insertEntry [BEq α] [PartialEquivBEq α] {l : List ((a : α) × β a)} {k : α} {v : β k} (h : DistinctKeys l) : DistinctKeys (insertEntry k v l) := by @@ -2015,34 +2112,9 @@ theorem containsKey_of_containsKey_insertList [BEq α] [PartialEquivBEq α] (h₂ : (toInsert.map Sigma.fst).contains k = false) : containsKey k l := by rw [containsKey_insertList, h₂] at h₁; simp at h₁; exact h₁ -theorem getValueCast_insertList [BEq α] [LawfulBEq α] - {l toInsert : List ((a : α) × β a)} {k : α} - {h} : getValueCast k (insertList l toInsert) h = - if h' : (toInsert.map Sigma.fst).contains k then List.get_by_key toInsert.reverse k (by simp; simp at h'; exact h') - else getValueCast k l (containsKey_of_containsKey_insertList h (Bool.eq_false_iff.2 h')) := by - induction toInsert generalizing l with - | nil => simp [insertList] - | cons elem toInsert ih => - simp [insertList] - rw [ih] - split - case isTrue eq => - split - case isTrue eq2 => rw [List.get_by_key_append_elem] - case isFalse eq2 => rw [List.contains_cons] at eq2; rw [eq] at eq2; simp at eq2 - case isFalse eq => - rw [getValueCast_insertEntry] - split - case isTrue eq2 => - split - case isTrue eq3 => rw [List.get_by_key_append_elem']; simp; simp at eq; exact eq - case isFalse eq3 => simp at eq3; simp at eq2; rw [eq2] at eq3; simp at eq3 - case isFalse eq2 => - split - case isTrue eq3 => simp at eq3; simp at eq; cases eq3 with | inl hl => rw [hl] at eq2; simp at eq2 | inr hr => cases hr with | intro a hr => specialize eq a hr.left hr.right; simp at eq - case isFalse eq3 => rfl - -theorem getValueCast?_insertList_start_mem [BEq α] [LawfulBEq α] (l toInsert: List ((a : α) × β a)) {k k': α} {v: β k} {distinct1: List.Pairwise (fun a b => ! a.1 == b.1) toInsert} {distinct2: List.Pairwise (fun a b => ! a.1 == b.1) l} {k_eq: k == k'} {mem: ⟨k,v⟩ ∈ l}: (∀ (a:α), ¬ (containsKey a l = true ∧ containsKey a toInsert = true)) → getValueCast? k' (insertList l toInsert) = some (cast (by congr; apply LawfulBEq.eq_of_beq k_eq) v) := by + + +theorem getValueCast?_insertList_start_mem [BEq α] [LawfulBEq α] (l toInsert: List ((a : α) × β a)) {k k': α} {v: β k} {distinct1: List.Pairwise (fun a b => ! a.1 == b.1) toInsert} {distinct2: List.Pairwise (fun a b => ! a.1 == b.1) l} {k_eq: k == k'} {mem: ⟨k,v⟩ ∈ l} {mem': containsKey k toInsert = false}: getValueCast? k' (insertList l toInsert) = some (cast (by congr; apply LawfulBEq.eq_of_beq k_eq) v) := by induction toInsert generalizing l with | nil => simp[insertList] @@ -2051,7 +2123,6 @@ theorem getValueCast?_insertList_start_mem [BEq α] [LawfulBEq α] (l toInsert: · exact mem · exact distinct2 | cons hd tl ih => - intro h simp [insertList] apply ih · exact List.Pairwise.of_cons distinct1 @@ -2060,53 +2131,122 @@ theorem getValueCast?_insertList_start_mem [BEq α] [LawfulBEq α] (l toInsert: · apply replaceEntry_distinct_keys distinct2 · simp constructor - · admit - · admit - · admit - · admit + · rename_i h + simp at h + rw [containsKey_eq_false_iff] at h + simp at h + apply h + · simp at distinct2 + apply distinct2 + · rw[containsKey_cons_eq_false] at mem' + rw [insertEntry_mem_of_different_keys] + · constructor + · apply mem + · simp + simp at mem' + apply Ne.symm + apply And.left mem' + · false_or_by_contra + rename_i h + rw [← h] at mem' + simp at mem' + · apply distinct2 + · simp[containsKey] at mem' + apply And.right mem' -theorem getValueCast?_insertList_toInsert_mem [BEq α] [LawfulBEq α] (l toInsert: List ((a : α) × β a)) {k k': α} {k_eq: k == k'} {v: β k} {distinct: List.Pairwise (fun a b => ! a.1 == b.1) toInsert} {distinct2: DistinctKeys l} {mem: ⟨k,v⟩ ∈ toInsert}: (∀ (a:α), ¬ (containsKey a l = true ∧ containsKey a toInsert = true)) → getValueCast? k' (insertList l toInsert) = some (cast (by congr;apply LawfulBEq.eq_of_beq k_eq) v) := by +theorem getValueCast?_insertList_toInsert_mem [BEq α] [LawfulBEq α] (l toInsert: List ((a : α) × β a)) {k k': α} {k_eq: k == k'} {v: β k} {distinct: List.Pairwise (fun a b => ! a.1 == b.1) toInsert} {distinct2: List.Pairwise (fun a b => ! a.1 == b.1) l} {mem: ⟨k,v⟩ ∈ toInsert}: getValueCast? k' (insertList l toInsert) = some (cast (by congr;apply LawfulBEq.eq_of_beq k_eq) v) := by induction toInsert generalizing l with | nil => simp at mem | cons hd tl ih => - intro h simp[insertList, insertEntry, cond_eq_if] - have contains_hd: ¬ (containsKey hd.fst l = true) := by - specialize h hd.fst - simp[containsKey] at h - simp[h] - simp[contains_hd] - by_cases hd_k': (hd.fst ==k') = true - · apply getValueCast?_insertList_start_mem - · exact List.Pairwise.of_cons distinct - · admit - · exact k_eq - · simp - admit - · admit - · apply ih - · exact List.Pairwise.of_cons distinct - · simp at mem - cases mem with - | inl mem => + split + · simp at mem + cases mem with + | inl mem => + apply getValueCast?_insertList_start_mem + · exact List.Pairwise.of_cons distinct + · apply replaceEntry_distinct_keys distinct2 + · exact k_eq + · rw[replaceEntry_mem_of_mem] + · right + rw [mem] + · rename_i h + exact h + · exact distinct2 + · have k_hd: k = hd.fst := by + simp[← mem] + rw [k_hd, containsKey_eq_false_iff] + simp + simp at distinct + apply And.left distinct + | inr mem => + apply ih + · simp + simp at distinct + apply And.right distinct + · apply replaceEntry_distinct_keys distinct2 + · exact mem + · rename_i contains_hd + by_cases hd_k': (hd.fst ==k') = true + · apply getValueCast?_insertList_start_mem + · exact List.Pairwise.of_cons distinct + · simp + simp at contains_hd + rw [containsKey_eq_false_iff] at contains_hd + simp at contains_hd + constructor + · apply contains_hd + · simp at distinct2 + apply distinct2 + · exact k_eq + · simp at mem + cases mem with + | inl mem => + simp [mem] + | inr mem => + simp at distinct + exfalso + apply And.left distinct ⟨k,v⟩ mem + simp + simp at hd_k' + simp at k_eq + simp[k_eq, hd_k'] + · simp at contains_hd simp at hd_k' - exfalso - apply hd_k' - rw [← mem] simp at k_eq - simp [k_eq] - | inr mem => + rw [k_eq, ← hd_k'] + rw[containsKey_eq_false_iff] + simp at distinct simp - admit - · simp at mem - cases mem with - | inl mem => - simp [← mem] at hd_k' - simp at k_eq - contradiction - | inr mem => - exact mem - · admit + apply And.left distinct + · apply ih + · exact List.Pairwise.of_cons distinct + · simp at mem + cases mem with + | inl mem => + simp at hd_k' + exfalso + apply hd_k' + rw [← mem] + simp at k_eq + simp [k_eq] + | inr mem => + simp + simp at contains_hd + rw [containsKey_eq_false_iff] at contains_hd + simp at contains_hd + constructor + · exact contains_hd + · simp at distinct2 + apply distinct2 + · simp at mem + cases mem with + | inl mem => + simp [← mem] at hd_k' + simp at k_eq + contradiction + | inr mem => + exact mem theorem insertList_perm [BEq α] [ReflBEq α] [PartialEquivBEq α] (l toInsert: List ((a : α) × β a)) (distinct_l: DistinctKeys l) (distinct_toInsert: DistinctKeys toInsert) (distinct_both: ∀ (a:α), ¬ (containsKey a l ∧ containsKey a toInsert)): diff --git a/src/Std/Data/DHashMap/Internal/RawLemmas.lean b/src/Std/Data/DHashMap/Internal/RawLemmas.lean index 3547e8608d7f..00cb3e9cc448 100644 --- a/src/Std/Data/DHashMap/Internal/RawLemmas.lean +++ b/src/Std/Data/DHashMap/Internal/RawLemmas.lean @@ -900,30 +900,38 @@ theorem contains_insertList [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l: L simp_to_model using List.containsKey_insertList theorem contains_insertList_of_contains_list [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l: List ((a:α) × (β a))} {k: α} {v: β k}: ⟨k,v⟩ ∈ l → (m.insertList l).contains k := by - sorry + rw [contains_insertList _ h] + rw [List.contains_iff_exists_mem_beq] + intro h + right + exists k + simp + exists ⟨k,v⟩ + theorem contains_of_contains_insertList [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l: List ((a:α) × (β a))} {k: α} : (m.insertList l).contains k → (l.map Sigma.fst).contains k = false → m.contains k := by simp_to_model using List.containsKey_of_containsKey_insertList -theorem get?_insertList_list [LawfulBEq α] [LawfulHashable α] {l: List ((a:α) × (β a))} {k k': α} {v: β k} {distinct: List.Pairwise (fun a b => ! a.1 == b.1) l} {k_eq: k == k'} {mem: ⟨k,v⟩ ∈ l} (h: m.1.WF): (∀ (a:α), ¬ (m.contains a = true ∧ List.containsKey a l = true)) → (m.insertList l).get? k' = some (cast (by congr; apply LawfulBEq.eq_of_beq k_eq) v) := by +theorem get?_insertList_mem_list [LawfulBEq α] [LawfulHashable α] {l: List ((a:α) × (β a))} {k k': α} {v: β k} {distinct: List.Pairwise (fun a b => ! a.1 == b.1) l} {k_eq: k == k'} {mem: ⟨k,v⟩ ∈ l} (h: m.1.WF): (m.insertList l).get? k' = some (cast (by congr; apply LawfulBEq.eq_of_beq k_eq) v) := by simp_to_model using getValueCast?_insertList_toInsert_mem +theorem get?_insertList_not_mem_list [LawfulBEq α] [LawfulHashable α] {l: List ((a:α) × (β a))} {k: α} {distinct: List.Pairwise (fun a b => ! a.1 == b.1) l} {mem: ¬ k ∈ (l.map (Sigma.fst))} (h: m.1.WF): (m.insertList l).get? k = m.get? k := by + simp_to_model + sorry -theorem get!_insertList_list [LawfulBEq α] [LawfulHashable α] {l: List ((a:α) × (β a))} {k k': α} {v: β k}[Inhabited (β k)] {distinct: List.Pairwise (fun a b => ! a.1 == b.1) l} {k_eq: k == k'} {mem: ⟨k,v⟩ ∈ l} (h: m.1.WF): (∀ (a:α), ¬ (m.contains a = true ∧ List.containsKey a l = true)) → (m.insertList l).get! k = v := by +theorem get!_insertList_mem_list [LawfulBEq α] [LawfulHashable α] {l: List ((a:α) × (β a))} {k k': α} {v: β k}[Inhabited (β k')] {distinct: List.Pairwise (fun a b => ! a.1 == b.1) l} {k_eq: k == k'} {mem: ⟨k,v⟩ ∈ l} (h: m.1.WF): (m.insertList l).get! k' = cast (by congr; apply LawfulBEq.eq_of_beq k_eq) v := by simp_to_model sorry -theorem get_insertList_list [LawfulBEq α] [LawfulHashable α] {l: List ((a:α) × (β a))} {k k': α} {v: β k}[Inhabited (β k)] {distinct: List.Pairwise (fun a b => ! a.1 == b.1) l} {k_eq: k == k'} {mem: ⟨k,v⟩ ∈ l} (h: m.1.WF): (∀ (a:α), ¬ (m.contains a = true ∧ List.containsKey a l = true)) → (m.insertList l).get k (contains_insertList_of_contains_list _ h mem)= v := by +theorem getD_insertList_mem_list [LawfulBEq α] [LawfulHashable α] {l: List ((a:α) × (β a))} {k k': α} {v: β k} {fallback: β k'} {distinct: List.Pairwise (fun a b => ! a.1 == b.1) l} {k_eq: k == k'} {mem: ⟨k,v⟩ ∈ l} (h: m.1.WF): (m.insertList l).getD k' fallback = cast (by congr; apply LawfulBEq.eq_of_beq k_eq) v := by simp_to_model sorry -theorem get_insertList [LawfulBEq α] [LawfulHashable α] (h : m.1.WF) {l: List ((a:α) × (β a))} {k: α} {h₁} : - get (m.insertList l) k h₁ = - if h₂ : (l.map Sigma.fst).contains k then List.get_by_key l.reverse k (by simp; simp at h₂; exact h₂) - else get m k (contains_of_contains_insertList _ h h₁ (Bool.eq_false_iff.2 h₂)) := by - simp_to_model using List.getValueCast_insertList +theorem get_insertList_mem_list [LawfulBEq α] [LawfulHashable α] {l: List ((a:α) × (β a))} {k k': α} {v: β k}[Inhabited (β k)] {distinct: List.Pairwise (fun a b => ! a.1 == b.1) l} {k_eq: k == k'} {mem: ⟨k,v⟩ ∈ l} (h: m.1.WF): (m.insertList l).get k (contains_insertList_of_contains_list _ h mem)= v := by + simp_to_model + sorry theorem size_insertList [EquivBEq α] [LawfulHashable α] {l: List ((a:α) × (β a))} {distinct: DistinctKeys l} (h: m.1.WF): (∀ (a:α), ¬ (m.contains a = true ∧ List.containsKey a l = true)) → (m.insertList l).1.size = m.1.size + l.length := by simp_to_model From 9013992d6c3b6bc2738d231998208bd09beb2a51 Mon Sep 17 00:00:00 2001 From: jt0202 Date: Tue, 19 Nov 2024 18:33:26 +0100 Subject: [PATCH 14/83] Finished DistinctKeys for get? --- .../DHashMap/Internal/List/Associative.lean | 19 ++++++++++++++++++- 1 file changed, 18 insertions(+), 1 deletion(-) diff --git a/src/Std/Data/DHashMap/Internal/List/Associative.lean b/src/Std/Data/DHashMap/Internal/List/Associative.lean index e885de35b5b8..6bb575829ea5 100644 --- a/src/Std/Data/DHashMap/Internal/List/Associative.lean +++ b/src/Std/Data/DHashMap/Internal/List/Associative.lean @@ -1061,6 +1061,22 @@ theorem containsKey_of_mem [BEq α] [ReflBEq α] {l : List ((a : α) × β a)} { theorem DistinctKeys.nil [BEq α] : DistinctKeys ([] : List ((a : α) × β a)) := ⟨by simp⟩ +theorem DistinctKeys.def [BEq α] {l : List ((a : α) × β a)}: DistinctKeys l ↔ List.Pairwise (fun a b => ! (a.1 == b.1)) l := by + have h: ∀ (l' : List ((a : α) × β a)), List.Pairwise (fun a b => (a==b)= false) (keys l') ↔ List.Pairwise (fun a b => !(a.1 == b.1)) l':= by + intro l' + induction l' with + | nil => simp + | cons hd tl ih=> + simp only [keys_eq_map, List.map_cons, pairwise_cons] + rw [← ih] + simp [keys_eq_map] + rw [← h] + constructor + · intro h + apply h.distinct + · intro h + exact ⟨h⟩ + open List theorem DistinctKeys.perm_keys [BEq α] [PartialEquivBEq α] {l l' : List ((a : α) × β a)} @@ -2154,7 +2170,8 @@ theorem getValueCast?_insertList_start_mem [BEq α] [LawfulBEq α] (l toInsert: · simp[containsKey] at mem' apply And.right mem' -theorem getValueCast?_insertList_toInsert_mem [BEq α] [LawfulBEq α] (l toInsert: List ((a : α) × β a)) {k k': α} {k_eq: k == k'} {v: β k} {distinct: List.Pairwise (fun a b => ! a.1 == b.1) toInsert} {distinct2: List.Pairwise (fun a b => ! a.1 == b.1) l} {mem: ⟨k,v⟩ ∈ toInsert}: getValueCast? k' (insertList l toInsert) = some (cast (by congr;apply LawfulBEq.eq_of_beq k_eq) v) := by +theorem getValueCast?_insertList_toInsert_mem [BEq α] [LawfulBEq α] (l toInsert: List ((a : α) × β a)) {k k': α} {k_eq: k == k'} {v: β k} {distinct: List.Pairwise (fun a b => ! a.1 == b.1) toInsert} {distinct2: DistinctKeys l} {mem: ⟨k,v⟩ ∈ toInsert}: getValueCast? k' (insertList l toInsert) = some (cast (by congr;apply LawfulBEq.eq_of_beq k_eq) v) := by + rw [DistinctKeys.def] at distinct2 induction toInsert generalizing l with | nil => simp at mem | cons hd tl ih => From 1d18281f483360bcb87d405419bc150eb76d7a48 Mon Sep 17 00:00:00 2001 From: jt0202 Date: Thu, 21 Nov 2024 21:08:34 +0100 Subject: [PATCH 15/83] Finish get verification --- .../DHashMap/Internal/List/Associative.lean | 95 +++++++++++++++++++ src/Std/Data/DHashMap/Internal/RawLemmas.lean | 50 ++++++---- 2 files changed, 126 insertions(+), 19 deletions(-) diff --git a/src/Std/Data/DHashMap/Internal/List/Associative.lean b/src/Std/Data/DHashMap/Internal/List/Associative.lean index 6bb575829ea5..7f3f3b3fcf41 100644 --- a/src/Std/Data/DHashMap/Internal/List/Associative.lean +++ b/src/Std/Data/DHashMap/Internal/List/Associative.lean @@ -2265,6 +2265,101 @@ theorem getValueCast?_insertList_toInsert_mem [BEq α] [LawfulBEq α] (l toInser | inr mem => exact mem +theorem getValueCast?_insertList_not_toInsert_mem [BEq α] [LawfulBEq α] (l toInsert: List ((a : α) × β a)) {k k': α} {k_eq: k == k'} {not_mem: ¬ k ∈ List.map Sigma.fst toInsert} {distinct: DistinctKeys l}: getValueCast? k' (insertList l toInsert) = getValueCast? k' l := by + simp at k_eq + rw [DistinctKeys.def] at distinct + induction toInsert generalizing l with + | nil => simp[insertList] + | cons hd tl ih => + simp [insertList] + rw [ih] + · rw [getValueCast?_insertEntry] + split + · rename_i h + simp at h + simp [k_eq, h] at not_mem + · rfl + · simp at not_mem + simp + apply And.right not_mem + · rw [← DistinctKeys.def] + apply DistinctKeys.insertEntry + rw [DistinctKeys.def] + exact distinct + +theorem getValueCast!_insertList_toInsert_mem [BEq α] [LawfulBEq α] (l toInsert: List ((a : α) × β a)) {k k': α} [Inhabited (β k')] {k_eq: k == k'} {v: β k} {distinct: List.Pairwise (fun a b => ! a.1 == b.1) toInsert} {distinct2: DistinctKeys l} {mem: ⟨k,v⟩ ∈ toInsert}: getValueCast! k' (insertList l toInsert) = (cast (by congr;apply LawfulBEq.eq_of_beq k_eq) v) := by + rw [getValueCast!_eq_getValueCast?, getValueCast?_insertList_toInsert_mem (mem:= mem) (l:=l) (toInsert:=toInsert)] + · simp + · exact k_eq + · exact distinct + · exact distinct2 + +theorem getValueCast!_insertList_not_toInsert_mem [BEq α] [LawfulBEq α] (l toInsert: List ((a : α) × β a)) {k k': α} [Inhabited (β k')] {k_eq: k == k'} {not_mem: ¬ k ∈ List.map Sigma.fst toInsert} {distinct: DistinctKeys l}: getValueCast! k' (insertList l toInsert) = getValueCast! k' l := by + rw [getValueCast!_eq_getValueCast?,getValueCast!_eq_getValueCast? ,getValueCast?_insertList_not_toInsert_mem (not_mem:=not_mem) (distinct:=distinct)] + exact k_eq + +theorem getValueCastD_insertList_toInsert_mem [BEq α] [LawfulBEq α] (l toInsert: List ((a : α) × β a)) {k k': α} {k_eq: k == k'} {v: β k} {fallback: β k'} {distinct: List.Pairwise (fun a b => ! a.1 == b.1) toInsert} {distinct2: DistinctKeys l} {mem: ⟨k,v⟩ ∈ toInsert}: getValueCastD k' (insertList l toInsert) fallback = (cast (by congr;apply LawfulBEq.eq_of_beq k_eq) v) := by + rw [getValueCastD_eq_getValueCast?, getValueCast?_insertList_toInsert_mem (mem:= mem) (l:=l) (toInsert:=toInsert)] + · simp + · exact k_eq + · exact distinct + · exact distinct2 + +theorem getValueCastD_insertList_not_toInsert_mem [BEq α] [LawfulBEq α] (l toInsert: List ((a : α) × β a)) {k k': α} {fallback: β k'} {k_eq: k == k'} {not_mem: ¬ k ∈ List.map Sigma.fst toInsert} {distinct: DistinctKeys l}: getValueCastD k' (insertList l toInsert) fallback = getValueCastD k' l fallback:= by + rw [getValueCastD_eq_getValueCast?, getValueCastD_eq_getValueCast?, getValueCast?_insertList_not_toInsert_mem (not_mem:=not_mem) (distinct:=distinct)] + exact k_eq + +theorem contains_insertList_of_mem [BEq α] [LawfulBEq α] (l toInsert: List ((a : α) × β a)) (k k': α) (v: β k) (k_eq: k == k') (mem: ⟨k,v⟩ ∈ toInsert): containsKey k' (insertList l toInsert) = true := by + rw [containsKey_insertList] + right + simp at * + exists ⟨k,v⟩ + +theorem contains_insertList_of_contains_first [BEq α] [LawfulBEq α] {l toInsert: List ((a : α) × β a)} {k k': α} (k_eq: k == k'): containsKey k l = true → containsKey k' (insertList l toInsert) = true:= by + intro h + rw [containsKey_insertList] + simp at * + simp [← k_eq, h] + +instance Inhabited_of_containsKey [BEq α] [LawfulBEq α] (l: List ((a : α) × β a)) (k k': α) (k_eq: k == k') (mem: containsKey k l = true): Inhabited (β k') := + ⟨ + match l with + | .nil => + (by simp at mem) + | .cons hd tl => + if h:hd.fst == k' + then cast (by congr; simp at h; apply h) hd.snd + else + have mem': containsKey k tl = true := by + simp[containsKey] at mem + cases mem with + | inl h' => simp at h; simp at k_eq; rw [k_eq] at h'; contradiction + | inr h' => exact h' + (Inhabited_of_containsKey tl k k' k_eq mem').default + ⟩ + + +theorem getValueCast_insertList_toInsert_mem [BEq α] [LawfulBEq α] (l toInsert: List ((a : α) × β a)) {k k': α} {v: β k} {k_eq: k == k'} {distinct: List.Pairwise (fun a b => ! a.1 == b.1) toInsert} {distinct2: DistinctKeys l} {mem: ⟨k,v⟩ ∈ toInsert}: getValueCast k' (insertList l toInsert) (contains_insertList_of_mem l toInsert k k' v k_eq mem) = cast (by congr; apply LawfulBEq.eq_of_beq k_eq) v := by + have inh: Inhabited (β k') := by + apply Inhabited_of_containsKey toInsert k k' k_eq + rw [containsKey_eq_true_iff_exists_mem] + exists ⟨k,v⟩ + simp[mem] + rw [getValueCast_eq_getValueCast!, getValueCast!_insertList_toInsert_mem] + · exact k_eq + · exact distinct + · exact distinct2 + · exact mem + +theorem getValueCast_insertList_not_toInsert_mem [BEq α] [LawfulBEq α] (l toInsert: List ((a : α) × β a)) {k k': α} {k_eq: k == k'} {not_mem: ¬ k ∈ List.map Sigma.fst toInsert} {distinct: DistinctKeys l} {contains: containsKey k l}: getValueCast k' (insertList l toInsert) (contains_insertList_of_contains_first k_eq contains) = getValueCast k' l (containsKey_of_beq contains k_eq) := by + have inh: Inhabited (β k') := by + apply Inhabited_of_containsKey l k k' k_eq + exact contains + rw [getValueCast_eq_getValueCast!, getValueCast_eq_getValueCast!, getValueCast!_insertList_not_toInsert_mem] + · exact k + · exact k_eq + · exact not_mem + · exact distinct theorem insertList_perm [BEq α] [ReflBEq α] [PartialEquivBEq α] (l toInsert: List ((a : α) × β a)) (distinct_l: DistinctKeys l) (distinct_toInsert: DistinctKeys toInsert) (distinct_both: ∀ (a:α), ¬ (containsKey a l ∧ containsKey a toInsert)): Perm (insertList l toInsert) (l++toInsert) := by diff --git a/src/Std/Data/DHashMap/Internal/RawLemmas.lean b/src/Std/Data/DHashMap/Internal/RawLemmas.lean index 00cb3e9cc448..2594e148d999 100644 --- a/src/Std/Data/DHashMap/Internal/RawLemmas.lean +++ b/src/Std/Data/DHashMap/Internal/RawLemmas.lean @@ -899,39 +899,49 @@ theorem contains_insertList [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l: L (m.insertList l).contains k ↔ m.contains k ∨ (l.map Sigma.fst).contains k := by simp_to_model using List.containsKey_insertList -theorem contains_insertList_of_contains_list [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l: List ((a:α) × (β a))} {k: α} {v: β k}: ⟨k,v⟩ ∈ l → (m.insertList l).contains k := by - rw [contains_insertList _ h] - rw [List.contains_iff_exists_mem_beq] - intro h - right - exists k - simp - exists ⟨k,v⟩ +theorem contains_insertList_of_mem_list [LawfulBEq α] [LawfulHashable α] (h : m.1.WF) {l: List ((a:α) × (β a))} {k k': α} (k_eq: k == k') {v: β k}: ⟨k,v⟩ ∈ l → (m.insertList l).contains k' := by + simp_to_model using contains_insertList_of_mem theorem contains_of_contains_insertList [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l: List ((a:α) × (β a))} {k: α} : (m.insertList l).contains k → (l.map Sigma.fst).contains k = false → m.contains k := by simp_to_model using List.containsKey_of_containsKey_insertList +theorem contains_insertList_of_contains_map [LawfulBEq α] [LawfulHashable α] (h : m.1.WF) {l: List ((a:α) × (β a))} {k k': α} (k_eq: k == k'): m.contains k = true → (m.insertList l).contains k' := by + simp_to_model using contains_insertList_of_contains_first + theorem get?_insertList_mem_list [LawfulBEq α] [LawfulHashable α] {l: List ((a:α) × (β a))} {k k': α} {v: β k} {distinct: List.Pairwise (fun a b => ! a.1 == b.1) l} {k_eq: k == k'} {mem: ⟨k,v⟩ ∈ l} (h: m.1.WF): (m.insertList l).get? k' = some (cast (by congr; apply LawfulBEq.eq_of_beq k_eq) v) := by simp_to_model using getValueCast?_insertList_toInsert_mem -theorem get?_insertList_not_mem_list [LawfulBEq α] [LawfulHashable α] {l: List ((a:α) × (β a))} {k: α} {distinct: List.Pairwise (fun a b => ! a.1 == b.1) l} {mem: ¬ k ∈ (l.map (Sigma.fst))} (h: m.1.WF): (m.insertList l).get? k = m.get? k := by - simp_to_model - sorry +theorem get?_insertList_not_mem_list [LawfulBEq α] [LawfulHashable α] {l: List ((a:α) × (β a))} {k k': α} {mem: ¬ k ∈ (l.map (Sigma.fst))} {k_eq: k == k'} (h: m.1.WF): (m.insertList l).get? k' = m.get? k' := by + simp_to_model using getValueCast?_insertList_not_toInsert_mem theorem get!_insertList_mem_list [LawfulBEq α] [LawfulHashable α] {l: List ((a:α) × (β a))} {k k': α} {v: β k}[Inhabited (β k')] {distinct: List.Pairwise (fun a b => ! a.1 == b.1) l} {k_eq: k == k'} {mem: ⟨k,v⟩ ∈ l} (h: m.1.WF): (m.insertList l).get! k' = cast (by congr; apply LawfulBEq.eq_of_beq k_eq) v := by - simp_to_model - sorry + simp_to_model using getValueCast!_insertList_toInsert_mem + +theorem get!_insertList_not_mem_list [LawfulBEq α] [LawfulHashable α] {l: List ((a:α) × (β a))} {k k': α} [Inhabited (β k')] {mem: ¬ k ∈ (l.map (Sigma.fst))} {k_eq: k == k'} (h: m.1.WF): (m.insertList l).get! k' = m.get! k' := by + simp_to_model using getValueCast!_insertList_not_toInsert_mem theorem getD_insertList_mem_list [LawfulBEq α] [LawfulHashable α] {l: List ((a:α) × (β a))} {k k': α} {v: β k} {fallback: β k'} {distinct: List.Pairwise (fun a b => ! a.1 == b.1) l} {k_eq: k == k'} {mem: ⟨k,v⟩ ∈ l} (h: m.1.WF): (m.insertList l).getD k' fallback = cast (by congr; apply LawfulBEq.eq_of_beq k_eq) v := by - simp_to_model - sorry + simp_to_model using getValueCastD_insertList_toInsert_mem -theorem get_insertList_mem_list [LawfulBEq α] [LawfulHashable α] {l: List ((a:α) × (β a))} {k k': α} {v: β k}[Inhabited (β k)] {distinct: List.Pairwise (fun a b => ! a.1 == b.1) l} {k_eq: k == k'} {mem: ⟨k,v⟩ ∈ l} (h: m.1.WF): (m.insertList l).get k (contains_insertList_of_contains_list _ h mem)= v := by +theorem getD_insertList_not_mem_list [LawfulBEq α] [LawfulHashable α] {l: List ((a:α) × (β a))} {k k': α} {fallback: β k'} {mem: ¬ k ∈ (l.map (Sigma.fst))} {k_eq: k == k'} (h: m.1.WF): (m.insertList l).getD k' fallback = m.getD k' fallback := by + simp_to_model using getValueCastD_insertList_not_toInsert_mem + + +theorem get_insertList_mem_list [LawfulBEq α] [LawfulHashable α] {l: List ((a:α) × (β a))} {k k': α} {v: β k} {distinct: List.Pairwise (fun a b => ! a.1 == b.1) l} {k_eq: k == k'} {mem: ⟨k,v⟩ ∈ l} (h: m.1.WF): (m.insertList l).get k' (contains_insertList_of_mem_list _ h k_eq mem)= cast (by congr; apply LawfulBEq.eq_of_beq k_eq ) v := by + simp_to_model using getValueCast_insertList_toInsert_mem + +theorem contains_of_beq [EquivBEq α][LawfulHashable α] {k k': α} (k_eq: k == k') (h: m.1.WF): + m.contains k = true → m.contains k' = true := by simp_to_model - sorry + intro h' + apply containsKey_of_beq h' k_eq + +theorem get_insertList_not_mem_list [LawfulBEq α] [LawfulHashable α] {l: List ((a:α) × (β a))} {k k': α} {v: β k}[Inhabited (β k)]{k_eq: k == k'} {mem_list: ¬ k ∈ (l.map (Sigma.fst))} (h: m.1.WF): + (h': m.contains k = true) → (m.insertList l).get k' (contains_insertList_of_contains_map m h k_eq h') = m.get k' (contains_of_beq m k_eq h h') := by + simp_to_model using getValueCast_insertList_not_toInsert_mem theorem size_insertList [EquivBEq α] [LawfulHashable α] {l: List ((a:α) × (β a))} {distinct: DistinctKeys l} (h: m.1.WF): (∀ (a:α), ¬ (m.contains a = true ∧ List.containsKey a l = true)) → (m.insertList l).1.size = m.1.size + l.length := by simp_to_model @@ -943,10 +953,12 @@ theorem size_insertList [EquivBEq α] [LawfulHashable α] {l: List ((a:α) × ( . exact distinct . apply distinct' -theorem insertList_notEmpty_if_m_notEmpty [EquivBEq α] [LawfulHashable α] {l: List ((a:α) × (β a))}(h: m.1.WF): (m.1.isEmpty = false) → (m.insertList l).1.isEmpty = false := by +theorem insertList_notEmpty_if_m_notEmpty [EquivBEq α] [LawfulHashable α] {l: List ((a:α) × (β a))}(h: m.1.WF): + (m.1.isEmpty = false) → (m.insertList l).1.isEmpty = false := by simp_to_model using List.insertList_not_isEmpty_if_start_not_isEmpty -theorem insertList_isEmpty [EquivBEq α] [LawfulHashable α] {l: List ((a:α) × (β a))}(h: m.1.WF): (m.insertList l).1.isEmpty ↔ m.1.isEmpty ∧ l.isEmpty := by +theorem insertList_isEmpty [EquivBEq α] [LawfulHashable α] {l: List ((a:α) × (β a))}(h: m.1.WF): + (m.insertList l).1.isEmpty ↔ m.1.isEmpty ∧ l.isEmpty := by simp_to_model using List.insertList_isEmpty end Raw₀ From f023e7f8bb054ab9e32839d4fc2d5e36964b1fa7 Mon Sep 17 00:00:00 2001 From: jt0202 Date: Thu, 21 Nov 2024 21:36:09 +0100 Subject: [PATCH 16/83] Some formatting --- src/Std/Data/DHashMap/Internal/RawLemmas.lean | 46 ++++++++++++------- 1 file changed, 30 insertions(+), 16 deletions(-) diff --git a/src/Std/Data/DHashMap/Internal/RawLemmas.lean b/src/Std/Data/DHashMap/Internal/RawLemmas.lean index 2594e148d999..a49b62e38369 100644 --- a/src/Std/Data/DHashMap/Internal/RawLemmas.lean +++ b/src/Std/Data/DHashMap/Internal/RawLemmas.lean @@ -859,20 +859,23 @@ theorem insertMany_val (m : Raw₀ α β) (l : List ((a : α) × β a)) : rw [ih] @[simp] -theorem insertMany_eq_insertList - (m : Raw₀ α β) (l : List ((a : α) × β a)) : (insertMany m l).val = insertList m l := by +theorem insertMany_eq_insertList (m : Raw₀ α β) (l : List ((a : α) × β a)) : + (insertMany m l).val = insertList m l := by rw [insertList_eq_foldl, insertMany_val] @[simp] -theorem insertList_nil: m.insertList [] = m := by +theorem insertList_nil: + m.insertList [] = m := by simp[insertList, Id.run] @[simp] -theorem insertList_singleton [EquivBEq α] [LawfulHashable α] {k: α} {v: β k}: m.insertList [⟨k,v⟩] = m.insert k v := by +theorem insertList_singleton [EquivBEq α] [LawfulHashable α] {k: α} {v: β k}: + m.insertList [⟨k,v⟩] = m.insert k v := by simp[insertList, Id.run] @[simp] -theorem insertList_cons {l: List ((a:α) × (β a))} {k: α} {v: β k}: m.insertList (⟨k,v⟩::l) = (m.insert k v).insertList l := by +theorem insertList_cons {l: List ((a:α) × (β a))} {k: α} {v: β k}: + m.insertList (⟨k,v⟩::l) = (m.insert k v).insertList l := by cases l with | nil => simp[insertList] | cons hd tl => simp[insertList] @@ -887,7 +890,7 @@ theorem insertList_insertList {l1 l2: List ((a:α) × (β a))}: apply ih theorem insertList_insert {l: List ((a:α) × (β a))} {k: α} {v: β k}: - (m.insertList l).insert k v = m.insertList (l ++ [⟨k,v⟩]) := by + (m.insertList l).insert k v = m.insertList (l ++ [⟨k,v⟩]) := by simp [insertList_eq_insertListₘ] induction l generalizing m with | nil => simp[insertListₘ] @@ -899,7 +902,8 @@ theorem contains_insertList [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l: L (m.insertList l).contains k ↔ m.contains k ∨ (l.map Sigma.fst).contains k := by simp_to_model using List.containsKey_insertList -theorem contains_insertList_of_mem_list [LawfulBEq α] [LawfulHashable α] (h : m.1.WF) {l: List ((a:α) × (β a))} {k k': α} (k_eq: k == k') {v: β k}: ⟨k,v⟩ ∈ l → (m.insertList l).contains k' := by +theorem contains_insertList_of_mem_list [LawfulBEq α] [LawfulHashable α] (h : m.1.WF) {l: List ((a:α) × (β a))} {k k': α} (k_eq: k == k') {v: β k}: + ⟨k,v⟩ ∈ l → (m.insertList l).contains k' := by simp_to_model using contains_insertList_of_mem @@ -907,30 +911,38 @@ theorem contains_of_contains_insertList [EquivBEq α] [LawfulHashable α] (h : m (m.insertList l).contains k → (l.map Sigma.fst).contains k = false → m.contains k := by simp_to_model using List.containsKey_of_containsKey_insertList -theorem contains_insertList_of_contains_map [LawfulBEq α] [LawfulHashable α] (h : m.1.WF) {l: List ((a:α) × (β a))} {k k': α} (k_eq: k == k'): m.contains k = true → (m.insertList l).contains k' := by +theorem contains_insertList_of_contains_map [LawfulBEq α] [LawfulHashable α] (h : m.1.WF) {l: List ((a:α) × (β a))} {k k': α} (k_eq: k == k'): + m.contains k = true → (m.insertList l).contains k' := by simp_to_model using contains_insertList_of_contains_first -theorem get?_insertList_mem_list [LawfulBEq α] [LawfulHashable α] {l: List ((a:α) × (β a))} {k k': α} {v: β k} {distinct: List.Pairwise (fun a b => ! a.1 == b.1) l} {k_eq: k == k'} {mem: ⟨k,v⟩ ∈ l} (h: m.1.WF): (m.insertList l).get? k' = some (cast (by congr; apply LawfulBEq.eq_of_beq k_eq) v) := by +theorem get?_insertList_mem_list [LawfulBEq α] [LawfulHashable α] {l: List ((a:α) × (β a))} {k k': α} {v: β k} {distinct: List.Pairwise (fun a b => ! a.1 == b.1) l} {k_eq: k == k'} {mem: ⟨k,v⟩ ∈ l} (h: m.1.WF): + (m.insertList l).get? k' = some (cast (by congr; apply LawfulBEq.eq_of_beq k_eq) v) := by simp_to_model using getValueCast?_insertList_toInsert_mem -theorem get?_insertList_not_mem_list [LawfulBEq α] [LawfulHashable α] {l: List ((a:α) × (β a))} {k k': α} {mem: ¬ k ∈ (l.map (Sigma.fst))} {k_eq: k == k'} (h: m.1.WF): (m.insertList l).get? k' = m.get? k' := by +theorem get?_insertList_not_mem_list [LawfulBEq α] [LawfulHashable α] {l: List ((a:α) × (β a))} {k k': α} {mem: ¬ k ∈ (l.map (Sigma.fst))} {k_eq: k == k'} (h: m.1.WF): + (m.insertList l).get? k' = m.get? k' := by simp_to_model using getValueCast?_insertList_not_toInsert_mem -theorem get!_insertList_mem_list [LawfulBEq α] [LawfulHashable α] {l: List ((a:α) × (β a))} {k k': α} {v: β k}[Inhabited (β k')] {distinct: List.Pairwise (fun a b => ! a.1 == b.1) l} {k_eq: k == k'} {mem: ⟨k,v⟩ ∈ l} (h: m.1.WF): (m.insertList l).get! k' = cast (by congr; apply LawfulBEq.eq_of_beq k_eq) v := by +theorem get!_insertList_mem_list [LawfulBEq α] [LawfulHashable α] {l: List ((a:α) × (β a))} {k k': α} {v: β k}[Inhabited (β k')] {distinct: List.Pairwise (fun a b => ! a.1 == b.1) l} {k_eq: k == k'} {mem: ⟨k,v⟩ ∈ l} (h: m.1.WF): + (m.insertList l).get! k' = cast (by congr; apply LawfulBEq.eq_of_beq k_eq) v := by simp_to_model using getValueCast!_insertList_toInsert_mem -theorem get!_insertList_not_mem_list [LawfulBEq α] [LawfulHashable α] {l: List ((a:α) × (β a))} {k k': α} [Inhabited (β k')] {mem: ¬ k ∈ (l.map (Sigma.fst))} {k_eq: k == k'} (h: m.1.WF): (m.insertList l).get! k' = m.get! k' := by +theorem get!_insertList_not_mem_list [LawfulBEq α] [LawfulHashable α] {l: List ((a:α) × (β a))} {k k': α} [Inhabited (β k')] {mem: ¬ k ∈ (l.map (Sigma.fst))} {k_eq: k == k'} (h: m.1.WF): + (m.insertList l).get! k' = m.get! k' := by simp_to_model using getValueCast!_insertList_not_toInsert_mem -theorem getD_insertList_mem_list [LawfulBEq α] [LawfulHashable α] {l: List ((a:α) × (β a))} {k k': α} {v: β k} {fallback: β k'} {distinct: List.Pairwise (fun a b => ! a.1 == b.1) l} {k_eq: k == k'} {mem: ⟨k,v⟩ ∈ l} (h: m.1.WF): (m.insertList l).getD k' fallback = cast (by congr; apply LawfulBEq.eq_of_beq k_eq) v := by +theorem getD_insertList_mem_list [LawfulBEq α] [LawfulHashable α] {l: List ((a:α) × (β a))} {k k': α} {v: β k} {fallback: β k'} {distinct: List.Pairwise (fun a b => ! a.1 == b.1) l} {k_eq: k == k'} {mem: ⟨k,v⟩ ∈ l} (h: m.1.WF): + (m.insertList l).getD k' fallback = cast (by congr; apply LawfulBEq.eq_of_beq k_eq) v := by simp_to_model using getValueCastD_insertList_toInsert_mem -theorem getD_insertList_not_mem_list [LawfulBEq α] [LawfulHashable α] {l: List ((a:α) × (β a))} {k k': α} {fallback: β k'} {mem: ¬ k ∈ (l.map (Sigma.fst))} {k_eq: k == k'} (h: m.1.WF): (m.insertList l).getD k' fallback = m.getD k' fallback := by +theorem getD_insertList_not_mem_list [LawfulBEq α] [LawfulHashable α] {l: List ((a:α) × (β a))} {k k': α} {fallback: β k'} {mem: ¬ k ∈ (l.map (Sigma.fst))} {k_eq: k == k'} (h: m.1.WF): + (m.insertList l).getD k' fallback = m.getD k' fallback := by simp_to_model using getValueCastD_insertList_not_toInsert_mem -theorem get_insertList_mem_list [LawfulBEq α] [LawfulHashable α] {l: List ((a:α) × (β a))} {k k': α} {v: β k} {distinct: List.Pairwise (fun a b => ! a.1 == b.1) l} {k_eq: k == k'} {mem: ⟨k,v⟩ ∈ l} (h: m.1.WF): (m.insertList l).get k' (contains_insertList_of_mem_list _ h k_eq mem)= cast (by congr; apply LawfulBEq.eq_of_beq k_eq ) v := by +theorem get_insertList_mem_list [LawfulBEq α] [LawfulHashable α] {l: List ((a:α) × (β a))} {k k': α} {v: β k} {distinct: List.Pairwise (fun a b => ! a.1 == b.1) l} {k_eq: k == k'} {mem: ⟨k,v⟩ ∈ l} (h: m.1.WF): + (m.insertList l).get k' (contains_insertList_of_mem_list _ h k_eq mem)= cast (by congr; apply LawfulBEq.eq_of_beq k_eq ) v := by simp_to_model using getValueCast_insertList_toInsert_mem theorem contains_of_beq [EquivBEq α][LawfulHashable α] {k k': α} (k_eq: k == k') (h: m.1.WF): @@ -943,7 +955,9 @@ theorem get_insertList_not_mem_list [LawfulBEq α] [LawfulHashable α] {l: List (h': m.contains k = true) → (m.insertList l).get k' (contains_insertList_of_contains_map m h k_eq h') = m.get k' (contains_of_beq m k_eq h h') := by simp_to_model using getValueCast_insertList_not_toInsert_mem -theorem size_insertList [EquivBEq α] [LawfulHashable α] {l: List ((a:α) × (β a))} {distinct: DistinctKeys l} (h: m.1.WF): (∀ (a:α), ¬ (m.contains a = true ∧ List.containsKey a l = true)) → (m.insertList l).1.size = m.1.size + l.length := by +theorem size_insertList [EquivBEq α] [LawfulHashable α] {l: List ((a:α) × (β a))} {distinct: DistinctKeys l} (h: m.1.WF): + (∀ (a:α), ¬ (m.contains a = true ∧ List.containsKey a l = true)) → + (m.insertList l).1.size = m.1.size + l.length := by simp_to_model rw [← List.length_append] intro distinct' From fdd6babfcb450e3116dcb9b86dd55aa63521af0a Mon Sep 17 00:00:00 2001 From: jt0202 Date: Thu, 21 Nov 2024 23:31:13 +0100 Subject: [PATCH 17/83] Pull insertList up to DHashMap --- src/Std/Data/DHashMap/Basic.lean | 5 +++++ src/Std/Data/DHashMap/Internal/Raw.lean | 16 ++++++++++++++++ src/Std/Data/DHashMap/Internal/RawLemmas.lean | 2 +- src/Std/Data/DHashMap/Internal/WF.lean | 1 + src/Std/Data/DHashMap/Raw.lean | 18 ++++++++++++++++++ src/Std/Data/DHashMap/RawLemmas.lean | 3 ++- 6 files changed, 43 insertions(+), 2 deletions(-) diff --git a/src/Std/Data/DHashMap/Basic.lean b/src/Std/Data/DHashMap/Basic.lean index 8ddbfd80216a..400ad02ea8a5 100644 --- a/src/Std/Data/DHashMap/Basic.lean +++ b/src/Std/Data/DHashMap/Basic.lean @@ -285,6 +285,11 @@ but will later become a primitive operation. ⟨(Raw₀.insertMany ⟨m.1, m.2.size_buckets_pos⟩ l).1, (Raw₀.insertMany ⟨m.1, m.2.size_buckets_pos⟩ l).2 _ Raw.WF.insert₀ m.2⟩ +@[inline, inherit_doc Raw.insertList] def insertList (m : DHashMap α β) + (l: List ((a:α) × (β a))): DHashMap α β := + ⟨(Raw₀.insertList ⟨m.1, m.2.size_buckets_pos⟩ l).1, + .insertList₀ m.2⟩ + @[inline, inherit_doc Raw.Const.insertMany] def Const.insertMany {β : Type v} {ρ : Type w} [ForIn Id ρ (α × β)] (m : DHashMap α (fun _ => β)) (l : ρ) : DHashMap α (fun _ => β) := diff --git a/src/Std/Data/DHashMap/Internal/Raw.lean b/src/Std/Data/DHashMap/Internal/Raw.lean index 810804db8e75..30703cea712b 100644 --- a/src/Std/Data/DHashMap/Internal/Raw.lean +++ b/src/Std/Data/DHashMap/Internal/Raw.lean @@ -199,6 +199,22 @@ theorem filter_val [BEq α] [Hashable α] {m : Raw₀ α β} {f : (a : α) → m.val.filter f = m.filter f := by simp [Raw.filter, m.2] +theorem insertMany_eq [BEq α] [Hashable α] {m : Raw α β} (h : m.WF) {ρ : Type w} [ForIn Id ρ ((a : α) × β a)] {l : ρ}: + m.insertMany l = Raw₀.insertMany ⟨m, h.size_buckets_pos⟩ l := by + simp[Raw.insertMany, h.size_buckets_pos] + +theorem insertMany_val [BEq α][Hashable α] {m: Raw₀ α β} {ρ : Type w} [ForIn Id ρ ((a : α) × β a)] {l : ρ}: + m.val.insertMany l = m.insertMany l := by + simp[Raw.insertMany, m.2] + +theorem insertList_eq [BEq α] [Hashable α] {m : Raw α β} (h : m.WF) {l: List ((a : α) × β a)}: + m.insertList l = Raw₀.insertList ⟨m, h.size_buckets_pos⟩ l := by + simp[Raw.insertList, h.size_buckets_pos] + +theorem insertList_val [BEq α] [Hashable α] {m : Raw₀ α β} {l: List ((a : α) × β a)}: + m.val.insertList l = m.insertList l := by + simp[Raw.insertList, m.2] + section variable {β : Type v} diff --git a/src/Std/Data/DHashMap/Internal/RawLemmas.lean b/src/Std/Data/DHashMap/Internal/RawLemmas.lean index a49b62e38369..d079b02b50d9 100644 --- a/src/Std/Data/DHashMap/Internal/RawLemmas.lean +++ b/src/Std/Data/DHashMap/Internal/RawLemmas.lean @@ -951,7 +951,7 @@ theorem contains_of_beq [EquivBEq α][LawfulHashable α] {k k': α} (k_eq: k == intro h' apply containsKey_of_beq h' k_eq -theorem get_insertList_not_mem_list [LawfulBEq α] [LawfulHashable α] {l: List ((a:α) × (β a))} {k k': α} {v: β k}[Inhabited (β k)]{k_eq: k == k'} {mem_list: ¬ k ∈ (l.map (Sigma.fst))} (h: m.1.WF): +theorem get_insertList_not_mem_list [LawfulBEq α] [LawfulHashable α] {l: List ((a:α) × (β a))} {k k': α} {k_eq: k == k'} {mem_list: ¬ k ∈ (l.map (Sigma.fst))} (h: m.1.WF): (h': m.contains k = true) → (m.insertList l).get k' (contains_insertList_of_contains_map m h k_eq h') = m.get k' (contains_of_beq m k_eq h h') := by simp_to_model using getValueCast_insertList_not_toInsert_mem diff --git a/src/Std/Data/DHashMap/Internal/WF.lean b/src/Std/Data/DHashMap/Internal/WF.lean index a20634c76906..09a333723b8f 100644 --- a/src/Std/Data/DHashMap/Internal/WF.lean +++ b/src/Std/Data/DHashMap/Internal/WF.lean @@ -757,6 +757,7 @@ theorem WF.out [BEq α] [Hashable α] [i₁ : EquivBEq α] [i₂ : LawfulHashabl · next h => exact Raw₀.wfImp_getThenInsertIfNew? (by apply h) · next h => exact Raw₀.wfImp_filter (by apply h) · next h => exact Raw₀.Const.wfImp_getThenInsertIfNew? (by apply h) + · next h => exact Raw₀.wfImp_insertList (by apply h) end Raw diff --git a/src/Std/Data/DHashMap/Raw.lean b/src/Std/Data/DHashMap/Raw.lean index 153cc4c65220..f71ed2e41e51 100644 --- a/src/Std/Data/DHashMap/Raw.lean +++ b/src/Std/Data/DHashMap/Raw.lean @@ -389,6 +389,16 @@ appearance. (Raw₀.insertMany ⟨m, h⟩ l).1 else m -- will never happen for well-formed inputs +/-- +Inserts multiple mappings into the hash map by iterating over the given list and calling +`insert`. If the same key appears multiple times, the last occurrence takes precedence. This is equal to `insertMany`, but allows easier proofs. +-/ +@[inline] def insertList [BEq α] [Hashable α] + (m: Raw α β)(l: List ((a : α) × β a)): Raw α β := + if h : 0 < m.buckets.size then + (Raw₀.insertList ⟨m,h⟩ l).1 + else m -- will never happen for well-formed inputs + @[inline, inherit_doc Raw.insertMany] def Const.insertMany {β : Type v} [BEq α] [Hashable α] {ρ : Type w} [ForIn Id ρ (α × β)] (m : Raw α (fun _ => β)) (l : ρ) : Raw α (fun _ => β) := if h : 0 < m.buckets.size then @@ -496,6 +506,8 @@ inductive WF : {α : Type u} → {β : α → Type v} → [BEq α] → [Hashable /-- Internal implementation detail of the hash map -/ | constGetThenInsertIfNew?₀ {α β} [BEq α] [Hashable α] {m : Raw α (fun _ => β)} {h a b} : WF m → WF (Raw₀.Const.getThenInsertIfNew? ⟨m, h⟩ a b).2.1 + /-- Internal implementation detail of the hash map -/ + | insertList₀ {α β} [BEq α] [Hashable α] {m : Raw α β} {l: List ((a: α) × (β a))} {h}: WF m → WF (Raw₀.insertList ⟨m,h⟩ l).1 /-- Internal implementation detail of the hash map -/ theorem WF.size_buckets_pos [BEq α] [Hashable α] (m : Raw α β) : WF m → 0 < m.buckets.size @@ -509,6 +521,7 @@ theorem WF.size_buckets_pos [BEq α] [Hashable α] (m : Raw α β) : WF m → 0 | getThenInsertIfNew?₀ _ => (Raw₀.getThenInsertIfNew? ⟨_, _⟩ _ _).2.2 | filter₀ _ => (Raw₀.filter _ ⟨_, _⟩).2 | constGetThenInsertIfNew?₀ _ => (Raw₀.Const.getThenInsertIfNew? ⟨_, _⟩ _ _).2.2 + | insertList₀ _ => (Raw₀.insertList ⟨_,_⟩ _).2 @[simp] theorem WF.empty [BEq α] [Hashable α] {c : Nat} : (Raw.empty c : Raw α β).WF := .empty₀ @@ -552,6 +565,11 @@ theorem WF.insertMany [BEq α] [Hashable α] {ρ : Type w} [ForIn Id ρ ((a : α simpa [Raw.insertMany, h.size_buckets_pos] using (Raw₀.insertMany ⟨m, h.size_buckets_pos⟩ l).2 _ WF.insert₀ h +theorem WF.insertList [BEq α] [Hashable α] {m : Raw α β} + {l : List ((a : α) × (β a))} (h : m.WF): (m.insertList l).WF := by + simpa [Raw.insertList, h.size_buckets_pos] using + .insertList₀ h + theorem WF.Const.insertMany {β : Type v} [BEq α] [Hashable α] {ρ : Type w} [ForIn Id ρ (α × β)] {m : Raw α (fun _ => β)} {l : ρ} (h : m.WF) : (Const.insertMany m l).WF := by simpa [Raw.Const.insertMany, h.size_buckets_pos] using diff --git a/src/Std/Data/DHashMap/RawLemmas.lean b/src/Std/Data/DHashMap/RawLemmas.lean index 42d4db498753..0cb2d343f174 100644 --- a/src/Std/Data/DHashMap/RawLemmas.lean +++ b/src/Std/Data/DHashMap/RawLemmas.lean @@ -39,6 +39,7 @@ private def baseNames : Array Name := ``getThenInsertIfNew?_snd_eq, ``getThenInsertIfNew?_snd_val, ``map_eq, ``map_val, ``filter_eq, ``filter_val, + ``insertList_eq, ``insertList_val, ``erase_eq, ``erase_val, ``filterMap_eq, ``filterMap_val, ``Const.getThenInsertIfNew?_snd_eq, ``Const.getThenInsertIfNew?_snd_val, @@ -1039,7 +1040,7 @@ theorem contains_keys [EquivBEq α] [LawfulHashable α] (h : m.WF) {k : α} : @[simp] theorem mem_keys [LawfulBEq α] [LawfulHashable α] (h : m.WF) {k : α} : - k ∈ m.keys ↔ k ∈ m := by + k ∈ m.keys ↔ k ∈ m := by rw [mem_iff_contains] simp_to_raw using Raw₀.mem_keys ⟨m, _⟩ h From 273a0d0934ca92d07a52b25c84583f53776bf940 Mon Sep 17 00:00:00 2001 From: monsterkrampe Date: Mon, 18 Nov 2024 18:26:53 +0100 Subject: [PATCH 18/83] Add results for getKey?_insertList --- .../DHashMap/Internal/List/Associative.lean | 63 +++++++++++++++++++ src/Std/Data/DHashMap/Internal/RawLemmas.lean | 63 +++++++++++++++++++ 2 files changed, 126 insertions(+) diff --git a/src/Std/Data/DHashMap/Internal/List/Associative.lean b/src/Std/Data/DHashMap/Internal/List/Associative.lean index 7f3f3b3fcf41..530c686094a8 100644 --- a/src/Std/Data/DHashMap/Internal/List/Associative.lean +++ b/src/Std/Data/DHashMap/Internal/List/Associative.lean @@ -7,6 +7,7 @@ prelude import Init.Data.BEq import Init.Data.Nat.Simproc import Init.Data.List.Perm +import Init.Data.List.Find import Std.Data.DHashMap.Internal.List.Defs /-! @@ -2111,6 +2112,15 @@ theorem insertList_perm_of_perm_first [BEq α] [EquivBEq α] (l1 l2 toInsert: Li exact h apply DistinctKeys.insertEntry distinct +theorem containsKey_eq_contains_map_fst [BEq α] [PartialEquivBEq α] (l : List ((a : α) × β a)) (k : α) : containsKey k l = +(l.map Sigma.fst).contains k := by + induction l with + | nil => simp + | cons hd tl ih => + rw [containsKey_cons, ih] + simp + rw [BEq.comm] + theorem containsKey_insertList [BEq α] [PartialEquivBEq α] (l toInsert : List ((a : α) × β a)) (k: α): containsKey k (List.insertList l toInsert) ↔ containsKey k l ∨ (toInsert.map Sigma.fst).contains k := by induction toInsert generalizing l with | nil => simp[insertList] @@ -2361,6 +2371,59 @@ theorem getValueCast_insertList_not_toInsert_mem [BEq α] [LawfulBEq α] (l toIn · exact not_mem · exact distinct +theorem getKey?_insertList [BEq α] [PartialEquivBEq α] {l toInsert : List ((a : α) × β a)} {k : α} : + getKey? k (insertList l toInsert) = + if (toInsert.map Sigma.fst).contains k then (toInsert.map Sigma.fst).reverse.find? (fun a => k == a) else getKey? k l := by + rw [← containsKey_eq_contains_map_fst] + induction toInsert generalizing l with + | nil => simp [insertList] + | cons hd tl ih => + unfold insertList + rw [ih] + rw [getKey?_insertEntry] + rw [containsKey_cons] + simp + cases eq : containsKey k tl with + | true => + simp + rw [containsKey_eq_contains_map_fst, List.contains_iff_exists_mem_beq] at eq + rw [Option.or_of_isSome] + simp + rcases eq with ⟨a, a_mem, a_eq⟩ + simp at a_mem + rcases a_mem with ⟨pair, pair_mem, pair_fst_a⟩ + exists a + constructor + . exists pair + . exact a_eq + | false => + simp + cases eq2 : hd.fst == k with + | true => + simp + rw [containsKey_eq_contains_map_fst, ← Bool.not_eq_true, List.contains_iff_exists_mem_beq] at eq + simp at eq + have : (tl.map Sigma.fst).reverse.find? (fun a => k == a) = none := by simp; exact eq + simp [this] + rw [BEq.comm] + exact eq2 + | false => simp + +theorem getKey?_insertList_lawful [BEq α] [LawfulBEq α] {l toInsert : List ((a : α) × β a)} {k : α} : + getKey? k (insertList l toInsert) = + if (toInsert.map Sigma.fst).contains k then some k else getKey? k l := by + rw [← containsKey_eq_contains_map_fst] + induction toInsert generalizing l with + | nil => simp [insertList] + | cons hd tl ih => + unfold insertList + rw [ih] + rw [getKey?_insertEntry] + rw [containsKey_cons] + cases containsKey k tl with + | true => simp + | false => simp; cases eq : hd.fst == k <;> simp; apply LawfulBEq.eq_of_beq; assumption + theorem insertList_perm [BEq α] [ReflBEq α] [PartialEquivBEq α] (l toInsert: List ((a : α) × β a)) (distinct_l: DistinctKeys l) (distinct_toInsert: DistinctKeys toInsert) (distinct_both: ∀ (a:α), ¬ (containsKey a l ∧ containsKey a toInsert)): Perm (insertList l toInsert) (l++toInsert) := by induction toInsert generalizing l with diff --git a/src/Std/Data/DHashMap/Internal/RawLemmas.lean b/src/Std/Data/DHashMap/Internal/RawLemmas.lean index d079b02b50d9..f9f306d964da 100644 --- a/src/Std/Data/DHashMap/Internal/RawLemmas.lean +++ b/src/Std/Data/DHashMap/Internal/RawLemmas.lean @@ -955,6 +955,69 @@ theorem get_insertList_not_mem_list [LawfulBEq α] [LawfulHashable α] {l: List (h': m.contains k = true) → (m.insertList l).get k' (contains_insertList_of_contains_map m h k_eq h') = m.get k' (contains_of_beq m k_eq h h') := by simp_to_model using getValueCast_insertList_not_toInsert_mem +theorem getKey?_insertList [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l : List ((a:α) × (β a))} {k : α} : + (m.insertList l).getKey? k = + if (l.map Sigma.fst).contains k then (l.map Sigma.fst).reverse.find? (fun a => k == a) else m.getKey? k := by + simp_to_model using List.getKey?_insertList + +-- TODO: have the lawful variants for other getKey theorems below +theorem getKey?_insertList_lawful [LawfulBEq α] [LawfulHashable α] (h : m.1.WF) {l : List ((a:α) × (β a))} {k : α} : + (m.insertList l).getKey? k = + if (l.map Sigma.fst).contains k then some k else m.getKey? k := by + simp_to_model using List.getKey?_insertList_lawful + +theorem getKey_insertList [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l : List ((a:α) × (β a))} {k +: α} {h₁} : + (m.insertList l).getKey k h₁ = + if h₂ : (l.map Sigma.fst).contains k then ((l.map Sigma.fst).reverse.find? (fun a => k == a)).get (by + -- TODO: maybe extract this result... + simp + rw [List.contains_iff_exists_mem_beq] at h₂ + rcases h₂ with ⟨a, a_mem, a_eq⟩ + simp at a_mem + rcases a_mem with ⟨pair, pair_mem, pair_fst_a⟩ + exists a + constructor + . exists pair + . exact a_eq) + else m.getKey k (contains_of_contains_insertList _ h h₁ (Bool.eq_false_iff.2 h₂)) := by + simp_to_model + sorry + +theorem getKey!_insertList [EquivBEq α] [LawfulHashable α] [Inhabited α] (h : m.1.WF) {l : List ((a:α) × (β a))} {k : α} : + (m.insertList l).getKey! k = + if h₂ : (l.map Sigma.fst).contains k then ((l.map Sigma.fst).reverse.find? (fun a => k == a)).get (by + -- TODO: maybe extract this result... + simp + rw [List.contains_iff_exists_mem_beq] at h₂ + rcases h₂ with ⟨a, a_mem, a_eq⟩ + simp at a_mem + rcases a_mem with ⟨pair, pair_mem, pair_fst_a⟩ + exists a + constructor + . exists pair + . exact a_eq) + else m.getKey! k := by + simp_to_model + sorry + +theorem getKeyD_insertList [EquivBEq α] [LawfulHashable α] [Inhabited α] (h : m.1.WF) {l : List ((a:α) × (β a))} {k fallback : α} : + (m.insertList l).getKeyD k fallback = + if h₂ : (l.map Sigma.fst).contains k then ((l.map Sigma.fst).reverse.find? (fun a => k == a)).get (by + -- TODO: maybe extract this result... + simp + rw [List.contains_iff_exists_mem_beq] at h₂ + rcases h₂ with ⟨a, a_mem, a_eq⟩ + simp at a_mem + rcases a_mem with ⟨pair, pair_mem, pair_fst_a⟩ + exists a + constructor + . exists pair + . exact a_eq) + else m.getKeyD k fallback := by + simp_to_model + sorry + theorem size_insertList [EquivBEq α] [LawfulHashable α] {l: List ((a:α) × (β a))} {distinct: DistinctKeys l} (h: m.1.WF): (∀ (a:α), ¬ (m.contains a = true ∧ List.containsKey a l = true)) → (m.insertList l).1.size = m.1.size + l.length := by From be70b67cc7449c54c2f47a568396525e26d5bad3 Mon Sep 17 00:00:00 2001 From: monsterkrampe Date: Tue, 19 Nov 2024 14:16:27 +0100 Subject: [PATCH 19/83] Add results for variants of getKey_insertList --- .../DHashMap/Internal/List/Associative.lean | 80 +++++++++++++++++++ src/Std/Data/DHashMap/Internal/RawLemmas.lean | 67 ++++++---------- 2 files changed, 106 insertions(+), 41 deletions(-) diff --git a/src/Std/Data/DHashMap/Internal/List/Associative.lean b/src/Std/Data/DHashMap/Internal/List/Associative.lean index 530c686094a8..3d4370938304 100644 --- a/src/Std/Data/DHashMap/Internal/List/Associative.lean +++ b/src/Std/Data/DHashMap/Internal/List/Associative.lean @@ -18,6 +18,21 @@ File contents: Verification of associative lists -/ +-- TODO: should this go somwhere else? +theorem List.find?_isSome_of_map_fst_contains {α : Type u} {β : α -> Type v} [BEq α] {l : List ((a : α) × β a)} {k : α} : + (l.map Sigma.fst).contains k -> + ((l.map Sigma.fst).reverse.find? (fun a => k == a)).isSome := by + intro h + simp + rw [List.contains_iff_exists_mem_beq] at h + rcases h with ⟨a, a_mem, a_eq⟩ + simp at a_mem + rcases a_mem with ⟨pair, pair_mem, pair_fst_a⟩ + exists a + constructor + . exists pair + . exact a_eq + set_option linter.missingDocs true set_option autoImplicit false @@ -2424,6 +2439,71 @@ theorem getKey?_insertList_lawful [BEq α] [LawfulBEq α] {l toInsert : List ((a | true => simp | false => simp; cases eq : hd.fst == k <;> simp; apply LawfulBEq.eq_of_beq; assumption +theorem getKey_insertList [BEq α] [PartialEquivBEq α] {l toInsert : List ((a : α) × β a)} {k : α} +{h₁} : + getKey k (insertList l toInsert) h₁ = + if h₂ : (toInsert.map Sigma.fst).contains k then ((toInsert.map Sigma.fst).reverse.find? (fun a => + k == a)).get (List.find?_isSome_of_map_fst_contains h₂) else getKey k l + (containsKey_of_containsKey_insertList h₁ (Bool.eq_false_iff.2 h₂)) := by + rw [← Option.some_inj] + rw [← getKey?_eq_some_getKey] + rw [getKey?_insertList] + split + . simp + . rw [getKey?_eq_some_getKey] + +theorem getKey_insertList_lawful [BEq α] [LawfulBEq α] {l toInsert : List ((a : α) × β a)} {k : α} +{h₁} : + getKey k (insertList l toInsert) h₁ = + if h₂ : (toInsert.map Sigma.fst).contains k then k else getKey k l + (containsKey_of_containsKey_insertList h₁ (Bool.eq_false_iff.2 h₂)) := by + rw [← Option.some_inj] + rw [← getKey?_eq_some_getKey] + rw [getKey?_insertList_lawful] + split + . simp + . rw [getKey?_eq_some_getKey] + +theorem getKey!_insertList [BEq α] [PartialEquivBEq α] [Inhabited α] {l toInsert : List ((a : α) × β a)} {k : α} : + getKey! k (insertList l toInsert) = + if h₂ : (toInsert.map Sigma.fst).contains k then ((toInsert.map Sigma.fst).reverse.find? (fun a => + k == a)).get (List.find?_isSome_of_map_fst_contains h₂) else getKey! k l := by + rw [getKey!_eq_getKey?] + rw [getKey?_insertList] + split + . rw [Option.get_eq_get!] + . rw [getKey!_eq_getKey?] + +theorem getKey!_insertList_lawful [BEq α] [LawfulBEq α] [Inhabited α] {l toInsert : List ((a : α) × β a)} {k : α} : + getKey! k (insertList l toInsert) = + if (toInsert.map Sigma.fst).contains k then k else getKey! k l := by + rw [getKey!_eq_getKey?] + rw [getKey?_insertList_lawful] + split + . simp + . rw [getKey!_eq_getKey?] + +theorem getKeyD_insertList [BEq α] [PartialEquivBEq α] {l toInsert : List ((a : α) × β +a)} {k fallback : α} : + getKeyD k (insertList l toInsert) fallback = + if h₂ : (toInsert.map Sigma.fst).contains k then ((toInsert.map Sigma.fst).reverse.find? (fun a => + k == a)).get (List.find?_isSome_of_map_fst_contains h₂) else getKeyD k l fallback := by + rw [getKeyD_eq_getKey?] + rw [getKey?_insertList] + split + . rw [Option.get_eq_getD] + . rw [getKeyD_eq_getKey?] + +theorem getKeyD_insertList_lawful [BEq α] [LawfulBEq α] {l toInsert : List ((a : α) × β a)} {k +fallback : α} : + getKeyD k (insertList l toInsert) fallback = + if (toInsert.map Sigma.fst).contains k then k else getKeyD k l fallback := by + rw [getKeyD_eq_getKey?] + rw [getKey?_insertList_lawful] + split + . simp + . rw [getKeyD_eq_getKey?] + theorem insertList_perm [BEq α] [ReflBEq α] [PartialEquivBEq α] (l toInsert: List ((a : α) × β a)) (distinct_l: DistinctKeys l) (distinct_toInsert: DistinctKeys toInsert) (distinct_both: ∀ (a:α), ¬ (containsKey a l ∧ containsKey a toInsert)): Perm (insertList l toInsert) (l++toInsert) := by induction toInsert generalizing l with diff --git a/src/Std/Data/DHashMap/Internal/RawLemmas.lean b/src/Std/Data/DHashMap/Internal/RawLemmas.lean index f9f306d964da..a90ff8b436d1 100644 --- a/src/Std/Data/DHashMap/Internal/RawLemmas.lean +++ b/src/Std/Data/DHashMap/Internal/RawLemmas.lean @@ -960,7 +960,6 @@ theorem getKey?_insertList [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l : L if (l.map Sigma.fst).contains k then (l.map Sigma.fst).reverse.find? (fun a => k == a) else m.getKey? k := by simp_to_model using List.getKey?_insertList --- TODO: have the lawful variants for other getKey theorems below theorem getKey?_insertList_lawful [LawfulBEq α] [LawfulHashable α] (h : m.1.WF) {l : List ((a:α) × (β a))} {k : α} : (m.insertList l).getKey? k = if (l.map Sigma.fst).contains k then some k else m.getKey? k := by @@ -969,54 +968,40 @@ theorem getKey?_insertList_lawful [LawfulBEq α] [LawfulHashable α] (h : m.1.WF theorem getKey_insertList [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l : List ((a:α) × (β a))} {k : α} {h₁} : (m.insertList l).getKey k h₁ = - if h₂ : (l.map Sigma.fst).contains k then ((l.map Sigma.fst).reverse.find? (fun a => k == a)).get (by - -- TODO: maybe extract this result... - simp - rw [List.contains_iff_exists_mem_beq] at h₂ - rcases h₂ with ⟨a, a_mem, a_eq⟩ - simp at a_mem - rcases a_mem with ⟨pair, pair_mem, pair_fst_a⟩ - exists a - constructor - . exists pair - . exact a_eq) + if h₂ : (l.map Sigma.fst).contains k then ((l.map Sigma.fst).reverse.find? (fun a => k == a)).get (List.find?_isSome_of_map_fst_contains h₂) else m.getKey k (contains_of_contains_insertList _ h h₁ (Bool.eq_false_iff.2 h₂)) := by - simp_to_model - sorry + simp_to_model using List.getKey_insertList + +theorem getKey_insertList_lawful [LawfulBEq α] [LawfulHashable α] (h : m.1.WF) {l : List ((a:α) × (β a))} {k +: α} {h₁} : + (m.insertList l).getKey k h₁ = + if h₂ : (l.map Sigma.fst).contains k then k + else m.getKey k (contains_of_contains_insertList _ h h₁ (Bool.eq_false_iff.2 h₂)) := by + simp_to_model using List.getKey_insertList_lawful theorem getKey!_insertList [EquivBEq α] [LawfulHashable α] [Inhabited α] (h : m.1.WF) {l : List ((a:α) × (β a))} {k : α} : (m.insertList l).getKey! k = - if h₂ : (l.map Sigma.fst).contains k then ((l.map Sigma.fst).reverse.find? (fun a => k == a)).get (by - -- TODO: maybe extract this result... - simp - rw [List.contains_iff_exists_mem_beq] at h₂ - rcases h₂ with ⟨a, a_mem, a_eq⟩ - simp at a_mem - rcases a_mem with ⟨pair, pair_mem, pair_fst_a⟩ - exists a - constructor - . exists pair - . exact a_eq) + if h₂ : (l.map Sigma.fst).contains k then ((l.map Sigma.fst).reverse.find? (fun a => k == a)).get (List.find?_isSome_of_map_fst_contains h₂) else m.getKey! k := by - simp_to_model - sorry + simp_to_model using List.getKey!_insertList -theorem getKeyD_insertList [EquivBEq α] [LawfulHashable α] [Inhabited α] (h : m.1.WF) {l : List ((a:α) × (β a))} {k fallback : α} : +theorem getKey!_insertList_lawful [LawfulBEq α] [LawfulHashable α] [Inhabited α] (h : m.1.WF) {l : List ((a:α) × (β a))} {k : α} : + (m.insertList l).getKey! k = + if (l.map Sigma.fst).contains k then k + else m.getKey! k := by + simp_to_model using List.getKey!_insertList_lawful + +theorem getKeyD_insertList [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l : List ((a:α) × (β a))} {k fallback : α} : (m.insertList l).getKeyD k fallback = - if h₂ : (l.map Sigma.fst).contains k then ((l.map Sigma.fst).reverse.find? (fun a => k == a)).get (by - -- TODO: maybe extract this result... - simp - rw [List.contains_iff_exists_mem_beq] at h₂ - rcases h₂ with ⟨a, a_mem, a_eq⟩ - simp at a_mem - rcases a_mem with ⟨pair, pair_mem, pair_fst_a⟩ - exists a - constructor - . exists pair - . exact a_eq) + if h₂ : (l.map Sigma.fst).contains k then ((l.map Sigma.fst).reverse.find? (fun a => k == a)).get (List.find?_isSome_of_map_fst_contains h₂) else m.getKeyD k fallback := by - simp_to_model - sorry + simp_to_model using List.getKeyD_insertList + +theorem getKeyD_insertList_lawful [LawfulBEq α] [LawfulHashable α] (h : m.1.WF) {l : List ((a:α) × (β a))} {k fallback : α} : + (m.insertList l).getKeyD k fallback = + if (l.map Sigma.fst).contains k then k + else m.getKeyD k fallback := by + simp_to_model using List.getKeyD_insertList_lawful theorem size_insertList [EquivBEq α] [LawfulHashable α] {l: List ((a:α) × (β a))} {distinct: DistinctKeys l} (h: m.1.WF): (∀ (a:α), ¬ (m.contains a = true ∧ List.containsKey a l = true)) → From a5f7e120b46708cdeb7126c8b739094816106a2d Mon Sep 17 00:00:00 2001 From: jt0202 Date: Tue, 26 Nov 2024 18:58:53 +0100 Subject: [PATCH 20/83] Multiple small changes --- src/Std/Data/DHashMap/Basic.lean | 5 - src/Std/Data/DHashMap/Internal/Defs.lean | 8 +- .../DHashMap/Internal/List/Associative.lean | 363 ++++++++---------- src/Std/Data/DHashMap/Internal/Model.lean | 4 +- src/Std/Data/DHashMap/Internal/Raw.lean | 8 - src/Std/Data/DHashMap/Internal/RawLemmas.lean | 74 ++-- src/Std/Data/DHashMap/Internal/WF.lean | 7 +- src/Std/Data/DHashMap/Raw.lean | 17 - src/Std/Data/DHashMap/RawLemmas.lean | 1 - 9 files changed, 194 insertions(+), 293 deletions(-) diff --git a/src/Std/Data/DHashMap/Basic.lean b/src/Std/Data/DHashMap/Basic.lean index 400ad02ea8a5..8ddbfd80216a 100644 --- a/src/Std/Data/DHashMap/Basic.lean +++ b/src/Std/Data/DHashMap/Basic.lean @@ -285,11 +285,6 @@ but will later become a primitive operation. ⟨(Raw₀.insertMany ⟨m.1, m.2.size_buckets_pos⟩ l).1, (Raw₀.insertMany ⟨m.1, m.2.size_buckets_pos⟩ l).2 _ Raw.WF.insert₀ m.2⟩ -@[inline, inherit_doc Raw.insertList] def insertList (m : DHashMap α β) - (l: List ((a:α) × (β a))): DHashMap α β := - ⟨(Raw₀.insertList ⟨m.1, m.2.size_buckets_pos⟩ l).1, - .insertList₀ m.2⟩ - @[inline, inherit_doc Raw.Const.insertMany] def Const.insertMany {β : Type v} {ρ : Type w} [ForIn Id ρ (α × β)] (m : DHashMap α (fun _ => β)) (l : ρ) : DHashMap α (fun _ => β) := diff --git a/src/Std/Data/DHashMap/Internal/Defs.lean b/src/Std/Data/DHashMap/Internal/Defs.lean index 51b602b026fc..6244f30b85d9 100644 --- a/src/Std/Data/DHashMap/Internal/Defs.lean +++ b/src/Std/Data/DHashMap/Internal/Defs.lean @@ -354,12 +354,8 @@ where /-- Internal implementation detail of the hash map -/ def insertList [BEq α] [Hashable α] - (m : Raw₀ α β) (l : List ((a : α) × β a)) : Raw₀ α β := Id.run do - let mut r := m - for ⟨a,b⟩ in l do - r:= r.insert a b - return r - + (m : Raw₀ α β) (l : List ((a : α) × β a)) : Raw₀ α β := + List.foldl (fun a b => insert a b.1 b.2) m l section variable {β : Type v} diff --git a/src/Std/Data/DHashMap/Internal/List/Associative.lean b/src/Std/Data/DHashMap/Internal/List/Associative.lean index 3d4370938304..632a5b5c445f 100644 --- a/src/Std/Data/DHashMap/Internal/List/Associative.lean +++ b/src/Std/Data/DHashMap/Internal/List/Associative.lean @@ -23,10 +23,10 @@ theorem List.find?_isSome_of_map_fst_contains {α : Type u} {β : α -> Type v} (l.map Sigma.fst).contains k -> ((l.map Sigma.fst).reverse.find? (fun a => k == a)).isSome := by intro h - simp + simp only [find?_isSome, mem_reverse, mem_map] rw [List.contains_iff_exists_mem_beq] at h rcases h with ⟨a, a_mem, a_eq⟩ - simp at a_mem + simp only [mem_map] at a_mem rcases a_mem with ⟨pair, pair_mem, pair_fst_a⟩ exists a constructor @@ -180,27 +180,27 @@ theorem getValueCast?_cons_self [BEq α] [LawfulBEq α] {l : List ((a : α) × getValueCast? k (⟨k, v⟩ :: l) = some v := by rw [getValueCast?_cons_of_true BEq.refl, cast_eq] -theorem getValueCast?_mem [BEq α] [LawfulBEq α] {l : List ((a : α) × β a)} {k k': α} {v : β k} {k_eq: k == k'} {mem: ⟨k,v⟩ ∈ l} {distinct: List.Pairwise (fun a b => ! a.1 == b.1) l}: getValueCast? k' l = some (cast (by congr; apply LawfulBEq.eq_of_beq k_eq) v) := by +theorem getValueCast?_mem [BEq α] [LawfulBEq α] {l : List ((a : α) × β a)} {k k': α} {v : β k} {k_eq: k == k'} {mem: ⟨k,v⟩ ∈ l} {distinct: List.Pairwise (fun a b => (a.1 == b.1) = false) l}: getValueCast? k' l = some (cast (by congr; apply LawfulBEq.eq_of_beq k_eq) v) := by induction l with | nil => simp at mem | cons hd tl ih => - simp[getValueCast?] - simp at mem + simp only [getValueCast?, beq_iff_eq] + simp only [List.mem_cons] at mem cases mem with | inl mem => rw [← mem] simp [k_eq] | inr mem => have hd_k': (hd.fst == k') = false := by - simp at distinct - simp + simp only [beq_eq_false_iff_ne, ne_eq, pairwise_cons] at distinct + simp only [beq_eq_false_iff_ne, ne_eq] rcases distinct with ⟨distinct, _⟩ specialize distinct ⟨k,v⟩ mem - simp at distinct - simp at k_eq + simp only at distinct + simp only [beq_iff_eq] at k_eq rw [k_eq] at distinct apply distinct - simp [hd_k'] + simp only [hd_k', Bool.false_eq_true, ↓reduceDIte] apply ih · exact mem · exact List.Pairwise.of_cons distinct @@ -273,7 +273,8 @@ theorem containsKey_eq_false_iff [BEq α][PartialEquivBEq α] {l : List ((a : α induction l with | nil => simp | cons hd tl ih => - simp[containsKey, ih] + simp only [containsKey, Bool.or_eq_false_iff, ih, List.mem_cons, forall_eq_or_imp, + and_congr_left_iff, Bool.coe_false_iff_false, Bool.not_eq_eq_eq_not, Bool.not_not] intro _ rw [Bool.eq_iff_iff] constructor @@ -316,27 +317,6 @@ theorem containsKey_eq_isSome_getEntry? [BEq α] {l : List ((a : α) × β a)} { · simp [getEntry?_cons_of_false h, h, ih] · simp [getEntry?_cons_of_true h, h] --- TODO: not needed; containsKey_congr instead -theorem containsKey_of_eq [BEq α] [PartialEquivBEq α]{l : List ((a : α) × β a)} {a b: α} (eq: a == b): containsKey a l ↔ containsKey b l := by - induction l with - | nil=> simp - | cons hd tl ih => - simp[containsKey] - by_cases hd_a: hd.fst == a - · have hd_b: hd.fst == b := by - apply PartialEquivBEq.trans hd_a eq - simp [hd_a, hd_b] - · have hd_b: (hd.fst == b) = false := by - false_or_by_contra - rename_i h - simp at h - have hd_a': hd.fst == a := by - apply PartialEquivBEq.trans h (PartialEquivBEq.symm eq) - contradiction - simp [hd_a, hd_b] - rw [Bool.eq_iff_iff] - apply ih - theorem isEmpty_eq_false_of_containsKey [BEq α] {l : List ((a : α) × β a)} {a : α} (h : containsKey a l = true) : l.isEmpty = false := by cases l <;> simp_all @@ -738,15 +718,15 @@ theorem replaceEntry_of_containsKey_eq_false [BEq α] {l : List ((a : α) × β rw [containsKey_cons_eq_false] at h rw [replaceEntry_cons_of_false h.1, ih h.2] -theorem replaceEntry_mem_of_mem [BEq α] [LawfulBEq α] {l : List ((a : α) × β a)} {k: α} {v : β k} {ele: (a: α) × β a} (mem: containsKey k l) {distinct: List.Pairwise (fun a b => ! a.1 == b.1) l}: ele ∈ replaceEntry k v l ↔ (ele ∈ l ∧ ¬ (ele.fst == k)) ∨ (ele = ⟨k,v⟩) := by +theorem replaceEntry_mem_of_mem [BEq α] [LawfulBEq α] {l : List ((a : α) × β a)} {k: α} {v : β k} {ele: (a: α) × β a} (mem: containsKey k l) {distinct: List.Pairwise (fun a b => (a.1 == b.1) = false) l}: ele ∈ replaceEntry k v l ↔ (ele ∈ l ∧ ¬ (ele.fst == k)) ∨ (ele = ⟨k,v⟩) := by induction l with | nil => simp at mem | cons hd tl ih => - simp [replaceEntry, cond_eq_if] + simp only [replaceEntry, cond_eq_if, beq_iff_eq, List.mem_cons] split - · simp + · simp only [List.mem_cons] rename_i hd_k - simp at hd_k + simp only [beq_iff_eq] at hd_k constructor · intro h cases h with @@ -757,7 +737,7 @@ theorem replaceEntry_mem_of_mem [BEq α] [LawfulBEq α] {l : List ((a : α) × · right exact h · rw [← hd_k] - simp at distinct + simp only [beq_eq_false_iff_ne, ne_eq, pairwise_cons] at distinct apply Ne.symm apply (And.left distinct ele h) · intro h @@ -774,9 +754,9 @@ theorem replaceEntry_mem_of_mem [BEq α] [LawfulBEq α] {l : List ((a : α) × | inr h => left apply h - · simp + · simp only [List.mem_cons] rename_i hd_k - simp at hd_k + simp only [beq_iff_eq] at hd_k rw [ih] · constructor · intro h @@ -785,12 +765,12 @@ theorem replaceEntry_mem_of_mem [BEq α] [LawfulBEq α] {l : List ((a : α) × left constructor · left - simp[h] - · simp[h, hd_k] + simp only [h] + · simp only [h, hd_k, not_false_eq_true] | inr h => cases h with | inl h => - simp at h + simp only [beq_iff_eq] at h left constructor · right @@ -806,68 +786,28 @@ theorem replaceEntry_mem_of_mem [BEq α] [LawfulBEq α] {l : List ((a : α) × cases h with | inl h => left - simp[h] + simp only [h] | inr h => right left - simp[h, ele_k] + simp [h, ele_k] | inr h => right right exact h · rw [containsKey_cons] at mem - simp[hd_k] at mem + simp only [Bool.or_eq_true, beq_iff_eq, hd_k, false_or] at mem exact mem - · simp at distinct - simp + · simp only [beq_eq_false_iff_ne, ne_eq, pairwise_cons] at distinct + simp only [beq_eq_false_iff_ne, ne_eq] apply And.right distinct - -theorem replaceEntry_distinct_keys [BEq α] [LawfulBEq α] {l : List ((a : α) × β a)} {k : α} {v : β k}: List.Pairwise (fun a b => ! a.1 == b.1) l → List.Pairwise (fun a b => ! a.1 == b.1) (replaceEntry k v l) := by - intro h - induction l with - | nil => simp - | cons hd tl ih => - simp[replaceEntry, cond_eq_if] - simp at h - rcases h with ⟨hd_pw, tl_pw⟩ - split - · simp - constructor - · rename_i hd_k - intro a a_tl - simp at hd_k - rw [← hd_k] - apply hd_pw a a_tl - · exact tl_pw - · rename_i hd_k - simp at hd_k - simp - constructor - · intro a a_mem - by_cases contains_k: containsKey k tl - · rw [replaceEntry_mem_of_mem contains_k] at a_mem - · cases a_mem with - | inl a_tl => - apply hd_pw a (And.left a_tl) - | inr a_k => - simp[a_k, hd_k] - · simp - exact tl_pw - · simp at contains_k - rw [replaceEntry_of_containsKey_eq_false contains_k] at a_mem - apply hd_pw a a_mem - · simp at ih - apply ih - exact tl_pw - - @[simp] theorem isEmpty_replaceEntry [BEq α] {l : List ((a : α) × β a)} {k : α} {v : β k} : (replaceEntry k v l).isEmpty = l.isEmpty := by induction l using assoc_induction · simp - · simp [replaceEntry_cons, cond_eq_if] + · simp only [replaceEntry_cons, cond_eq_if, List.isEmpty_cons] split <;> simp theorem getEntry?_replaceEntry_of_containsKey_eq_false [BEq α] {l : List ((a : α) × β a)} {a k : α} @@ -885,7 +825,7 @@ theorem getEntry?_replaceEntry_of_false [BEq α] [PartialEquivBEq α] {l : List · rw [replaceEntry_cons_of_false h', getEntry?_cons, getEntry?_cons, ih] · rw [replaceEntry_cons_of_true h'] have hk : (k' == a) = false := BEq.neq_of_beq_of_neq h' h - simp [getEntry?_cons_of_false h, getEntry?_cons_of_false hk] + simp only [getEntry?_cons_of_false h, getEntry?_cons_of_false hk] theorem getEntry?_replaceEntry_of_true [BEq α] [PartialEquivBEq α] {l : List ((a : α) × β a)} {a k : α} {v : β k} (hl : containsKey k l = true) (h : k == a) : @@ -1077,8 +1017,8 @@ theorem containsKey_of_mem [BEq α] [ReflBEq α] {l : List ((a : α) × β a)} { theorem DistinctKeys.nil [BEq α] : DistinctKeys ([] : List ((a : α) × β a)) := ⟨by simp⟩ -theorem DistinctKeys.def [BEq α] {l : List ((a : α) × β a)}: DistinctKeys l ↔ List.Pairwise (fun a b => ! (a.1 == b.1)) l := by - have h: ∀ (l' : List ((a : α) × β a)), List.Pairwise (fun a b => (a==b)= false) (keys l') ↔ List.Pairwise (fun a b => !(a.1 == b.1)) l':= by +theorem DistinctKeys.def [BEq α] {l : List ((a : α) × β a)}: DistinctKeys l ↔ List.Pairwise (fun a b => (a.1 == b.1) = false) l := by + have h: ∀ (l' : List ((a : α) × β a)), List.Pairwise (fun a b => (a==b)= false) (keys l') ↔ List.Pairwise (fun a b => (a.1 == b.1) = false) l':= by intro l' induction l' with | nil => simp @@ -1204,9 +1144,9 @@ theorem insertEntry_of_containsKey_eq_false_perm [BEq α] {l : List ((a : α) × (h : containsKey k l = false) : Perm (insertEntry k v l) (⟨k,v⟩::l) := by cases l with | nil => simp - | cons hd tl => simp[insertEntry, h] + | cons hd tl => simp [insertEntry, h] -theorem insertEntry_mem_of_different_keys [BEq α] [LawfulBEq α] {l : List ((a : α) × β a)} {k : α} {v : β k} {a: ((a : α) × β a)} {h: ¬ ⟨k,v⟩ = a} {distinct: List.Pairwise (fun a b => ! (a.1 == b.1) ) l}: ⟨k, v⟩ ∈ insertEntry a.fst a.snd l ↔ ⟨k,v⟩ ∈ l ∧ (k == a.fst) = false := by +theorem insertEntry_mem_of_different_keys [BEq α] [LawfulBEq α] {l : List ((a : α) × β a)} {k : α} {v : β k} {a: ((a : α) × β a)} {h: ¬ ⟨k,v⟩ = a} {distinct: List.Pairwise (fun a b => (a.1 == b.1) = false) l}: ⟨k, v⟩ ∈ insertEntry a.fst a.snd l ↔ ⟨k,v⟩ ∈ l ∧ (k == a.fst) = false := by induction l with | nil => simp @@ -1217,37 +1157,37 @@ theorem insertEntry_mem_of_different_keys [BEq α] [LawfulBEq α] {l : List ((a split · rename_i contains_a rw [replaceEntry_mem_of_mem contains_a (distinct:=distinct)] - simp + simp only [List.mem_cons, beq_iff_eq, beq_eq_false_iff_ne, ne_eq, or_iff_left_iff_imp] intro h contradiction - · simp + · simp only [List.mem_cons, beq_eq_false_iff_ne, ne_eq] constructor · intro h' cases h' with | inl h' => contradiction | inr h' => rename_i contains_a - simp at contains_a + simp only [Bool.not_eq_true] at contains_a rw [containsKey_eq_false_iff] at contains_a cases h' with | inl h' => constructor · left exact h' - · simp at contains_a + · simp only [List.mem_cons, beq_eq_false_iff_ne, ne_eq, forall_eq_or_imp] at contains_a apply Ne.symm rw [← h'] at contains_a - simp at contains_a + simp only at contains_a apply And.left contains_a | inr h' => constructor · right exact h' - · simp at contains_a + · simp only [List.mem_cons, beq_eq_false_iff_ne, ne_eq, forall_eq_or_imp] at contains_a apply Ne.symm rcases contains_a with ⟨_, contains_a⟩ specialize contains_a ⟨k,v⟩ - simp at contains_a + simp only at contains_a apply contains_a h' · intro h' rcases h' with ⟨h', _⟩ @@ -2120,7 +2060,7 @@ theorem insertList_perm_of_perm_first [BEq α] [EquivBEq α] (l1 l2 toInsert: Li induction toInsert generalizing l1 l2 with | nil => simp[insertList, h] | cons hd tl ih => - simp[insertList] + simp only [insertList] apply ih apply insertEntry_of_perm exact distinct @@ -2133,17 +2073,17 @@ theorem containsKey_eq_contains_map_fst [BEq α] [PartialEquivBEq α] (l : List | nil => simp | cons hd tl ih => rw [containsKey_cons, ih] - simp + simp only [List.map_cons, List.contains_cons] rw [BEq.comm] theorem containsKey_insertList [BEq α] [PartialEquivBEq α] (l toInsert : List ((a : α) × β a)) (k: α): containsKey k (List.insertList l toInsert) ↔ containsKey k l ∨ (toInsert.map Sigma.fst).contains k := by induction toInsert generalizing l with - | nil => simp[insertList] + | nil => simp only [insertList, List.map_nil, List.elem_nil, Bool.false_eq_true, or_false] | cons hd tl ih => unfold insertList rw [ih] rw [containsKey_insertEntry] - simp + simp only [Bool.or_eq_true, List.map_cons, List.contains_cons] rw [BEq.comm] conv => left; left; rw [or_comm] rw [or_assoc] @@ -2155,59 +2095,65 @@ theorem containsKey_of_containsKey_insertList [BEq α] [PartialEquivBEq α] -theorem getValueCast?_insertList_start_mem [BEq α] [LawfulBEq α] (l toInsert: List ((a : α) × β a)) {k k': α} {v: β k} {distinct1: List.Pairwise (fun a b => ! a.1 == b.1) toInsert} {distinct2: List.Pairwise (fun a b => ! a.1 == b.1) l} {k_eq: k == k'} {mem: ⟨k,v⟩ ∈ l} {mem': containsKey k toInsert = false}: getValueCast? k' (insertList l toInsert) = some (cast (by congr; apply LawfulBEq.eq_of_beq k_eq) v) := by +theorem getValueCast?_insertList_start_mem [BEq α] [LawfulBEq α] (l toInsert: List ((a : α) × β a)) {k k': α} {v: β k} {distinct1: List.Pairwise (fun a b => (a.1 == b.1) = false) toInsert} {distinct2: List.Pairwise (fun a b => (a.1 == b.1) = false) l} {k_eq: k == k'} {mem: ⟨k,v⟩ ∈ l} {mem': containsKey k toInsert = false}: getValueCast? k' (insertList l toInsert) = some (cast (by congr; apply LawfulBEq.eq_of_beq k_eq) v) := by induction toInsert generalizing l with | nil => - simp[insertList] + simp only [insertList] apply getValueCast?_mem · exact k_eq · exact mem · exact distinct2 | cons hd tl ih => - simp [insertList] + simp only [insertList] apply ih · exact List.Pairwise.of_cons distinct1 · simp only[insertEntry, cond_eq_if] split - · apply replaceEntry_distinct_keys distinct2 - · simp + · rw [← DistinctKeys.def] + apply DistinctKeys.replaceEntry + rw [DistinctKeys.def] + exact distinct2 + · simp only [beq_eq_false_iff_ne, ne_eq, pairwise_cons] constructor · rename_i h - simp at h + simp only [Bool.not_eq_true] at h rw [containsKey_eq_false_iff] at h - simp at h + simp only [beq_eq_false_iff_ne, ne_eq] at h apply h - · simp at distinct2 + · simp only [beq_eq_false_iff_ne, ne_eq] at distinct2 apply distinct2 · rw[containsKey_cons_eq_false] at mem' rw [insertEntry_mem_of_different_keys] · constructor · apply mem - · simp - simp at mem' + · simp only [beq_eq_false_iff_ne, ne_eq] + simp only [beq_eq_false_iff_ne, ne_eq] at mem' apply Ne.symm apply And.left mem' · false_or_by_contra rename_i h rw [← h] at mem' - simp at mem' + simp only [beq_self_eq_true, Bool.true_eq_false, false_and] at mem' · apply distinct2 - · simp[containsKey] at mem' + · simp only [containsKey, Bool.or_eq_false_iff, beq_eq_false_iff_ne, ne_eq] at mem' apply And.right mem' -theorem getValueCast?_insertList_toInsert_mem [BEq α] [LawfulBEq α] (l toInsert: List ((a : α) × β a)) {k k': α} {k_eq: k == k'} {v: β k} {distinct: List.Pairwise (fun a b => ! a.1 == b.1) toInsert} {distinct2: DistinctKeys l} {mem: ⟨k,v⟩ ∈ toInsert}: getValueCast? k' (insertList l toInsert) = some (cast (by congr;apply LawfulBEq.eq_of_beq k_eq) v) := by +theorem getValueCast?_insertList_toInsert_mem [BEq α] [LawfulBEq α] (l toInsert: List ((a : α) × β a)) {k k': α} {k_eq: k == k'} {v: β k} {distinct: List.Pairwise (fun a b => (a.1 == b.1) = false) toInsert} {distinct2: DistinctKeys l} {mem: ⟨k,v⟩ ∈ toInsert}: getValueCast? k' (insertList l toInsert) = some (cast (by congr;apply LawfulBEq.eq_of_beq k_eq) v) := by rw [DistinctKeys.def] at distinct2 induction toInsert generalizing l with | nil => simp at mem | cons hd tl ih => - simp[insertList, insertEntry, cond_eq_if] + simp only [insertList, insertEntry, cond_eq_if] split - · simp at mem + · simp only [List.mem_cons] at mem cases mem with | inl mem => apply getValueCast?_insertList_start_mem · exact List.Pairwise.of_cons distinct - · apply replaceEntry_distinct_keys distinct2 + · rw [← DistinctKeys.def] + apply DistinctKeys.replaceEntry + rw [DistinctKeys.def] + exact distinct2 · exact k_eq · rw[replaceEntry_mem_of_mem] · right @@ -2216,32 +2162,35 @@ theorem getValueCast?_insertList_toInsert_mem [BEq α] [LawfulBEq α] (l toInser exact h · exact distinct2 · have k_hd: k = hd.fst := by - simp[← mem] + simp only [← mem] rw [k_hd, containsKey_eq_false_iff] - simp - simp at distinct + simp only [beq_eq_false_iff_ne, ne_eq] + simp only [beq_eq_false_iff_ne, ne_eq, pairwise_cons] at distinct apply And.left distinct | inr mem => apply ih - · simp - simp at distinct + · simp only [beq_eq_false_iff_ne, ne_eq] + simp only [beq_eq_false_iff_ne, ne_eq, pairwise_cons] at distinct apply And.right distinct - · apply replaceEntry_distinct_keys distinct2 + · rw [← DistinctKeys.def] + apply DistinctKeys.replaceEntry + rw [DistinctKeys.def] + exact distinct2 · exact mem · rename_i contains_hd by_cases hd_k': (hd.fst ==k') = true · apply getValueCast?_insertList_start_mem · exact List.Pairwise.of_cons distinct - · simp - simp at contains_hd + · simp only [beq_eq_false_iff_ne, ne_eq, pairwise_cons] + simp only [Bool.not_eq_true] at contains_hd rw [containsKey_eq_false_iff] at contains_hd - simp at contains_hd + simp only [beq_eq_false_iff_ne, ne_eq] at contains_hd constructor · apply contains_hd - · simp at distinct2 + · simp only [beq_eq_false_iff_ne, ne_eq] at distinct2 apply distinct2 · exact k_eq - · simp at mem + · simp only [List.mem_cons] at mem cases mem with | inl mem => simp [mem] @@ -2249,70 +2198,72 @@ theorem getValueCast?_insertList_toInsert_mem [BEq α] [LawfulBEq α] (l toInser simp at distinct exfalso apply And.left distinct ⟨k,v⟩ mem - simp - simp at hd_k' - simp at k_eq - simp[k_eq, hd_k'] - · simp at contains_hd - simp at hd_k' - simp at k_eq + simp only + simp only [beq_iff_eq] at hd_k' + simp only [beq_iff_eq] at k_eq + simp only [hd_k', k_eq] + · simp only [Bool.not_eq_true] at contains_hd + simp only [beq_iff_eq] at hd_k' + simp only [beq_iff_eq] at k_eq rw [k_eq, ← hd_k'] rw[containsKey_eq_false_iff] - simp at distinct - simp + simp only [beq_eq_false_iff_ne, ne_eq, pairwise_cons] at distinct + simp only [beq_eq_false_iff_ne, ne_eq] apply And.left distinct · apply ih · exact List.Pairwise.of_cons distinct - · simp at mem + · simp only [List.mem_cons] at mem cases mem with | inl mem => - simp at hd_k' + simp only [beq_iff_eq] at hd_k' exfalso apply hd_k' rw [← mem] - simp at k_eq + simp only [beq_iff_eq] at k_eq simp [k_eq] | inr mem => - simp - simp at contains_hd + simp only [beq_eq_false_iff_ne, ne_eq, pairwise_cons] + simp only [Bool.not_eq_true] at contains_hd rw [containsKey_eq_false_iff] at contains_hd - simp at contains_hd + simp only [beq_eq_false_iff_ne, ne_eq] at contains_hd constructor · exact contains_hd - · simp at distinct2 + · simp only [beq_eq_false_iff_ne, ne_eq] at distinct2 apply distinct2 - · simp at mem + · simp only [List.mem_cons] at mem cases mem with | inl mem => - simp [← mem] at hd_k' - simp at k_eq + simp only [← mem, beq_iff_eq] at hd_k' + simp only [beq_iff_eq] at k_eq contradiction | inr mem => exact mem theorem getValueCast?_insertList_not_toInsert_mem [BEq α] [LawfulBEq α] (l toInsert: List ((a : α) × β a)) {k k': α} {k_eq: k == k'} {not_mem: ¬ k ∈ List.map Sigma.fst toInsert} {distinct: DistinctKeys l}: getValueCast? k' (insertList l toInsert) = getValueCast? k' l := by - simp at k_eq + simp only [beq_iff_eq] at k_eq rw [DistinctKeys.def] at distinct induction toInsert generalizing l with - | nil => simp[insertList] + | nil => simp only [insertList] | cons hd tl ih => - simp [insertList] + simp only [insertList] rw [ih] · rw [getValueCast?_insertEntry] split · rename_i h - simp at h - simp [k_eq, h] at not_mem + simp only [beq_iff_eq] at h + simp only [List.map_cons, h, k_eq, List.mem_cons, List.mem_map, true_or, + not_true_eq_false] at not_mem · rfl - · simp at not_mem - simp + · simp only [List.map_cons, List.mem_cons, List.mem_map, not_or, not_exists, + not_and] at not_mem + simp only [List.mem_map, not_exists, not_and] apply And.right not_mem · rw [← DistinctKeys.def] apply DistinctKeys.insertEntry rw [DistinctKeys.def] exact distinct -theorem getValueCast!_insertList_toInsert_mem [BEq α] [LawfulBEq α] (l toInsert: List ((a : α) × β a)) {k k': α} [Inhabited (β k')] {k_eq: k == k'} {v: β k} {distinct: List.Pairwise (fun a b => ! a.1 == b.1) toInsert} {distinct2: DistinctKeys l} {mem: ⟨k,v⟩ ∈ toInsert}: getValueCast! k' (insertList l toInsert) = (cast (by congr;apply LawfulBEq.eq_of_beq k_eq) v) := by +theorem getValueCast!_insertList_toInsert_mem [BEq α] [LawfulBEq α] (l toInsert: List ((a : α) × β a)) {k k': α} [Inhabited (β k')] {k_eq: k == k'} {v: β k} {distinct:List.Pairwise (fun a b => (a.1 == b.1) = false) toInsert} {distinct2: DistinctKeys l} {mem: ⟨k,v⟩ ∈ toInsert}: getValueCast! k' (insertList l toInsert) = (cast (by congr;apply LawfulBEq.eq_of_beq k_eq) v) := by rw [getValueCast!_eq_getValueCast?, getValueCast?_insertList_toInsert_mem (mem:= mem) (l:=l) (toInsert:=toInsert)] · simp · exact k_eq @@ -2323,7 +2274,7 @@ theorem getValueCast!_insertList_not_toInsert_mem [BEq α] [LawfulBEq α] (l toI rw [getValueCast!_eq_getValueCast?,getValueCast!_eq_getValueCast? ,getValueCast?_insertList_not_toInsert_mem (not_mem:=not_mem) (distinct:=distinct)] exact k_eq -theorem getValueCastD_insertList_toInsert_mem [BEq α] [LawfulBEq α] (l toInsert: List ((a : α) × β a)) {k k': α} {k_eq: k == k'} {v: β k} {fallback: β k'} {distinct: List.Pairwise (fun a b => ! a.1 == b.1) toInsert} {distinct2: DistinctKeys l} {mem: ⟨k,v⟩ ∈ toInsert}: getValueCastD k' (insertList l toInsert) fallback = (cast (by congr;apply LawfulBEq.eq_of_beq k_eq) v) := by +theorem getValueCastD_insertList_toInsert_mem [BEq α] [LawfulBEq α] (l toInsert: List ((a : α) × β a)) {k k': α} {k_eq: k == k'} {v: β k} {fallback: β k'} {distinct: List.Pairwise (fun a b => (a.1 == b.1) = false) toInsert} {distinct2: DistinctKeys l} {mem: ⟨k,v⟩ ∈ toInsert}: getValueCastD k' (insertList l toInsert) fallback = (cast (by congr;apply LawfulBEq.eq_of_beq k_eq) v) := by rw [getValueCastD_eq_getValueCast?, getValueCast?_insertList_toInsert_mem (mem:= mem) (l:=l) (toInsert:=toInsert)] · simp · exact k_eq @@ -2334,17 +2285,22 @@ theorem getValueCastD_insertList_not_toInsert_mem [BEq α] [LawfulBEq α] (l toI rw [getValueCastD_eq_getValueCast?, getValueCastD_eq_getValueCast?, getValueCast?_insertList_not_toInsert_mem (not_mem:=not_mem) (distinct:=distinct)] exact k_eq -theorem contains_insertList_of_mem [BEq α] [LawfulBEq α] (l toInsert: List ((a : α) × β a)) (k k': α) (v: β k) (k_eq: k == k') (mem: ⟨k,v⟩ ∈ toInsert): containsKey k' (insertList l toInsert) = true := by +theorem contains_insertList_of_mem [BEq α] [PartialEquivBEq α] (l toInsert: List ((a : α) × β a)) (k k': α) (v: β k) (k_eq: k == k') (mem: ⟨k,v⟩ ∈ toInsert): containsKey k' (insertList l toInsert) = true := by rw [containsKey_insertList] right - simp at * - exists ⟨k,v⟩ + simp [List.contains_eq_any_beq] + exists k + constructor + · exists ⟨k,v⟩ + · apply PartialEquivBEq.symm + exact k_eq -theorem contains_insertList_of_contains_first [BEq α] [LawfulBEq α] {l toInsert: List ((a : α) × β a)} {k k': α} (k_eq: k == k'): containsKey k l = true → containsKey k' (insertList l toInsert) = true:= by +theorem contains_insertList_of_contains_first [BEq α] [PartialEquivBEq α] {l toInsert: List ((a : α) × β a)} {k k': α} (k_eq: k == k'): containsKey k l = true → containsKey k' (insertList l toInsert) = true:= by intro h rw [containsKey_insertList] - simp at * - simp [← k_eq, h] + left + rw [containsKey_congr (PartialEquivBEq.symm k_eq)] + exact h instance Inhabited_of_containsKey [BEq α] [LawfulBEq α] (l: List ((a : α) × β a)) (k k': α) (k_eq: k == k') (mem: containsKey k l = true): Inhabited (β k') := ⟨ @@ -2353,23 +2309,23 @@ instance Inhabited_of_containsKey [BEq α] [LawfulBEq α] (l: List ((a : α) × (by simp at mem) | .cons hd tl => if h:hd.fst == k' - then cast (by congr; simp at h; apply h) hd.snd + then cast (by congr; simp only [beq_iff_eq] at h; apply h) hd.snd else have mem': containsKey k tl = true := by - simp[containsKey] at mem + simp only [containsKey, Bool.or_eq_true, beq_iff_eq] at mem cases mem with - | inl h' => simp at h; simp at k_eq; rw [k_eq] at h'; contradiction + | inl h' => simp only [beq_iff_eq] at h; simp only [beq_iff_eq] at k_eq; rw [k_eq] at h'; contradiction | inr h' => exact h' (Inhabited_of_containsKey tl k k' k_eq mem').default ⟩ -theorem getValueCast_insertList_toInsert_mem [BEq α] [LawfulBEq α] (l toInsert: List ((a : α) × β a)) {k k': α} {v: β k} {k_eq: k == k'} {distinct: List.Pairwise (fun a b => ! a.1 == b.1) toInsert} {distinct2: DistinctKeys l} {mem: ⟨k,v⟩ ∈ toInsert}: getValueCast k' (insertList l toInsert) (contains_insertList_of_mem l toInsert k k' v k_eq mem) = cast (by congr; apply LawfulBEq.eq_of_beq k_eq) v := by +theorem getValueCast_insertList_toInsert_mem [BEq α] [LawfulBEq α] (l toInsert: List ((a : α) × β a)) {k k': α} {v: β k} {k_eq: k == k'} {distinct: List.Pairwise (fun a b => (a.1 == b.1) = false) toInsert} {distinct2: DistinctKeys l} {mem: ⟨k,v⟩ ∈ toInsert}: getValueCast k' (insertList l toInsert) (contains_insertList_of_mem l toInsert k k' v k_eq mem) = cast (by congr; apply LawfulBEq.eq_of_beq k_eq) v := by have inh: Inhabited (β k') := by apply Inhabited_of_containsKey toInsert k k' k_eq rw [containsKey_eq_true_iff_exists_mem] exists ⟨k,v⟩ - simp[mem] + simp only [mem, beq_self_eq_true, and_self] rw [getValueCast_eq_getValueCast!, getValueCast!_insertList_toInsert_mem] · exact k_eq · exact distinct @@ -2397,29 +2353,34 @@ theorem getKey?_insertList [BEq α] [PartialEquivBEq α] {l toInsert : List ((a rw [ih] rw [getKey?_insertEntry] rw [containsKey_cons] - simp + simp only [Bool.or_eq_true, List.map_cons, List.reverse_cons, List.find?_append, + List.find?_singleton] cases eq : containsKey k tl with | true => - simp + simp only [↓reduceIte, or_true] rw [containsKey_eq_contains_map_fst, List.contains_iff_exists_mem_beq] at eq rw [Option.or_of_isSome] - simp + simp only [List.find?_isSome, List.mem_reverse, List.mem_map] rcases eq with ⟨a, a_mem, a_eq⟩ - simp at a_mem + simp only [List.mem_map] at a_mem rcases a_mem with ⟨pair, pair_mem, pair_fst_a⟩ exists a constructor . exists pair . exact a_eq | false => - simp + simp only [Bool.false_eq_true, ↓reduceIte, or_false] cases eq2 : hd.fst == k with | true => - simp + simp only [↓reduceIte] rw [containsKey_eq_contains_map_fst, ← Bool.not_eq_true, List.contains_iff_exists_mem_beq] at eq - simp at eq - have : (tl.map Sigma.fst).reverse.find? (fun a => k == a) = none := by simp; exact eq - simp [this] + simp only [List.mem_map, not_exists, not_and, Bool.not_eq_true, forall_exists_index, + and_imp, forall_apply_eq_imp_iff₂] at eq + have : (tl.map Sigma.fst).reverse.find? (fun a => k == a) = none := by + simp only [List.find?_eq_none, List.mem_reverse, List.mem_map, Bool.not_eq_true, + forall_exists_index, and_imp, forall_apply_eq_imp_iff₂] + exact eq + simp only [this, Option.none_or, Option.some_eq_ite_none_right, and_true] rw [BEq.comm] exact eq2 | false => simp @@ -2437,7 +2398,8 @@ theorem getKey?_insertList_lawful [BEq α] [LawfulBEq α] {l toInsert : List ((a rw [containsKey_cons] cases containsKey k tl with | true => simp - | false => simp; cases eq : hd.fst == k <;> simp; apply LawfulBEq.eq_of_beq; assumption + | false => simp only [Bool.false_eq_true, ↓reduceIte, beq_iff_eq, Bool.or_false]; cases eq : hd.fst == k <;> simp only [Bool.false_eq_true, + ↓reduceIte, Bool.or_self]; apply LawfulBEq.eq_of_beq; assumption theorem getKey_insertList [BEq α] [PartialEquivBEq α] {l toInsert : List ((a : α) × β a)} {k : α} {h₁} : @@ -2504,42 +2466,43 @@ fallback : α} : . simp . rw [getKeyD_eq_getKey?] -theorem insertList_perm [BEq α] [ReflBEq α] [PartialEquivBEq α] (l toInsert: List ((a : α) × β a)) (distinct_l: DistinctKeys l) (distinct_toInsert: DistinctKeys toInsert) (distinct_both: ∀ (a:α), ¬ (containsKey a l ∧ containsKey a toInsert)): +theorem insertList_perm [BEq α] [ReflBEq α] [PartialEquivBEq α] (l toInsert: List ((a : α) × β a)) (distinct_l: DistinctKeys l) (distinct_toInsert: List.Pairwise (fun a b => (a.1 == b.1) = false) toInsert) (distinct_both: ∀ (a:α), ¬ (containsKey a l ∧ containsKey a toInsert)): Perm (insertList l toInsert) (l++toInsert) := by + rw [← DistinctKeys.def] at distinct_toInsert induction toInsert generalizing l with - | nil => simp[insertList] + | nil => simp only [insertList, List.append_nil, Perm.refl] | cons hd tl ih => - simp[insertList] + simp only [insertList] specialize ih (insertEntry hd.fst hd.snd l) apply Perm.trans · apply ih - · simp[insertEntry] + · simp only [insertEntry] specialize distinct_both hd.1 - simp[containsKey] at distinct_both - simp [distinct_both] + simp only [containsKey, BEq.refl, Bool.true_or, and_true, Bool.not_eq_true] at distinct_both + simp only [distinct_both, cond_false, distinctKeys_cons_iff, and_true] exact distinct_l · apply DistinctKeys.tail distinct_toInsert · intro a specialize distinct_both a - simp[containsKey] at distinct_both - simp[containsKey] + simp only [containsKey, Bool.or_eq_true, not_and, not_or, Bool.not_eq_true] at distinct_both + simp only [containsKey_insertEntry, Bool.or_eq_true, not_and, Bool.not_eq_true] by_cases hd_a: hd.fst == a · rw [distinctKeys_cons_iff] at distinct_toInsert have contains_a_tl: containsKey a tl = false := by false_or_by_contra rename_i h - simp at h - rw [containsKey_of_eq (PartialEquivBEq.symm hd_a)] at h - simp [h] at distinct_toInsert + simp only [Bool.not_eq_false] at h + rw [containsKey_congr (PartialEquivBEq.symm hd_a)] at h + simp only [h, Bool.true_eq_false, and_false] at distinct_toInsert simp [contains_a_tl] - · simp[hd_a] - simp at hd_a - simp [hd_a] at distinct_both + · simp only [hd_a, Bool.false_eq_true, false_or] + simp only [Bool.not_eq_true] at hd_a + simp only [hd_a, true_and] at distinct_both apply distinct_both - · simp[insertEntry] + · simp only [insertEntry] specialize distinct_both hd.1 - simp[containsKey] at distinct_both - simp [distinct_both] + simp only [containsKey, BEq.refl, Bool.true_or, and_true, Bool.not_eq_true] at distinct_both + simp only [distinct_both, cond_false, List.cons_append] have h_hd: hd = ⟨hd.fst, hd.snd⟩:= rfl rw [← h_hd] apply Perm.symm @@ -2548,14 +2511,14 @@ theorem insertList_perm [BEq α] [ReflBEq α] [PartialEquivBEq α] (l toInsert: theorem insertList_not_isEmpty_if_start_not_isEmpty [BEq α]{l toInsert: List ((a : α) × β a)} {h:l.isEmpty = false}: (List.insertList l toInsert).isEmpty = false := by induction toInsert generalizing l with - | nil => simp[insertList, h] + | nil => simp [insertList, h] | cons hd tl ih => - simp[insertList, ih] + simp [insertList, ih] theorem insertList_isEmpty [BEq α]{l toInsert: List ((a : α) × β a)}: (List.insertList l toInsert).isEmpty ↔ l.isEmpty ∧ toInsert.isEmpty := by induction toInsert with - | nil => simp[insertList] + | nil => simp [insertList] | cons hd tl ih => simp only [insertList, List.isEmpty_cons, Bool.false_eq_true, and_false, iff_false] diff --git a/src/Std/Data/DHashMap/Internal/Model.lean b/src/Std/Data/DHashMap/Internal/Model.lean index c7f2b38264f0..0f0ae3589b0d 100644 --- a/src/Std/Data/DHashMap/Internal/Model.lean +++ b/src/Std/Data/DHashMap/Internal/Model.lean @@ -462,11 +462,11 @@ theorem filter_eq_filterₘ (m : Raw₀ α β) (f : (a : α) → β a → Bool) m.filter f = m.filterₘ f := rfl theorem insertList_eq_insertListₘ [BEq α] [Hashable α](m : Raw₀ α β) (l: List ((a : α) × β a)): insertList m l = insertListₘ m l := by - simp [insertList, Id.run] + simp only [insertList] induction l generalizing m with | nil => simp[insertListₘ] | cons hd tl ih => - simp [insertListₘ] + simp only [List.foldl_cons,insertListₘ] rw [ih] section diff --git a/src/Std/Data/DHashMap/Internal/Raw.lean b/src/Std/Data/DHashMap/Internal/Raw.lean index 30703cea712b..2aa64a6a1371 100644 --- a/src/Std/Data/DHashMap/Internal/Raw.lean +++ b/src/Std/Data/DHashMap/Internal/Raw.lean @@ -207,14 +207,6 @@ theorem insertMany_val [BEq α][Hashable α] {m: Raw₀ α β} {ρ : Type w} [Fo m.val.insertMany l = m.insertMany l := by simp[Raw.insertMany, m.2] -theorem insertList_eq [BEq α] [Hashable α] {m : Raw α β} (h : m.WF) {l: List ((a : α) × β a)}: - m.insertList l = Raw₀.insertList ⟨m, h.size_buckets_pos⟩ l := by - simp[Raw.insertList, h.size_buckets_pos] - -theorem insertList_val [BEq α] [Hashable α] {m : Raw₀ α β} {l: List ((a : α) × β a)}: - m.val.insertList l = m.insertList l := by - simp[Raw.insertList, m.2] - section variable {β : Type v} diff --git a/src/Std/Data/DHashMap/Internal/RawLemmas.lean b/src/Std/Data/DHashMap/Internal/RawLemmas.lean index a90ff8b436d1..7a0ff34c3957 100644 --- a/src/Std/Data/DHashMap/Internal/RawLemmas.lean +++ b/src/Std/Data/DHashMap/Internal/RawLemmas.lean @@ -830,7 +830,7 @@ theorem contains_keys [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {k : α} : simp_to_model using List.containsKey_eq_keys_contains.symm @[simp] -theorem mem_keys [LawfulBEq α] [LawfulHashable α] (h : m.1.WF) {k : α} : +theorem mem_keys [LawfulBEq α] (h : m.1.WF) {k : α} : k ∈ m.1.keys ↔ m.contains k := by simp_to_model rw [List.containsKey_eq_keys_contains] @@ -839,13 +839,9 @@ theorem distinct_keys [EquivBEq α] [LawfulHashable α] (h : m.1.WF) : m.1.keys.Pairwise (fun a b => (a == b) = false) := by simp_to_model using (Raw.WF.out h).distinct.distinct -theorem insertList_eq_foldl - (m : Raw₀ α β) (l : List ((a : α) × β a)) : - insertList m l = l.foldl (init := m) fun m' p => m'.insert p.1 p.2 := by - simp [insertList, Id.run] - -theorem insertMany_val (m : Raw₀ α β) (l : List ((a : α) × β a)) : - (insertMany m l).val = l.foldl (init := m) fun m' p => m'.insert p.1 p.2 := by +@[simp] +theorem insertMany_eq_insertList (m : Raw₀ α β) (l : List ((a : α) × β a)) : + (insertMany m l).val = insertList m l := by simp only [insertMany, Id.run, Id.pure_eq, Id.bind_eq, List.forIn_yield_eq_foldl] suffices ∀ (t : { m' // ∀ (P : Raw₀ α β → Prop), (∀ {m'' : Raw₀ α β} {a : α} {b : β a}, P m'' → P (m''.insert a b)) → P m → P m' }), @@ -858,51 +854,29 @@ theorem insertMany_val (m : Raw₀ α β) (l : List ((a : α) × β a)) : simp rw [ih] -@[simp] -theorem insertMany_eq_insertList (m : Raw₀ α β) (l : List ((a : α) × β a)) : - (insertMany m l).val = insertList m l := by - rw [insertList_eq_foldl, insertMany_val] - @[simp] theorem insertList_nil: m.insertList [] = m := by - simp[insertList, Id.run] + simp [insertList, Id.run] @[simp] theorem insertList_singleton [EquivBEq α] [LawfulHashable α] {k: α} {v: β k}: m.insertList [⟨k,v⟩] = m.insert k v := by - simp[insertList, Id.run] + simp [insertList, Id.run] @[simp] theorem insertList_cons {l: List ((a:α) × (β a))} {k: α} {v: β k}: m.insertList (⟨k,v⟩::l) = (m.insert k v).insertList l := by cases l with - | nil => simp[insertList] - | cons hd tl => simp[insertList] - -theorem insertList_insertList {l1 l2: List ((a:α) × (β a))}: - (m.insertList l1).insertList l2 = m.insertList (l1++l2) := by - simp[insertList_eq_insertListₘ] - induction l1 generalizing m with - | nil => simp[insertListₘ] - | cons hd tl ih => - simp[insertListₘ] - apply ih - -theorem insertList_insert {l: List ((a:α) × (β a))} {k: α} {v: β k}: - (m.insertList l).insert k v = m.insertList (l ++ [⟨k,v⟩]) := by - simp [insertList_eq_insertListₘ] - induction l generalizing m with - | nil => simp[insertListₘ] - | cons hd tl ih => - simp[insertListₘ] - rw [ih] + | nil => simp [insertList] + | cons hd tl => simp [insertList] + theorem contains_insertList [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l: List ((a:α) × (β a))} {k: α}: (m.insertList l).contains k ↔ m.contains k ∨ (l.map Sigma.fst).contains k := by simp_to_model using List.containsKey_insertList -theorem contains_insertList_of_mem_list [LawfulBEq α] [LawfulHashable α] (h : m.1.WF) {l: List ((a:α) × (β a))} {k k': α} (k_eq: k == k') {v: β k}: +theorem contains_insertList_of_mem_list [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l: List ((a:α) × (β a))} {k k': α} (k_eq: k == k') {v: β k}: ⟨k,v⟩ ∈ l → (m.insertList l).contains k' := by simp_to_model using contains_insertList_of_mem @@ -911,37 +885,37 @@ theorem contains_of_contains_insertList [EquivBEq α] [LawfulHashable α] (h : m (m.insertList l).contains k → (l.map Sigma.fst).contains k = false → m.contains k := by simp_to_model using List.containsKey_of_containsKey_insertList -theorem contains_insertList_of_contains_map [LawfulBEq α] [LawfulHashable α] (h : m.1.WF) {l: List ((a:α) × (β a))} {k k': α} (k_eq: k == k'): +theorem contains_insertList_of_contains_map [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l: List ((a:α) × (β a))} {k k': α} (k_eq: k == k'): m.contains k = true → (m.insertList l).contains k' := by simp_to_model using contains_insertList_of_contains_first -theorem get?_insertList_mem_list [LawfulBEq α] [LawfulHashable α] {l: List ((a:α) × (β a))} {k k': α} {v: β k} {distinct: List.Pairwise (fun a b => ! a.1 == b.1) l} {k_eq: k == k'} {mem: ⟨k,v⟩ ∈ l} (h: m.1.WF): +theorem get?_insertList_mem_list [LawfulBEq α] {l: List ((a:α) × (β a))} {k k': α} {v: β k} {distinct: List.Pairwise (fun a b => (a.1 == b.1) = false) l} {k_eq: k == k'} {mem: ⟨k,v⟩ ∈ l} (h: m.1.WF): (m.insertList l).get? k' = some (cast (by congr; apply LawfulBEq.eq_of_beq k_eq) v) := by simp_to_model using getValueCast?_insertList_toInsert_mem -theorem get?_insertList_not_mem_list [LawfulBEq α] [LawfulHashable α] {l: List ((a:α) × (β a))} {k k': α} {mem: ¬ k ∈ (l.map (Sigma.fst))} {k_eq: k == k'} (h: m.1.WF): +theorem get?_insertList_not_mem_list [LawfulBEq α] {l: List ((a:α) × (β a))} {k k': α} {mem: ¬ k ∈ (l.map (Sigma.fst))} {k_eq: k == k'} (h: m.1.WF): (m.insertList l).get? k' = m.get? k' := by simp_to_model using getValueCast?_insertList_not_toInsert_mem -theorem get!_insertList_mem_list [LawfulBEq α] [LawfulHashable α] {l: List ((a:α) × (β a))} {k k': α} {v: β k}[Inhabited (β k')] {distinct: List.Pairwise (fun a b => ! a.1 == b.1) l} {k_eq: k == k'} {mem: ⟨k,v⟩ ∈ l} (h: m.1.WF): +theorem get!_insertList_mem_list [LawfulBEq α] {l: List ((a:α) × (β a))} {k k': α} {v: β k}[Inhabited (β k')] {distinct: List.Pairwise (fun a b => (a.1 == b.1) = false) l} {k_eq: k == k'} {mem: ⟨k,v⟩ ∈ l} (h: m.1.WF): (m.insertList l).get! k' = cast (by congr; apply LawfulBEq.eq_of_beq k_eq) v := by simp_to_model using getValueCast!_insertList_toInsert_mem -theorem get!_insertList_not_mem_list [LawfulBEq α] [LawfulHashable α] {l: List ((a:α) × (β a))} {k k': α} [Inhabited (β k')] {mem: ¬ k ∈ (l.map (Sigma.fst))} {k_eq: k == k'} (h: m.1.WF): +theorem get!_insertList_not_mem_list [LawfulBEq α] {l: List ((a:α) × (β a))} {k k': α} [Inhabited (β k')] {mem: ¬ k ∈ (l.map (Sigma.fst))} {k_eq: k == k'} (h: m.1.WF): (m.insertList l).get! k' = m.get! k' := by simp_to_model using getValueCast!_insertList_not_toInsert_mem -theorem getD_insertList_mem_list [LawfulBEq α] [LawfulHashable α] {l: List ((a:α) × (β a))} {k k': α} {v: β k} {fallback: β k'} {distinct: List.Pairwise (fun a b => ! a.1 == b.1) l} {k_eq: k == k'} {mem: ⟨k,v⟩ ∈ l} (h: m.1.WF): +theorem getD_insertList_mem_list [LawfulBEq α] {l: List ((a:α) × (β a))} {k k': α} {v: β k} {fallback: β k'} {distinct: List.Pairwise (fun a b => (a.1 == b.1) = false) l} {k_eq: k == k'} {mem: ⟨k,v⟩ ∈ l} (h: m.1.WF): (m.insertList l).getD k' fallback = cast (by congr; apply LawfulBEq.eq_of_beq k_eq) v := by simp_to_model using getValueCastD_insertList_toInsert_mem -theorem getD_insertList_not_mem_list [LawfulBEq α] [LawfulHashable α] {l: List ((a:α) × (β a))} {k k': α} {fallback: β k'} {mem: ¬ k ∈ (l.map (Sigma.fst))} {k_eq: k == k'} (h: m.1.WF): +theorem getD_insertList_not_mem_list [LawfulBEq α] {l: List ((a:α) × (β a))} {k k': α} {fallback: β k'} {mem: ¬ k ∈ (l.map (Sigma.fst))} {k_eq: k == k'} (h: m.1.WF): (m.insertList l).getD k' fallback = m.getD k' fallback := by simp_to_model using getValueCastD_insertList_not_toInsert_mem -theorem get_insertList_mem_list [LawfulBEq α] [LawfulHashable α] {l: List ((a:α) × (β a))} {k k': α} {v: β k} {distinct: List.Pairwise (fun a b => ! a.1 == b.1) l} {k_eq: k == k'} {mem: ⟨k,v⟩ ∈ l} (h: m.1.WF): +theorem get_insertList_mem_list [LawfulBEq α] {l: List ((a:α) × (β a))} {k k': α} {v: β k} {distinct: List.Pairwise (fun a b => (a.1 == b.1) = false) l} {k_eq: k == k'} {mem: ⟨k,v⟩ ∈ l} (h: m.1.WF): (m.insertList l).get k' (contains_insertList_of_mem_list _ h k_eq mem)= cast (by congr; apply LawfulBEq.eq_of_beq k_eq ) v := by simp_to_model using getValueCast_insertList_toInsert_mem @@ -951,7 +925,7 @@ theorem contains_of_beq [EquivBEq α][LawfulHashable α] {k k': α} (k_eq: k == intro h' apply containsKey_of_beq h' k_eq -theorem get_insertList_not_mem_list [LawfulBEq α] [LawfulHashable α] {l: List ((a:α) × (β a))} {k k': α} {k_eq: k == k'} {mem_list: ¬ k ∈ (l.map (Sigma.fst))} (h: m.1.WF): +theorem get_insertList_not_mem_list [LawfulBEq α] {l: List ((a:α) × (β a))} {k k': α} {k_eq: k == k'} {mem_list: ¬ k ∈ (l.map (Sigma.fst))} (h: m.1.WF): (h': m.contains k = true) → (m.insertList l).get k' (contains_insertList_of_contains_map m h k_eq h') = m.get k' (contains_of_beq m k_eq h h') := by simp_to_model using getValueCast_insertList_not_toInsert_mem @@ -960,7 +934,7 @@ theorem getKey?_insertList [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l : L if (l.map Sigma.fst).contains k then (l.map Sigma.fst).reverse.find? (fun a => k == a) else m.getKey? k := by simp_to_model using List.getKey?_insertList -theorem getKey?_insertList_lawful [LawfulBEq α] [LawfulHashable α] (h : m.1.WF) {l : List ((a:α) × (β a))} {k : α} : +theorem getKey?_insertList_lawful [LawfulBEq α] (h : m.1.WF) {l : List ((a:α) × (β a))} {k : α} : (m.insertList l).getKey? k = if (l.map Sigma.fst).contains k then some k else m.getKey? k := by simp_to_model using List.getKey?_insertList_lawful @@ -972,7 +946,7 @@ theorem getKey_insertList [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l : Li else m.getKey k (contains_of_contains_insertList _ h h₁ (Bool.eq_false_iff.2 h₂)) := by simp_to_model using List.getKey_insertList -theorem getKey_insertList_lawful [LawfulBEq α] [LawfulHashable α] (h : m.1.WF) {l : List ((a:α) × (β a))} {k +theorem getKey_insertList_lawful [LawfulBEq α] (h : m.1.WF) {l : List ((a:α) × (β a))} {k : α} {h₁} : (m.insertList l).getKey k h₁ = if h₂ : (l.map Sigma.fst).contains k then k @@ -985,7 +959,7 @@ theorem getKey!_insertList [EquivBEq α] [LawfulHashable α] [Inhabited α] (h : else m.getKey! k := by simp_to_model using List.getKey!_insertList -theorem getKey!_insertList_lawful [LawfulBEq α] [LawfulHashable α] [Inhabited α] (h : m.1.WF) {l : List ((a:α) × (β a))} {k : α} : +theorem getKey!_insertList_lawful [LawfulBEq α] [Inhabited α] (h : m.1.WF) {l : List ((a:α) × (β a))} {k : α} : (m.insertList l).getKey! k = if (l.map Sigma.fst).contains k then k else m.getKey! k := by @@ -997,13 +971,13 @@ theorem getKeyD_insertList [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l : L else m.getKeyD k fallback := by simp_to_model using List.getKeyD_insertList -theorem getKeyD_insertList_lawful [LawfulBEq α] [LawfulHashable α] (h : m.1.WF) {l : List ((a:α) × (β a))} {k fallback : α} : +theorem getKeyD_insertList_lawful [LawfulBEq α] (h : m.1.WF) {l : List ((a:α) × (β a))} {k fallback : α} : (m.insertList l).getKeyD k fallback = if (l.map Sigma.fst).contains k then k else m.getKeyD k fallback := by simp_to_model using List.getKeyD_insertList_lawful -theorem size_insertList [EquivBEq α] [LawfulHashable α] {l: List ((a:α) × (β a))} {distinct: DistinctKeys l} (h: m.1.WF): +theorem size_insertList [EquivBEq α] [LawfulHashable α] {l: List ((a:α) × (β a))} {distinct: List.Pairwise (fun a b => (a.1 == b.1) = false) l} (h: m.1.WF): (∀ (a:α), ¬ (m.contains a = true ∧ List.containsKey a l = true)) → (m.insertList l).1.size = m.1.size + l.length := by simp_to_model diff --git a/src/Std/Data/DHashMap/Internal/WF.lean b/src/Std/Data/DHashMap/Internal/WF.lean index 09a333723b8f..2b53f475ec8d 100644 --- a/src/Std/Data/DHashMap/Internal/WF.lean +++ b/src/Std/Data/DHashMap/Internal/WF.lean @@ -710,10 +710,10 @@ theorem wfImp_filter [BEq α] [Hashable α] [EquivBEq α] [LawfulHashable α] {m theorem wfImp_insertListₘ [BEq α] [Hashable α] [EquivBEq α][LawfulHashable α] {m : Raw₀ α β} {l: List ((a : α) × β a)} (h : Raw.WFImp m.1): Raw.WFImp (m.insertListₘ l).1 := by induction l generalizing m with | nil => - simp[insertListₘ] + simp only [insertListₘ] exact h | cons hd tl ih => - simp[insertListₘ] + simp only [insertListₘ] apply ih apply wfImp_insert h @@ -723,7 +723,7 @@ induction l generalizing m with | nil => simp[insertListₘ, List.insertList] | cons hd tl ih => - simp[insertListₘ, List.insertList] + simp only [insertListₘ, List.insertList] apply Perm.trans apply ih (wfImp_insert h) apply List.insertList_perm_of_perm_first @@ -757,7 +757,6 @@ theorem WF.out [BEq α] [Hashable α] [i₁ : EquivBEq α] [i₂ : LawfulHashabl · next h => exact Raw₀.wfImp_getThenInsertIfNew? (by apply h) · next h => exact Raw₀.wfImp_filter (by apply h) · next h => exact Raw₀.Const.wfImp_getThenInsertIfNew? (by apply h) - · next h => exact Raw₀.wfImp_insertList (by apply h) end Raw diff --git a/src/Std/Data/DHashMap/Raw.lean b/src/Std/Data/DHashMap/Raw.lean index f71ed2e41e51..dc6a4e0e2250 100644 --- a/src/Std/Data/DHashMap/Raw.lean +++ b/src/Std/Data/DHashMap/Raw.lean @@ -389,15 +389,6 @@ appearance. (Raw₀.insertMany ⟨m, h⟩ l).1 else m -- will never happen for well-formed inputs -/-- -Inserts multiple mappings into the hash map by iterating over the given list and calling -`insert`. If the same key appears multiple times, the last occurrence takes precedence. This is equal to `insertMany`, but allows easier proofs. --/ -@[inline] def insertList [BEq α] [Hashable α] - (m: Raw α β)(l: List ((a : α) × β a)): Raw α β := - if h : 0 < m.buckets.size then - (Raw₀.insertList ⟨m,h⟩ l).1 - else m -- will never happen for well-formed inputs @[inline, inherit_doc Raw.insertMany] def Const.insertMany {β : Type v} [BEq α] [Hashable α] {ρ : Type w} [ForIn Id ρ (α × β)] (m : Raw α (fun _ => β)) (l : ρ) : Raw α (fun _ => β) := @@ -506,8 +497,6 @@ inductive WF : {α : Type u} → {β : α → Type v} → [BEq α] → [Hashable /-- Internal implementation detail of the hash map -/ | constGetThenInsertIfNew?₀ {α β} [BEq α] [Hashable α] {m : Raw α (fun _ => β)} {h a b} : WF m → WF (Raw₀.Const.getThenInsertIfNew? ⟨m, h⟩ a b).2.1 - /-- Internal implementation detail of the hash map -/ - | insertList₀ {α β} [BEq α] [Hashable α] {m : Raw α β} {l: List ((a: α) × (β a))} {h}: WF m → WF (Raw₀.insertList ⟨m,h⟩ l).1 /-- Internal implementation detail of the hash map -/ theorem WF.size_buckets_pos [BEq α] [Hashable α] (m : Raw α β) : WF m → 0 < m.buckets.size @@ -521,7 +510,6 @@ theorem WF.size_buckets_pos [BEq α] [Hashable α] (m : Raw α β) : WF m → 0 | getThenInsertIfNew?₀ _ => (Raw₀.getThenInsertIfNew? ⟨_, _⟩ _ _).2.2 | filter₀ _ => (Raw₀.filter _ ⟨_, _⟩).2 | constGetThenInsertIfNew?₀ _ => (Raw₀.Const.getThenInsertIfNew? ⟨_, _⟩ _ _).2.2 - | insertList₀ _ => (Raw₀.insertList ⟨_,_⟩ _).2 @[simp] theorem WF.empty [BEq α] [Hashable α] {c : Nat} : (Raw.empty c : Raw α β).WF := .empty₀ @@ -565,11 +553,6 @@ theorem WF.insertMany [BEq α] [Hashable α] {ρ : Type w} [ForIn Id ρ ((a : α simpa [Raw.insertMany, h.size_buckets_pos] using (Raw₀.insertMany ⟨m, h.size_buckets_pos⟩ l).2 _ WF.insert₀ h -theorem WF.insertList [BEq α] [Hashable α] {m : Raw α β} - {l : List ((a : α) × (β a))} (h : m.WF): (m.insertList l).WF := by - simpa [Raw.insertList, h.size_buckets_pos] using - .insertList₀ h - theorem WF.Const.insertMany {β : Type v} [BEq α] [Hashable α] {ρ : Type w} [ForIn Id ρ (α × β)] {m : Raw α (fun _ => β)} {l : ρ} (h : m.WF) : (Const.insertMany m l).WF := by simpa [Raw.Const.insertMany, h.size_buckets_pos] using diff --git a/src/Std/Data/DHashMap/RawLemmas.lean b/src/Std/Data/DHashMap/RawLemmas.lean index 0cb2d343f174..9dbb0e3b2d0d 100644 --- a/src/Std/Data/DHashMap/RawLemmas.lean +++ b/src/Std/Data/DHashMap/RawLemmas.lean @@ -39,7 +39,6 @@ private def baseNames : Array Name := ``getThenInsertIfNew?_snd_eq, ``getThenInsertIfNew?_snd_val, ``map_eq, ``map_val, ``filter_eq, ``filter_val, - ``insertList_eq, ``insertList_val, ``erase_eq, ``erase_val, ``filterMap_eq, ``filterMap_val, ``Const.getThenInsertIfNew?_snd_eq, ``Const.getThenInsertIfNew?_snd_val, From 77bc339006c2648a992b5e2d7afc1d3059188a72 Mon Sep 17 00:00:00 2001 From: jt0202 Date: Tue, 26 Nov 2024 20:48:51 +0100 Subject: [PATCH 21/83] Replace insertList with insertMany in lemmas --- src/Std/Data/DHashMap/Internal/Defs.lean | 4 - .../DHashMap/Internal/List/Associative.lean | 25 ++---- src/Std/Data/DHashMap/Internal/Model.lean | 11 ++- src/Std/Data/DHashMap/Internal/RawLemmas.lean | 80 ++++++++----------- src/Std/Data/DHashMap/Internal/WF.lean | 17 ++-- 5 files changed, 55 insertions(+), 82 deletions(-) diff --git a/src/Std/Data/DHashMap/Internal/Defs.lean b/src/Std/Data/DHashMap/Internal/Defs.lean index 6244f30b85d9..e4a03ce53ee8 100644 --- a/src/Std/Data/DHashMap/Internal/Defs.lean +++ b/src/Std/Data/DHashMap/Internal/Defs.lean @@ -352,10 +352,6 @@ where r := ⟨r.1.insert a b, fun _ h hm => h (r.2 _ h hm)⟩ return r -/-- Internal implementation detail of the hash map -/ -def insertList [BEq α] [Hashable α] - (m : Raw₀ α β) (l : List ((a : α) × β a)) : Raw₀ α β := - List.foldl (fun a b => insert a b.1 b.2) m l section variable {β : Type v} diff --git a/src/Std/Data/DHashMap/Internal/List/Associative.lean b/src/Std/Data/DHashMap/Internal/List/Associative.lean index 632a5b5c445f..befe1fbed638 100644 --- a/src/Std/Data/DHashMap/Internal/List/Associative.lean +++ b/src/Std/Data/DHashMap/Internal/List/Associative.lean @@ -1017,21 +1017,10 @@ theorem containsKey_of_mem [BEq α] [ReflBEq α] {l : List ((a : α) × β a)} { theorem DistinctKeys.nil [BEq α] : DistinctKeys ([] : List ((a : α) × β a)) := ⟨by simp⟩ -theorem DistinctKeys.def [BEq α] {l : List ((a : α) × β a)}: DistinctKeys l ↔ List.Pairwise (fun a b => (a.1 == b.1) = false) l := by - have h: ∀ (l' : List ((a : α) × β a)), List.Pairwise (fun a b => (a==b)= false) (keys l') ↔ List.Pairwise (fun a b => (a.1 == b.1) = false) l':= by - intro l' - induction l' with - | nil => simp - | cons hd tl ih=> - simp only [keys_eq_map, List.map_cons, pairwise_cons] - rw [← ih] - simp [keys_eq_map] - rw [← h] - constructor - · intro h - apply h.distinct - · intro h - exact ⟨h⟩ +theorem DistinctKeys.def [BEq α] {l : List ((a : α) × β a)}: DistinctKeys l ↔ List.Pairwise (fun a b => (a.1 == b.1) = false) l := + ⟨fun h => by simpa [keys_eq_map, List.pairwise_map] using h.distinct, + fun h => ⟨by simpa [keys_eq_map, List.pairwise_map] using h⟩⟩ + open List @@ -2076,7 +2065,7 @@ theorem containsKey_eq_contains_map_fst [BEq α] [PartialEquivBEq α] (l : List simp only [List.map_cons, List.contains_cons] rw [BEq.comm] -theorem containsKey_insertList [BEq α] [PartialEquivBEq α] (l toInsert : List ((a : α) × β a)) (k: α): containsKey k (List.insertList l toInsert) ↔ containsKey k l ∨ (toInsert.map Sigma.fst).contains k := by +theorem containsKey_insertList [BEq α] [PartialEquivBEq α] (l toInsert : List ((a : α) × β a)) (k : α): containsKey k (List.insertList l toInsert) ↔ containsKey k l ∨ (toInsert.map Sigma.fst).contains k := by induction toInsert generalizing l with | nil => simp only [insertList, List.map_nil, List.elem_nil, Bool.false_eq_true, or_false] | cons hd tl ih => @@ -2274,14 +2263,14 @@ theorem getValueCast!_insertList_not_toInsert_mem [BEq α] [LawfulBEq α] (l toI rw [getValueCast!_eq_getValueCast?,getValueCast!_eq_getValueCast? ,getValueCast?_insertList_not_toInsert_mem (not_mem:=not_mem) (distinct:=distinct)] exact k_eq -theorem getValueCastD_insertList_toInsert_mem [BEq α] [LawfulBEq α] (l toInsert: List ((a : α) × β a)) {k k': α} {k_eq: k == k'} {v: β k} {fallback: β k'} {distinct: List.Pairwise (fun a b => (a.1 == b.1) = false) toInsert} {distinct2: DistinctKeys l} {mem: ⟨k,v⟩ ∈ toInsert}: getValueCastD k' (insertList l toInsert) fallback = (cast (by congr;apply LawfulBEq.eq_of_beq k_eq) v) := by +theorem getValueCastD_insertList_toInsert_mem [BEq α] [LawfulBEq α] (l toInsert: List ((a : α) × β a)) {k k': α} {k_eq: k == k'} {v: β k} {fallback : β k'} {distinct: List.Pairwise (fun a b => (a.1 == b.1) = false) toInsert} {distinct2: DistinctKeys l} {mem: ⟨k,v⟩ ∈ toInsert}: getValueCastD k' (insertList l toInsert) fallback = (cast (by congr;apply LawfulBEq.eq_of_beq k_eq) v) := by rw [getValueCastD_eq_getValueCast?, getValueCast?_insertList_toInsert_mem (mem:= mem) (l:=l) (toInsert:=toInsert)] · simp · exact k_eq · exact distinct · exact distinct2 -theorem getValueCastD_insertList_not_toInsert_mem [BEq α] [LawfulBEq α] (l toInsert: List ((a : α) × β a)) {k k': α} {fallback: β k'} {k_eq: k == k'} {not_mem: ¬ k ∈ List.map Sigma.fst toInsert} {distinct: DistinctKeys l}: getValueCastD k' (insertList l toInsert) fallback = getValueCastD k' l fallback:= by +theorem getValueCastD_insertList_not_toInsert_mem [BEq α] [LawfulBEq α] (l toInsert: List ((a : α) × β a)) {k k': α} {fallback: β k'} {k_eq: k == k'} {not_mem: ¬ k ∈ List.map Sigma.fst toInsert} {distinct: DistinctKeys l}: getValueCastD k' (insertList l toInsert) fallback = getValueCastD k' l fallback := by rw [getValueCastD_eq_getValueCast?, getValueCastD_eq_getValueCast?, getValueCast?_insertList_not_toInsert_mem (not_mem:=not_mem) (distinct:=distinct)] exact k_eq diff --git a/src/Std/Data/DHashMap/Internal/Model.lean b/src/Std/Data/DHashMap/Internal/Model.lean index 0f0ae3589b0d..3c0dff49372c 100644 --- a/src/Std/Data/DHashMap/Internal/Model.lean +++ b/src/Std/Data/DHashMap/Internal/Model.lean @@ -461,13 +461,18 @@ theorem map_eq_mapₘ (m : Raw₀ α β) (f : (a : α) → β a → δ a) : theorem filter_eq_filterₘ (m : Raw₀ α β) (f : (a : α) → β a → Bool) : m.filter f = m.filterₘ f := rfl -theorem insertList_eq_insertListₘ [BEq α] [Hashable α](m : Raw₀ α β) (l: List ((a : α) × β a)): insertList m l = insertListₘ m l := by - simp only [insertList] +theorem insertMany_eq_insertListₘ [BEq α] [Hashable α](m : Raw₀ α β) (l: List ((a : α) × β a)): insertMany m l = insertListₘ m l := by + simp only [insertMany, Id.run, Id.pure_eq, Id.bind_eq, List.forIn_yield_eq_foldl] + suffices ∀ (t : { m' // ∀ (P : Raw₀ α β → Prop), + (∀ {m'' : Raw₀ α β} {a : α} {b : β a}, P m'' → P (m''.insert a b)) → P m → P m' }), + (List.foldl (fun m' p => ⟨m'.val.insert p.1 p.2, fun P h₁ h₂ => h₁ (m'.2 _ h₁ h₂)⟩) t l).val = + t.val.insertListₘ l from this _ + intro t induction l generalizing m with | nil => simp[insertListₘ] | cons hd tl ih => simp only [List.foldl_cons,insertListₘ] - rw [ih] + apply ih section diff --git a/src/Std/Data/DHashMap/Internal/RawLemmas.lean b/src/Std/Data/DHashMap/Internal/RawLemmas.lean index 7a0ff34c3957..1335fb879c18 100644 --- a/src/Std/Data/DHashMap/Internal/RawLemmas.lean +++ b/src/Std/Data/DHashMap/Internal/RawLemmas.lean @@ -70,7 +70,7 @@ variable [BEq α] [Hashable α] /-- Internal implementation detail of the hash map -/ scoped macro "wf_trivial" : tactic => `(tactic| repeat (first - | apply Raw₀.wfImp_insert | apply Raw₀.wfImp_insertIfNew | apply Raw₀.wfImp_insertList | apply Raw₀.wfImp_erase + | apply Raw₀.wfImp_insert | apply Raw₀.wfImp_insertIfNew | apply Raw₀.wfImp_insertMany | apply Raw₀.wfImp_erase | apply Raw.WF.out | assumption | apply Raw₀.wfImp_empty | apply Raw.WFImp.distinct | apply Raw.WF.empty₀)) @@ -90,7 +90,7 @@ private def queryNames : Array Name := ``Raw.pairwise_keys_iff_pairwise_keys] private def modifyNames : Array Name := - #[``toListModel_insert, ``toListModel_erase, ``toListModel_insertIfNew, ``toListModel_insertList] + #[``toListModel_insert, ``toListModel_erase, ``toListModel_insertIfNew, ``toListModel_insertMany] private def congrNames : MacroM (Array (TSyntax `term)) := do return #[← `(_root_.List.Perm.isEmpty_eq), ← `(containsKey_of_perm), @@ -839,84 +839,71 @@ theorem distinct_keys [EquivBEq α] [LawfulHashable α] (h : m.1.WF) : m.1.keys.Pairwise (fun a b => (a == b) = false) := by simp_to_model using (Raw.WF.out h).distinct.distinct -@[simp] -theorem insertMany_eq_insertList (m : Raw₀ α β) (l : List ((a : α) × β a)) : - (insertMany m l).val = insertList m l := by - simp only [insertMany, Id.run, Id.pure_eq, Id.bind_eq, List.forIn_yield_eq_foldl] - suffices ∀ (t : { m' // ∀ (P : Raw₀ α β → Prop), - (∀ {m'' : Raw₀ α β} {a : α} {b : β a}, P m'' → P (m''.insert a b)) → P m → P m' }), - (List.foldl (fun m' p => ⟨m'.val.insert p.1 p.2, fun P h₁ h₂ => h₁ (m'.2 _ h₁ h₂)⟩) t l).val = - List.foldl (fun m' p => m'.insert p.fst p.snd) t.val l from this _ - intro t - induction l generalizing m with - | nil => simp - | cons h t ih => - simp - rw [ih] @[simp] theorem insertList_nil: - m.insertList [] = m := by - simp [insertList, Id.run] + m.insertMany [] = m := by + simp [insertMany, Id.run] @[simp] theorem insertList_singleton [EquivBEq α] [LawfulHashable α] {k: α} {v: β k}: - m.insertList [⟨k,v⟩] = m.insert k v := by - simp [insertList, Id.run] + m.insertMany [⟨k,v⟩] = m.insert k v := by + simp [insertMany, Id.run] @[simp] theorem insertList_cons {l: List ((a:α) × (β a))} {k: α} {v: β k}: - m.insertList (⟨k,v⟩::l) = (m.insert k v).insertList l := by + (m.insertMany (⟨k,v⟩::l)).1 = ((m.insert k v).insertMany l).1 := by + simp only [insertMany_eq_insertListₘ] cases l with - | nil => simp [insertList] - | cons hd tl => simp [insertList] + | nil => simp [insertListₘ] + | cons hd tl => simp [insertListₘ] theorem contains_insertList [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l: List ((a:α) × (β a))} {k: α}: - (m.insertList l).contains k ↔ m.contains k ∨ (l.map Sigma.fst).contains k := by + (m.insertMany l).1.contains k ↔ m.contains k ∨ (l.map Sigma.fst).contains k := by simp_to_model using List.containsKey_insertList theorem contains_insertList_of_mem_list [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l: List ((a:α) × (β a))} {k k': α} (k_eq: k == k') {v: β k}: - ⟨k,v⟩ ∈ l → (m.insertList l).contains k' := by + ⟨k,v⟩ ∈ l → (m.insertMany l).1.contains k' := by simp_to_model using contains_insertList_of_mem theorem contains_of_contains_insertList [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l: List ((a:α) × (β a))} {k: α} : - (m.insertList l).contains k → (l.map Sigma.fst).contains k = false → m.contains k := by + (m.insertMany l).1.contains k → (l.map Sigma.fst).contains k = false → m.contains k := by simp_to_model using List.containsKey_of_containsKey_insertList theorem contains_insertList_of_contains_map [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l: List ((a:α) × (β a))} {k k': α} (k_eq: k == k'): - m.contains k = true → (m.insertList l).contains k' := by + m.contains k = true → (m.insertMany l).1.contains k' := by simp_to_model using contains_insertList_of_contains_first theorem get?_insertList_mem_list [LawfulBEq α] {l: List ((a:α) × (β a))} {k k': α} {v: β k} {distinct: List.Pairwise (fun a b => (a.1 == b.1) = false) l} {k_eq: k == k'} {mem: ⟨k,v⟩ ∈ l} (h: m.1.WF): - (m.insertList l).get? k' = some (cast (by congr; apply LawfulBEq.eq_of_beq k_eq) v) := by + (m.insertMany l).1.get? k' = some (cast (by congr; apply LawfulBEq.eq_of_beq k_eq) v) := by simp_to_model using getValueCast?_insertList_toInsert_mem theorem get?_insertList_not_mem_list [LawfulBEq α] {l: List ((a:α) × (β a))} {k k': α} {mem: ¬ k ∈ (l.map (Sigma.fst))} {k_eq: k == k'} (h: m.1.WF): - (m.insertList l).get? k' = m.get? k' := by + (m.insertMany l).1.get? k' = m.get? k' := by simp_to_model using getValueCast?_insertList_not_toInsert_mem theorem get!_insertList_mem_list [LawfulBEq α] {l: List ((a:α) × (β a))} {k k': α} {v: β k}[Inhabited (β k')] {distinct: List.Pairwise (fun a b => (a.1 == b.1) = false) l} {k_eq: k == k'} {mem: ⟨k,v⟩ ∈ l} (h: m.1.WF): - (m.insertList l).get! k' = cast (by congr; apply LawfulBEq.eq_of_beq k_eq) v := by + (m.insertMany l).1.get! k' = cast (by congr; apply LawfulBEq.eq_of_beq k_eq) v := by simp_to_model using getValueCast!_insertList_toInsert_mem theorem get!_insertList_not_mem_list [LawfulBEq α] {l: List ((a:α) × (β a))} {k k': α} [Inhabited (β k')] {mem: ¬ k ∈ (l.map (Sigma.fst))} {k_eq: k == k'} (h: m.1.WF): - (m.insertList l).get! k' = m.get! k' := by + (m.insertMany l).1.get! k' = m.get! k' := by simp_to_model using getValueCast!_insertList_not_toInsert_mem theorem getD_insertList_mem_list [LawfulBEq α] {l: List ((a:α) × (β a))} {k k': α} {v: β k} {fallback: β k'} {distinct: List.Pairwise (fun a b => (a.1 == b.1) = false) l} {k_eq: k == k'} {mem: ⟨k,v⟩ ∈ l} (h: m.1.WF): - (m.insertList l).getD k' fallback = cast (by congr; apply LawfulBEq.eq_of_beq k_eq) v := by + (m.insertMany l).1.getD k' fallback = cast (by congr; apply LawfulBEq.eq_of_beq k_eq) v := by simp_to_model using getValueCastD_insertList_toInsert_mem theorem getD_insertList_not_mem_list [LawfulBEq α] {l: List ((a:α) × (β a))} {k k': α} {fallback: β k'} {mem: ¬ k ∈ (l.map (Sigma.fst))} {k_eq: k == k'} (h: m.1.WF): - (m.insertList l).getD k' fallback = m.getD k' fallback := by + (m.insertMany l).1.getD k' fallback = m.getD k' fallback := by simp_to_model using getValueCastD_insertList_not_toInsert_mem theorem get_insertList_mem_list [LawfulBEq α] {l: List ((a:α) × (β a))} {k k': α} {v: β k} {distinct: List.Pairwise (fun a b => (a.1 == b.1) = false) l} {k_eq: k == k'} {mem: ⟨k,v⟩ ∈ l} (h: m.1.WF): - (m.insertList l).get k' (contains_insertList_of_mem_list _ h k_eq mem)= cast (by congr; apply LawfulBEq.eq_of_beq k_eq ) v := by + (m.insertMany l).1.get k' (contains_insertList_of_mem_list _ h k_eq mem)= cast (by congr; apply LawfulBEq.eq_of_beq k_eq ) v := by simp_to_model using getValueCast_insertList_toInsert_mem theorem contains_of_beq [EquivBEq α][LawfulHashable α] {k k': α} (k_eq: k == k') (h: m.1.WF): @@ -926,60 +913,60 @@ theorem contains_of_beq [EquivBEq α][LawfulHashable α] {k k': α} (k_eq: k == apply containsKey_of_beq h' k_eq theorem get_insertList_not_mem_list [LawfulBEq α] {l: List ((a:α) × (β a))} {k k': α} {k_eq: k == k'} {mem_list: ¬ k ∈ (l.map (Sigma.fst))} (h: m.1.WF): - (h': m.contains k = true) → (m.insertList l).get k' (contains_insertList_of_contains_map m h k_eq h') = m.get k' (contains_of_beq m k_eq h h') := by + (h': m.contains k = true) → (m.insertMany l).1.get k' (contains_insertList_of_contains_map m h k_eq h') = m.get k' (contains_of_beq m k_eq h h') := by simp_to_model using getValueCast_insertList_not_toInsert_mem theorem getKey?_insertList [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l : List ((a:α) × (β a))} {k : α} : - (m.insertList l).getKey? k = + (m.insertMany l).1.getKey? k = if (l.map Sigma.fst).contains k then (l.map Sigma.fst).reverse.find? (fun a => k == a) else m.getKey? k := by simp_to_model using List.getKey?_insertList theorem getKey?_insertList_lawful [LawfulBEq α] (h : m.1.WF) {l : List ((a:α) × (β a))} {k : α} : - (m.insertList l).getKey? k = + (m.insertMany l).1.getKey? k = if (l.map Sigma.fst).contains k then some k else m.getKey? k := by simp_to_model using List.getKey?_insertList_lawful theorem getKey_insertList [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l : List ((a:α) × (β a))} {k : α} {h₁} : - (m.insertList l).getKey k h₁ = + (m.insertMany l).1.getKey k h₁ = if h₂ : (l.map Sigma.fst).contains k then ((l.map Sigma.fst).reverse.find? (fun a => k == a)).get (List.find?_isSome_of_map_fst_contains h₂) else m.getKey k (contains_of_contains_insertList _ h h₁ (Bool.eq_false_iff.2 h₂)) := by simp_to_model using List.getKey_insertList theorem getKey_insertList_lawful [LawfulBEq α] (h : m.1.WF) {l : List ((a:α) × (β a))} {k : α} {h₁} : - (m.insertList l).getKey k h₁ = + (m.insertMany l).1.getKey k h₁ = if h₂ : (l.map Sigma.fst).contains k then k else m.getKey k (contains_of_contains_insertList _ h h₁ (Bool.eq_false_iff.2 h₂)) := by simp_to_model using List.getKey_insertList_lawful theorem getKey!_insertList [EquivBEq α] [LawfulHashable α] [Inhabited α] (h : m.1.WF) {l : List ((a:α) × (β a))} {k : α} : - (m.insertList l).getKey! k = + (m.insertMany l).1.getKey! k = if h₂ : (l.map Sigma.fst).contains k then ((l.map Sigma.fst).reverse.find? (fun a => k == a)).get (List.find?_isSome_of_map_fst_contains h₂) else m.getKey! k := by simp_to_model using List.getKey!_insertList theorem getKey!_insertList_lawful [LawfulBEq α] [Inhabited α] (h : m.1.WF) {l : List ((a:α) × (β a))} {k : α} : - (m.insertList l).getKey! k = + (m.insertMany l).1.getKey! k = if (l.map Sigma.fst).contains k then k else m.getKey! k := by simp_to_model using List.getKey!_insertList_lawful theorem getKeyD_insertList [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l : List ((a:α) × (β a))} {k fallback : α} : - (m.insertList l).getKeyD k fallback = + (m.insertMany l).1.getKeyD k fallback = if h₂ : (l.map Sigma.fst).contains k then ((l.map Sigma.fst).reverse.find? (fun a => k == a)).get (List.find?_isSome_of_map_fst_contains h₂) else m.getKeyD k fallback := by simp_to_model using List.getKeyD_insertList theorem getKeyD_insertList_lawful [LawfulBEq α] (h : m.1.WF) {l : List ((a:α) × (β a))} {k fallback : α} : - (m.insertList l).getKeyD k fallback = + (m.insertMany l).1.getKeyD k fallback = if (l.map Sigma.fst).contains k then k else m.getKeyD k fallback := by simp_to_model using List.getKeyD_insertList_lawful theorem size_insertList [EquivBEq α] [LawfulHashable α] {l: List ((a:α) × (β a))} {distinct: List.Pairwise (fun a b => (a.1 == b.1) = false) l} (h: m.1.WF): (∀ (a:α), ¬ (m.contains a = true ∧ List.containsKey a l = true)) → - (m.insertList l).1.size = m.1.size + l.length := by + (m.insertMany l).1.1.size = m.1.size + l.length := by simp_to_model rw [← List.length_append] intro distinct' @@ -990,12 +977,13 @@ theorem size_insertList [EquivBEq α] [LawfulHashable α] {l: List ((a:α) × ( . apply distinct' theorem insertList_notEmpty_if_m_notEmpty [EquivBEq α] [LawfulHashable α] {l: List ((a:α) × (β a))}(h: m.1.WF): - (m.1.isEmpty = false) → (m.insertList l).1.isEmpty = false := by + (m.1.isEmpty = false) → (m.insertMany l).1.1.isEmpty = false := by simp_to_model using List.insertList_not_isEmpty_if_start_not_isEmpty theorem insertList_isEmpty [EquivBEq α] [LawfulHashable α] {l: List ((a:α) × (β a))}(h: m.1.WF): - (m.insertList l).1.isEmpty ↔ m.1.isEmpty ∧ l.isEmpty := by + (m.insertMany l).1.1.isEmpty ↔ m.1.isEmpty ∧ l.isEmpty := by simp_to_model using List.insertList_isEmpty + end Raw₀ end Std.DHashMap.Internal diff --git a/src/Std/Data/DHashMap/Internal/WF.lean b/src/Std/Data/DHashMap/Internal/WF.lean index 2b53f475ec8d..d6c324d3733f 100644 --- a/src/Std/Data/DHashMap/Internal/WF.lean +++ b/src/Std/Data/DHashMap/Internal/WF.lean @@ -729,17 +729,6 @@ induction l generalizing m with apply List.insertList_perm_of_perm_first apply toListModel_insert h apply (wfImp_insert h).distinct - -/-! # `insertList` -/ -theorem wfImp_insertList [BEq α] [Hashable α] [EquivBEq α][LawfulHashable α] {m : Raw₀ α β} {l: List ((a : α) × β a)} (h : Raw.WFImp m.1): Raw.WFImp (m.insertList l).1 := by - rw [insertList_eq_insertListₘ] - apply wfImp_insertListₘ h - -theorem toListModel_insertList [BEq α] [Hashable α] [EquivBEq α][LawfulHashable α] {m : Raw₀ α β} {l: List ((a : α) × β a)} (h : Raw.WFImp m.1): - Perm (toListModel (insertList m l).1.buckets) (List.insertList (toListModel m.1.buckets) l) := by - rw [insertList_eq_insertListₘ] - apply toListModel_insertListₘ h - end Raw₀ namespace Raw @@ -768,6 +757,12 @@ theorem wfImp_insertMany [BEq α] [Hashable α] [EquivBEq α] [LawfulHashable α Raw.WFImp (m.insertMany l).1.1 := Raw.WF.out ((m.insertMany l).2 _ Raw.WF.insert₀ (.wf m.2 h)) +theorem toListModel_insertMany [BEq α] [Hashable α] [EquivBEq α] [LawfulHashable α] {m : Raw₀ α β} {l: List ((a:α) × (β a))} (h : Raw.WFImp m.1): + Perm (toListModel (insertMany m l).1.1.buckets) (List.insertList (toListModel m.1.buckets) l) := by + rw[insertMany_eq_insertListₘ] + apply toListModel_insertListₘ + exact h + theorem Const.wfImp_insertMany {β : Type v} [BEq α] [Hashable α] [EquivBEq α] [LawfulHashable α] {ρ : Type w} [ForIn Id ρ (α × β)] {m : Raw₀ α (fun _ => β)} {l : ρ} (h : Raw.WFImp m.1) : Raw.WFImp (Const.insertMany m l).1.1 := From cea372326adca6a513cc6f415b2ef643400ccf66 Mon Sep 17 00:00:00 2001 From: monsterkrampe Date: Wed, 27 Nov 2024 16:20:28 +0100 Subject: [PATCH 22/83] Hotfix for build after rebase --- src/Std/Data/DHashMap/Internal/List/Associative.lean | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/src/Std/Data/DHashMap/Internal/List/Associative.lean b/src/Std/Data/DHashMap/Internal/List/Associative.lean index befe1fbed638..05c4331d134c 100644 --- a/src/Std/Data/DHashMap/Internal/List/Associative.lean +++ b/src/Std/Data/DHashMap/Internal/List/Associative.lean @@ -2278,9 +2278,9 @@ theorem contains_insertList_of_mem [BEq α] [PartialEquivBEq α] (l toInsert: Li rw [containsKey_insertList] right simp [List.contains_eq_any_beq] - exists k + exists ⟨k, v⟩ constructor - · exists ⟨k,v⟩ + · exact mem · apply PartialEquivBEq.symm exact k_eq From 6f01fd394247d340fbbbdb75ca69cb9d4d8c35f7 Mon Sep 17 00:00:00 2001 From: monsterkrampe Date: Wed, 27 Nov 2024 17:28:04 +0100 Subject: [PATCH 23/83] Split getKey? result into two cases --- .../DHashMap/Internal/List/Associative.lean | 189 +++++++----------- src/Std/Data/DHashMap/Internal/RawLemmas.lean | 80 +++----- 2 files changed, 107 insertions(+), 162 deletions(-) diff --git a/src/Std/Data/DHashMap/Internal/List/Associative.lean b/src/Std/Data/DHashMap/Internal/List/Associative.lean index 05c4331d134c..e3b381d48382 100644 --- a/src/Std/Data/DHashMap/Internal/List/Associative.lean +++ b/src/Std/Data/DHashMap/Internal/List/Associative.lean @@ -2331,129 +2331,88 @@ theorem getValueCast_insertList_not_toInsert_mem [BEq α] [LawfulBEq α] (l toIn · exact not_mem · exact distinct -theorem getKey?_insertList [BEq α] [PartialEquivBEq α] {l toInsert : List ((a : α) × β a)} {k : α} : - getKey? k (insertList l toInsert) = - if (toInsert.map Sigma.fst).contains k then (toInsert.map Sigma.fst).reverse.find? (fun a => k == a) else getKey? k l := by - rw [← containsKey_eq_contains_map_fst] +theorem getKey?_insertList_not_mem [BEq α] [PartialEquivBEq α] + {l toInsert : List ((a : α) × β a)} {k : α} + {distinct: toInsert.Pairwise (fun a b => (a.1 == b.1) = false)} {not_mem : (toInsert.map Sigma.fst).contains k = false} : + getKey? k (insertList l toInsert) = getKey? k l := by + rw [← containsKey_eq_contains_map_fst] at not_mem induction toInsert generalizing l with | nil => simp [insertList] | cons hd tl ih => + rw [containsKey_cons] at not_mem + simp only [Bool.or_eq_false_iff] at not_mem unfold insertList rw [ih] - rw [getKey?_insertEntry] - rw [containsKey_cons] - simp only [Bool.or_eq_true, List.map_cons, List.reverse_cons, List.find?_append, - List.find?_singleton] - cases eq : containsKey k tl with - | true => - simp only [↓reduceIte, or_true] - rw [containsKey_eq_contains_map_fst, List.contains_iff_exists_mem_beq] at eq - rw [Option.or_of_isSome] - simp only [List.find?_isSome, List.mem_reverse, List.mem_map] - rcases eq with ⟨a, a_mem, a_eq⟩ - simp only [List.mem_map] at a_mem - rcases a_mem with ⟨pair, pair_mem, pair_fst_a⟩ - exists a - constructor - . exists pair - . exact a_eq - | false => - simp only [Bool.false_eq_true, ↓reduceIte, or_false] - cases eq2 : hd.fst == k with - | true => - simp only [↓reduceIte] - rw [containsKey_eq_contains_map_fst, ← Bool.not_eq_true, List.contains_iff_exists_mem_beq] at eq - simp only [List.mem_map, not_exists, not_and, Bool.not_eq_true, forall_exists_index, - and_imp, forall_apply_eq_imp_iff₂] at eq - have : (tl.map Sigma.fst).reverse.find? (fun a => k == a) = none := by - simp only [List.find?_eq_none, List.mem_reverse, List.mem_map, Bool.not_eq_true, - forall_exists_index, and_imp, forall_apply_eq_imp_iff₂] - exact eq - simp only [this, Option.none_or, Option.some_eq_ite_none_right, and_true] - rw [BEq.comm] - exact eq2 - | false => simp - -theorem getKey?_insertList_lawful [BEq α] [LawfulBEq α] {l toInsert : List ((a : α) × β a)} {k : α} : - getKey? k (insertList l toInsert) = - if (toInsert.map Sigma.fst).contains k then some k else getKey? k l := by - rw [← containsKey_eq_contains_map_fst] + . rw [getKey?_insertEntry] + simp [not_mem.left] + . exact not_mem.right + . simp only [pairwise_cons] at distinct; exact distinct.right + +theorem getKey?_insertList_mem [BEq α] [PartialEquivBEq α] + {l toInsert : List ((a : α) × β a)} + {k k' : α} {k_beq : k == k'} + {distinct: toInsert.Pairwise (fun a b => (a.1 == b.1) = false)} + {mem : k ∈ (toInsert.map Sigma.fst)} : + getKey? k' (insertList l toInsert) = some k := by induction toInsert generalizing l with - | nil => simp [insertList] + | nil => simp at mem | cons hd tl ih => + simp only [pairwise_cons] at distinct unfold insertList - rw [ih] - rw [getKey?_insertEntry] - rw [containsKey_cons] - cases containsKey k tl with - | true => simp - | false => simp only [Bool.false_eq_true, ↓reduceIte, beq_iff_eq, Bool.or_false]; cases eq : hd.fst == k <;> simp only [Bool.false_eq_true, - ↓reduceIte, Bool.or_self]; apply LawfulBEq.eq_of_beq; assumption - -theorem getKey_insertList [BEq α] [PartialEquivBEq α] {l toInsert : List ((a : α) × β a)} {k : α} -{h₁} : - getKey k (insertList l toInsert) h₁ = - if h₂ : (toInsert.map Sigma.fst).contains k then ((toInsert.map Sigma.fst).reverse.find? (fun a => - k == a)).get (List.find?_isSome_of_map_fst_contains h₂) else getKey k l - (containsKey_of_containsKey_insertList h₁ (Bool.eq_false_iff.2 h₂)) := by - rw [← Option.some_inj] - rw [← getKey?_eq_some_getKey] - rw [getKey?_insertList] - split - . simp - . rw [getKey?_eq_some_getKey] - -theorem getKey_insertList_lawful [BEq α] [LawfulBEq α] {l toInsert : List ((a : α) × β a)} {k : α} -{h₁} : - getKey k (insertList l toInsert) h₁ = - if h₂ : (toInsert.map Sigma.fst).contains k then k else getKey k l - (containsKey_of_containsKey_insertList h₁ (Bool.eq_false_iff.2 h₂)) := by - rw [← Option.some_inj] - rw [← getKey?_eq_some_getKey] - rw [getKey?_insertList_lawful] - split - . simp - . rw [getKey?_eq_some_getKey] - -theorem getKey!_insertList [BEq α] [PartialEquivBEq α] [Inhabited α] {l toInsert : List ((a : α) × β a)} {k : α} : - getKey! k (insertList l toInsert) = - if h₂ : (toInsert.map Sigma.fst).contains k then ((toInsert.map Sigma.fst).reverse.find? (fun a => - k == a)).get (List.find?_isSome_of_map_fst_contains h₂) else getKey! k l := by - rw [getKey!_eq_getKey?] - rw [getKey?_insertList] - split - . rw [Option.get_eq_get!] - . rw [getKey!_eq_getKey?] - -theorem getKey!_insertList_lawful [BEq α] [LawfulBEq α] [Inhabited α] {l toInsert : List ((a : α) × β a)} {k : α} : - getKey! k (insertList l toInsert) = - if (toInsert.map Sigma.fst).contains k then k else getKey! k l := by - rw [getKey!_eq_getKey?] - rw [getKey?_insertList_lawful] - split - . simp - . rw [getKey!_eq_getKey?] - -theorem getKeyD_insertList [BEq α] [PartialEquivBEq α] {l toInsert : List ((a : α) × β -a)} {k fallback : α} : - getKeyD k (insertList l toInsert) fallback = - if h₂ : (toInsert.map Sigma.fst).contains k then ((toInsert.map Sigma.fst).reverse.find? (fun a => - k == a)).get (List.find?_isSome_of_map_fst_contains h₂) else getKeyD k l fallback := by - rw [getKeyD_eq_getKey?] - rw [getKey?_insertList] - split - . rw [Option.get_eq_getD] - . rw [getKeyD_eq_getKey?] - -theorem getKeyD_insertList_lawful [BEq α] [LawfulBEq α] {l toInsert : List ((a : α) × β a)} {k -fallback : α} : - getKeyD k (insertList l toInsert) fallback = - if (toInsert.map Sigma.fst).contains k then k else getKeyD k l fallback := by - rw [getKeyD_eq_getKey?] - rw [getKey?_insertList_lawful] - split - . simp - . rw [getKeyD_eq_getKey?] + simp only [List.map_cons, List.mem_cons, List.mem_map] at mem + cases mem with + | inl mem => + rw [getKey?_insertList_not_mem] + . rw [getKey?_insertEntry] + rw [← mem, k_beq] + simp + . exact distinct.right + . rw [← containsKey_eq_contains_map_fst] + rw [containsKey_eq_false_iff] + intro b b_mem + have distinct_l := distinct.left + specialize distinct_l b b_mem + apply BEq.neq_of_beq_of_neq + . rw [BEq.comm]; exact k_beq + . rw [mem]; exact distinct_l + | inr mem => + rw [ih] + . exact distinct.right + . simp [mem] + +/- theorem getKey_insertList [BEq α] [PartialEquivBEq α] {l toInsert : List ((a : α) × β a)} {k : α} -/ +/- {h₁} : -/ +/- getKey k (insertList l toInsert) h₁ = -/ +/- if h₂ : (toInsert.map Sigma.fst).contains k then ((toInsert.map Sigma.fst).reverse.find? (fun a => -/ +/- k == a)).get (List.find?_isSome_of_map_fst_contains h₂) else getKey k l -/ +/- (containsKey_of_containsKey_insertList h₁ (Bool.eq_false_iff.2 h₂)) := by -/ +/- rw [← Option.some_inj] -/ +/- rw [← getKey?_eq_some_getKey] -/ +/- rw [getKey?_insertList] -/ +/- split -/ +/- . simp -/ +/- . rw [getKey?_eq_some_getKey] -/ + +/- theorem getKey!_insertList [BEq α] [PartialEquivBEq α] [Inhabited α] {l toInsert : List ((a : α) × β a)} {k : α} : -/ +/- getKey! k (insertList l toInsert) = -/ +/- if h₂ : (toInsert.map Sigma.fst).contains k then ((toInsert.map Sigma.fst).reverse.find? (fun a => -/ +/- k == a)).get (List.find?_isSome_of_map_fst_contains h₂) else getKey! k l := by -/ +/- rw [getKey!_eq_getKey?] -/ +/- rw [getKey?_insertList] -/ +/- split -/ +/- . rw [Option.get_eq_get!] -/ +/- . rw [getKey!_eq_getKey?] -/ + +/- theorem getKeyD_insertList [BEq α] [PartialEquivBEq α] {l toInsert : List ((a : α) × β -/ +/- a)} {k fallback : α} : -/ +/- getKeyD k (insertList l toInsert) fallback = -/ +/- if h₂ : (toInsert.map Sigma.fst).contains k then ((toInsert.map Sigma.fst).reverse.find? (fun a => -/ +/- k == a)).get (List.find?_isSome_of_map_fst_contains h₂) else getKeyD k l fallback := by -/ +/- rw [getKeyD_eq_getKey?] -/ +/- rw [getKey?_insertList] -/ +/- split -/ +/- . rw [Option.get_eq_getD] -/ +/- . rw [getKeyD_eq_getKey?] -/ theorem insertList_perm [BEq α] [ReflBEq α] [PartialEquivBEq α] (l toInsert: List ((a : α) × β a)) (distinct_l: DistinctKeys l) (distinct_toInsert: List.Pairwise (fun a b => (a.1 == b.1) = false) toInsert) (distinct_both: ∀ (a:α), ¬ (containsKey a l ∧ containsKey a toInsert)): Perm (insertList l toInsert) (l++toInsert) := by diff --git a/src/Std/Data/DHashMap/Internal/RawLemmas.lean b/src/Std/Data/DHashMap/Internal/RawLemmas.lean index 1335fb879c18..14d5f51aa208 100644 --- a/src/Std/Data/DHashMap/Internal/RawLemmas.lean +++ b/src/Std/Data/DHashMap/Internal/RawLemmas.lean @@ -916,53 +916,39 @@ theorem get_insertList_not_mem_list [LawfulBEq α] {l: List ((a:α) × (β a))} (h': m.contains k = true) → (m.insertMany l).1.get k' (contains_insertList_of_contains_map m h k_eq h') = m.get k' (contains_of_beq m k_eq h h') := by simp_to_model using getValueCast_insertList_not_toInsert_mem -theorem getKey?_insertList [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l : List ((a:α) × (β a))} {k : α} : - (m.insertMany l).1.getKey? k = - if (l.map Sigma.fst).contains k then (l.map Sigma.fst).reverse.find? (fun a => k == a) else m.getKey? k := by - simp_to_model using List.getKey?_insertList - -theorem getKey?_insertList_lawful [LawfulBEq α] (h : m.1.WF) {l : List ((a:α) × (β a))} {k : α} : - (m.insertMany l).1.getKey? k = - if (l.map Sigma.fst).contains k then some k else m.getKey? k := by - simp_to_model using List.getKey?_insertList_lawful - -theorem getKey_insertList [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l : List ((a:α) × (β a))} {k -: α} {h₁} : - (m.insertMany l).1.getKey k h₁ = - if h₂ : (l.map Sigma.fst).contains k then ((l.map Sigma.fst).reverse.find? (fun a => k == a)).get (List.find?_isSome_of_map_fst_contains h₂) - else m.getKey k (contains_of_contains_insertList _ h h₁ (Bool.eq_false_iff.2 h₂)) := by - simp_to_model using List.getKey_insertList - -theorem getKey_insertList_lawful [LawfulBEq α] (h : m.1.WF) {l : List ((a:α) × (β a))} {k -: α} {h₁} : - (m.insertMany l).1.getKey k h₁ = - if h₂ : (l.map Sigma.fst).contains k then k - else m.getKey k (contains_of_contains_insertList _ h h₁ (Bool.eq_false_iff.2 h₂)) := by - simp_to_model using List.getKey_insertList_lawful - -theorem getKey!_insertList [EquivBEq α] [LawfulHashable α] [Inhabited α] (h : m.1.WF) {l : List ((a:α) × (β a))} {k : α} : - (m.insertMany l).1.getKey! k = - if h₂ : (l.map Sigma.fst).contains k then ((l.map Sigma.fst).reverse.find? (fun a => k == a)).get (List.find?_isSome_of_map_fst_contains h₂) - else m.getKey! k := by - simp_to_model using List.getKey!_insertList - -theorem getKey!_insertList_lawful [LawfulBEq α] [Inhabited α] (h : m.1.WF) {l : List ((a:α) × (β a))} {k : α} : - (m.insertMany l).1.getKey! k = - if (l.map Sigma.fst).contains k then k - else m.getKey! k := by - simp_to_model using List.getKey!_insertList_lawful - -theorem getKeyD_insertList [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l : List ((a:α) × (β a))} {k fallback : α} : - (m.insertMany l).1.getKeyD k fallback = - if h₂ : (l.map Sigma.fst).contains k then ((l.map Sigma.fst).reverse.find? (fun a => k == a)).get (List.find?_isSome_of_map_fst_contains h₂) - else m.getKeyD k fallback := by - simp_to_model using List.getKeyD_insertList - -theorem getKeyD_insertList_lawful [LawfulBEq α] (h : m.1.WF) {l : List ((a:α) × (β a))} {k fallback : α} : - (m.insertMany l).1.getKeyD k fallback = - if (l.map Sigma.fst).contains k then k - else m.getKeyD k fallback := by - simp_to_model using List.getKeyD_insertList_lawful +theorem getKey?_insertList_not_mem [EquivBEq α] [LawfulHashable α] (h : m.1.WF) + {l : List ((a:α) × (β a))} {k : α} + {distinct: l.Pairwise (fun a b => (a.1 == b.1) = false)} + {not_mem : (l.map Sigma.fst).contains k = false} : + (m.insertMany l).1.getKey? k = m.getKey? k := by + simp_to_model using List.getKey?_insertList_not_mem + +theorem getKey?_insertList_mem [EquivBEq α] [LawfulHashable α] (h : m.1.WF) + {l : List ((a:α) × (β a))} + {k k' : α} {k_beq : k == k'} + {distinct: l.Pairwise (fun a b => (a.1 == b.1) = false)} + {mem : k ∈ (l.map Sigma.fst)}: + (m.insertMany l).1.getKey? k' = some k := by + simp_to_model using List.getKey?_insertList_mem + +/- theorem getKey_insertList [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l : List ((a:α) × (β a))} {k -/ +/- : α} {h₁} : -/ +/- (m.insertMany l).1.getKey k h₁ = -/ +/- if h₂ : (l.map Sigma.fst).contains k then ((l.map Sigma.fst).reverse.find? (fun a => k == a)).get (List.find?_isSome_of_map_fst_contains h₂) -/ +/- else m.getKey k (contains_of_contains_insertList _ h h₁ (Bool.eq_false_iff.2 h₂)) := by -/ +/- simp_to_model using List.getKey_insertList -/ + +/- theorem getKey!_insertList [EquivBEq α] [LawfulHashable α] [Inhabited α] (h : m.1.WF) {l : List ((a:α) × (β a))} {k : α} : -/ +/- (m.insertMany l).1.getKey! k = -/ +/- if h₂ : (l.map Sigma.fst).contains k then ((l.map Sigma.fst).reverse.find? (fun a => k == a)).get (List.find?_isSome_of_map_fst_contains h₂) -/ +/- else m.getKey! k := by -/ +/- simp_to_model using List.getKey!_insertList -/ + +/- theorem getKeyD_insertList [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l : List ((a:α) × (β a))} {k fallback : α} : -/ +/- (m.insertMany l).1.getKeyD k fallback = -/ +/- if h₂ : (l.map Sigma.fst).contains k then ((l.map Sigma.fst).reverse.find? (fun a => k == a)).get (List.find?_isSome_of_map_fst_contains h₂) -/ +/- else m.getKeyD k fallback := by -/ +/- simp_to_model using List.getKeyD_insertList -/ theorem size_insertList [EquivBEq α] [LawfulHashable α] {l: List ((a:α) × (β a))} {distinct: List.Pairwise (fun a b => (a.1 == b.1) = false) l} (h: m.1.WF): (∀ (a:α), ¬ (m.contains a = true ∧ List.containsKey a l = true)) → From c85b1561ad7b52a6dc1aa32c645dd27922585260 Mon Sep 17 00:00:00 2001 From: monsterkrampe Date: Thu, 28 Nov 2024 09:13:44 +0100 Subject: [PATCH 24/83] Adjust remaining getKey results --- .../DHashMap/Internal/List/Associative.lean | 111 ++++++++++++------ src/Std/Data/DHashMap/Internal/RawLemmas.lean | 66 ++++++++--- 2 files changed, 124 insertions(+), 53 deletions(-) diff --git a/src/Std/Data/DHashMap/Internal/List/Associative.lean b/src/Std/Data/DHashMap/Internal/List/Associative.lean index e3b381d48382..2254802bb8c4 100644 --- a/src/Std/Data/DHashMap/Internal/List/Associative.lean +++ b/src/Std/Data/DHashMap/Internal/List/Associative.lean @@ -2333,7 +2333,8 @@ theorem getValueCast_insertList_not_toInsert_mem [BEq α] [LawfulBEq α] (l toIn theorem getKey?_insertList_not_mem [BEq α] [PartialEquivBEq α] {l toInsert : List ((a : α) × β a)} {k : α} - {distinct: toInsert.Pairwise (fun a b => (a.1 == b.1) = false)} {not_mem : (toInsert.map Sigma.fst).contains k = false} : + {distinct: toInsert.Pairwise (fun a b => (a.1 == b.1) = false)} + {not_mem : (toInsert.map Sigma.fst).contains k = false} : getKey? k (insertList l toInsert) = getKey? k l := by rw [← containsKey_eq_contains_map_fst] at not_mem induction toInsert generalizing l with @@ -2380,39 +2381,81 @@ theorem getKey?_insertList_mem [BEq α] [PartialEquivBEq α] . exact distinct.right . simp [mem] -/- theorem getKey_insertList [BEq α] [PartialEquivBEq α] {l toInsert : List ((a : α) × β a)} {k : α} -/ -/- {h₁} : -/ -/- getKey k (insertList l toInsert) h₁ = -/ -/- if h₂ : (toInsert.map Sigma.fst).contains k then ((toInsert.map Sigma.fst).reverse.find? (fun a => -/ -/- k == a)).get (List.find?_isSome_of_map_fst_contains h₂) else getKey k l -/ -/- (containsKey_of_containsKey_insertList h₁ (Bool.eq_false_iff.2 h₂)) := by -/ -/- rw [← Option.some_inj] -/ -/- rw [← getKey?_eq_some_getKey] -/ -/- rw [getKey?_insertList] -/ -/- split -/ -/- . simp -/ -/- . rw [getKey?_eq_some_getKey] -/ - -/- theorem getKey!_insertList [BEq α] [PartialEquivBEq α] [Inhabited α] {l toInsert : List ((a : α) × β a)} {k : α} : -/ -/- getKey! k (insertList l toInsert) = -/ -/- if h₂ : (toInsert.map Sigma.fst).contains k then ((toInsert.map Sigma.fst).reverse.find? (fun a => -/ -/- k == a)).get (List.find?_isSome_of_map_fst_contains h₂) else getKey! k l := by -/ -/- rw [getKey!_eq_getKey?] -/ -/- rw [getKey?_insertList] -/ -/- split -/ -/- . rw [Option.get_eq_get!] -/ -/- . rw [getKey!_eq_getKey?] -/ - -/- theorem getKeyD_insertList [BEq α] [PartialEquivBEq α] {l toInsert : List ((a : α) × β -/ -/- a)} {k fallback : α} : -/ -/- getKeyD k (insertList l toInsert) fallback = -/ -/- if h₂ : (toInsert.map Sigma.fst).contains k then ((toInsert.map Sigma.fst).reverse.find? (fun a => -/ -/- k == a)).get (List.find?_isSome_of_map_fst_contains h₂) else getKeyD k l fallback := by -/ -/- rw [getKeyD_eq_getKey?] -/ -/- rw [getKey?_insertList] -/ -/- split -/ -/- . rw [Option.get_eq_getD] -/ -/- . rw [getKeyD_eq_getKey?] -/ +theorem getKey_insertList_not_mem [BEq α] [PartialEquivBEq α] + {l toInsert : List ((a : α) × β a)} {k : α} + {distinct: toInsert.Pairwise (fun a b => (a.1 == b.1) = false)} + {not_mem : (toInsert.map Sigma.fst).contains k = false} + {h} : + getKey k (insertList l toInsert) h = + getKey k l (containsKey_of_containsKey_insertList h not_mem) := by + rw [← Option.some_inj] + rw [← getKey?_eq_some_getKey] + rw [getKey?_insertList_not_mem] + . rw [getKey?_eq_some_getKey] + . exact distinct + . exact not_mem + +theorem getKey_insertList_mem [BEq α] [PartialEquivBEq α] + {l toInsert : List ((a : α) × β a)} + {k k' : α} {k_beq : k == k'} + {distinct: toInsert.Pairwise (fun a b => (a.1 == b.1) = false)} + {mem : k ∈ (toInsert.map Sigma.fst)} + {h} : + getKey k' (insertList l toInsert) h = k := by + rw [← Option.some_inj] + rw [← getKey?_eq_some_getKey] + rw [getKey?_insertList_mem] + . exact k_beq + . exact distinct + . exact mem + +theorem getKey!_insertList_not_mem [BEq α] [PartialEquivBEq α] [Inhabited α] + {l toInsert : List ((a : α) × β a)} {k : α} + {distinct: toInsert.Pairwise (fun a b => (a.1 == b.1) = false)} + {not_mem : (toInsert.map Sigma.fst).contains k = false} : + getKey! k (insertList l toInsert) = getKey! k l := by + rw [getKey!_eq_getKey?] + rw [getKey?_insertList_not_mem] + . rw [getKey!_eq_getKey?] + . exact distinct + . exact not_mem + +theorem getKey!_insertList_mem [BEq α] [PartialEquivBEq α] [Inhabited α] + {l toInsert : List ((a : α) × β a)} + {k k' : α} {k_beq : k == k'} + {distinct: toInsert.Pairwise (fun a b => (a.1 == b.1) = false)} + {mem : k ∈ (toInsert.map Sigma.fst)} : + getKey! k' (insertList l toInsert) = k := by + rw [getKey!_eq_getKey?] + rw [getKey?_insertList_mem] + . rw [Option.get!_some] + . exact k_beq + . exact distinct + . exact mem + +theorem getKeyD_insertList_not_mem [BEq α] [PartialEquivBEq α] + {l toInsert : List ((a : α) × β a)} {k fallback : α} + {distinct: toInsert.Pairwise (fun a b => (a.1 == b.1) = false)} + {not_mem : (toInsert.map Sigma.fst).contains k = false} : + getKeyD k (insertList l toInsert) fallback = getKeyD k l fallback := by + rw [getKeyD_eq_getKey?] + rw [getKey?_insertList_not_mem] + . rw [getKeyD_eq_getKey?] + . exact distinct + . exact not_mem + +theorem getKeyD_insertList_mem [BEq α] [PartialEquivBEq α] + {l toInsert : List ((a : α) × β a)} + {k k' fallback : α} {k_beq : k == k'} + {distinct: toInsert.Pairwise (fun a b => (a.1 == b.1) = false)} + {mem : k ∈ (toInsert.map Sigma.fst)} : + getKeyD k' (insertList l toInsert) fallback = k := by + rw [getKeyD_eq_getKey?] + rw [getKey?_insertList_mem] + . rw [Option.getD_some] + . exact k_beq + . exact distinct + . exact mem theorem insertList_perm [BEq α] [ReflBEq α] [PartialEquivBEq α] (l toInsert: List ((a : α) × β a)) (distinct_l: DistinctKeys l) (distinct_toInsert: List.Pairwise (fun a b => (a.1 == b.1) = false) toInsert) (distinct_both: ∀ (a:α), ¬ (containsKey a l ∧ containsKey a toInsert)): Perm (insertList l toInsert) (l++toInsert) := by diff --git a/src/Std/Data/DHashMap/Internal/RawLemmas.lean b/src/Std/Data/DHashMap/Internal/RawLemmas.lean index 14d5f51aa208..ff367dc87d3d 100644 --- a/src/Std/Data/DHashMap/Internal/RawLemmas.lean +++ b/src/Std/Data/DHashMap/Internal/RawLemmas.lean @@ -927,28 +927,56 @@ theorem getKey?_insertList_mem [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l : List ((a:α) × (β a))} {k k' : α} {k_beq : k == k'} {distinct: l.Pairwise (fun a b => (a.1 == b.1) = false)} - {mem : k ∈ (l.map Sigma.fst)}: + {mem : k ∈ (l.map Sigma.fst)} : (m.insertMany l).1.getKey? k' = some k := by simp_to_model using List.getKey?_insertList_mem -/- theorem getKey_insertList [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l : List ((a:α) × (β a))} {k -/ -/- : α} {h₁} : -/ -/- (m.insertMany l).1.getKey k h₁ = -/ -/- if h₂ : (l.map Sigma.fst).contains k then ((l.map Sigma.fst).reverse.find? (fun a => k == a)).get (List.find?_isSome_of_map_fst_contains h₂) -/ -/- else m.getKey k (contains_of_contains_insertList _ h h₁ (Bool.eq_false_iff.2 h₂)) := by -/ -/- simp_to_model using List.getKey_insertList -/ - -/- theorem getKey!_insertList [EquivBEq α] [LawfulHashable α] [Inhabited α] (h : m.1.WF) {l : List ((a:α) × (β a))} {k : α} : -/ -/- (m.insertMany l).1.getKey! k = -/ -/- if h₂ : (l.map Sigma.fst).contains k then ((l.map Sigma.fst).reverse.find? (fun a => k == a)).get (List.find?_isSome_of_map_fst_contains h₂) -/ -/- else m.getKey! k := by -/ -/- simp_to_model using List.getKey!_insertList -/ - -/- theorem getKeyD_insertList [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l : List ((a:α) × (β a))} {k fallback : α} : -/ -/- (m.insertMany l).1.getKeyD k fallback = -/ -/- if h₂ : (l.map Sigma.fst).contains k then ((l.map Sigma.fst).reverse.find? (fun a => k == a)).get (List.find?_isSome_of_map_fst_contains h₂) -/ -/- else m.getKeyD k fallback := by -/ -/- simp_to_model using List.getKeyD_insertList -/ +theorem getKey_insertList_not_mem [EquivBEq α] [LawfulHashable α] (h : m.1.WF) + {l : List ((a:α) × (β a))} {k : α} + {distinct: l.Pairwise (fun a b => (a.1 == b.1) = false)} + {not_mem : (l.map Sigma.fst).contains k = false} + {h'} : + (m.insertMany l).1.getKey k h' = m.getKey k (contains_of_contains_insertList _ h h' not_mem) := by + simp_to_model using List.getKey_insertList_not_mem + +theorem getKey_insertList_mem [EquivBEq α] [LawfulHashable α] (h : m.1.WF) + {l : List ((a:α) × (β a))} + {k k' : α} {k_beq : k == k'} + {distinct: l.Pairwise (fun a b => (a.1 == b.1) = false)} + {mem : k ∈ (l.map Sigma.fst)} + {h'} : + (m.insertMany l).1.getKey k' h' = k := by + simp_to_model using List.getKey_insertList_mem + +theorem getKey!_insertList_not_mem [EquivBEq α] [LawfulHashable α] [Inhabited α] (h : m.1.WF) + {l : List ((a:α) × (β a))} {k : α} + {distinct: l.Pairwise (fun a b => (a.1 == b.1) = false)} + {not_mem : (l.map Sigma.fst).contains k = false} : + (m.insertMany l).1.getKey! k = m.getKey! k := by + simp_to_model using List.getKey!_insertList_not_mem + +theorem getKey!_insertList_mem [EquivBEq α] [LawfulHashable α] [Inhabited α] (h : m.1.WF) + {l : List ((a:α) × (β a))} + {k k' : α} {k_beq : k == k'} + {distinct: l.Pairwise (fun a b => (a.1 == b.1) = false)} + {mem : k ∈ (l.map Sigma.fst)} : + (m.insertMany l).1.getKey! k' = k := by + simp_to_model using List.getKey!_insertList_mem + +theorem getKeyD_insertList_not_mem [EquivBEq α] [LawfulHashable α] (h : m.1.WF) + {l : List ((a:α) × (β a))} {k fallback : α} + {distinct: l.Pairwise (fun a b => (a.1 == b.1) = false)} + {not_mem : (l.map Sigma.fst).contains k = false} : + (m.insertMany l).1.getKeyD k fallback = m.getKeyD k fallback := by + simp_to_model using List.getKeyD_insertList_not_mem + +theorem getKeyD_insertList_mem [EquivBEq α] [LawfulHashable α] (h : m.1.WF) + {l : List ((a:α) × (β a))} + {k k' fallback : α} {k_beq : k == k'} + {distinct: l.Pairwise (fun a b => (a.1 == b.1) = false)} + {mem : k ∈ (l.map Sigma.fst)} : + (m.insertMany l).1.getKeyD k' fallback = k := by + simp_to_model using List.getKeyD_insertList_mem theorem size_insertList [EquivBEq α] [LawfulHashable α] {l: List ((a:α) × (β a))} {distinct: List.Pairwise (fun a b => (a.1 == b.1) = false) l} (h: m.1.WF): (∀ (a:α), ¬ (m.contains a = true ∧ List.containsKey a l = true)) → From 3a4c796bbb3e16611a1097ed2684a77b527b5adc Mon Sep 17 00:00:00 2001 From: monsterkrampe Date: Thu, 28 Nov 2024 09:24:49 +0100 Subject: [PATCH 25/83] Remove obsolete auxiliary result --- .../Data/DHashMap/Internal/List/Associative.lean | 16 ---------------- 1 file changed, 16 deletions(-) diff --git a/src/Std/Data/DHashMap/Internal/List/Associative.lean b/src/Std/Data/DHashMap/Internal/List/Associative.lean index 2254802bb8c4..35e0a1c6d891 100644 --- a/src/Std/Data/DHashMap/Internal/List/Associative.lean +++ b/src/Std/Data/DHashMap/Internal/List/Associative.lean @@ -17,22 +17,6 @@ the contents of this file. File contents: Verification of associative lists -/ - --- TODO: should this go somwhere else? -theorem List.find?_isSome_of_map_fst_contains {α : Type u} {β : α -> Type v} [BEq α] {l : List ((a : α) × β a)} {k : α} : - (l.map Sigma.fst).contains k -> - ((l.map Sigma.fst).reverse.find? (fun a => k == a)).isSome := by - intro h - simp only [find?_isSome, mem_reverse, mem_map] - rw [List.contains_iff_exists_mem_beq] at h - rcases h with ⟨a, a_mem, a_eq⟩ - simp only [mem_map] at a_mem - rcases a_mem with ⟨pair, pair_mem, pair_fst_a⟩ - exists a - constructor - . exists pair - . exact a_eq - set_option linter.missingDocs true set_option autoImplicit false From 0cb4f2a5b638436a6583bfe07579ef837d316593 Mon Sep 17 00:00:00 2001 From: monsterkrampe Date: Fri, 29 Nov 2024 10:52:04 +0100 Subject: [PATCH 26/83] Use getKey?_of_perm for getKey?_insertList --- .../DHashMap/Internal/List/Associative.lean | 110 ++++++++++++------ 1 file changed, 74 insertions(+), 36 deletions(-) diff --git a/src/Std/Data/DHashMap/Internal/List/Associative.lean b/src/Std/Data/DHashMap/Internal/List/Associative.lean index 35e0a1c6d891..89b6e7eea457 100644 --- a/src/Std/Data/DHashMap/Internal/List/Associative.lean +++ b/src/Std/Data/DHashMap/Internal/List/Associative.lean @@ -2029,6 +2029,14 @@ def insertList [BEq α] (l toInsert: List ((a : α) × β a)) : List ((a : α) | .nil => l | .cons ⟨k, v⟩ toInsert => insertList (insertEntry k v l) toInsert +theorem DistinctKeys.insertList [BEq α] [PartialEquivBEq α] {l₁ l₂ : List ((a : α) × β a)} (h : DistinctKeys l₁) : + DistinctKeys (insertList l₁ l₂) := by + induction l₂ using assoc_induction generalizing l₁ + · simpa [insertList] + · rename_i k v t ih + rw [insertList.eq_def] + exact ih h.insertEntry + theorem insertList_perm_of_perm_first [BEq α] [EquivBEq α] (l1 l2 toInsert: List ((a : α) × β a)) (h: Perm l1 l2) (distinct: DistinctKeys l1): Perm (insertList l1 toInsert) (insertList l2 toInsert) := by induction toInsert generalizing l1 l2 with | nil => simp[insertList, h] @@ -2040,6 +2048,26 @@ theorem insertList_perm_of_perm_first [BEq α] [EquivBEq α] (l1 l2 toInsert: Li exact h apply DistinctKeys.insertEntry distinct +theorem insertList_cons_perm [BEq α] [EquivBEq α] (l₁ l₂ : List ((a : α) × β a)) + (p : (a : α) × β a) (hl₁ : DistinctKeys l₁) (hl₂ : DistinctKeys (p :: l₂)) : + (insertList l₁ (p :: l₂)).Perm (insertEntry p.1 p.2 (insertList l₁ l₂)) := by + induction l₂ generalizing p l₁ + · simp [insertList] + · rename_i h t ih + rw [insertList] + refine (ih _ _ hl₁.insertEntry hl₂.tail).trans + ((insertEntry_of_perm hl₁.insertList + (ih _ _ hl₁ (distinctKeys_of_sublist (by simp) hl₂))).trans + (List.Perm.trans ?_ (insertEntry_of_perm hl₁.insertList.insertEntry + (ih _ _ hl₁ (distinctKeys_of_sublist (by simp) hl₂)).symm))) + apply getEntry?_ext hl₁.insertList.insertEntry.insertEntry + hl₁.insertList.insertEntry.insertEntry (fun k => ?_) + simp only [getEntry?_insertEntry] + split <;> rename_i hp <;> split <;> rename_i hh <;> try rfl + rw [DistinctKeys.def] at hl₂ + have := List.rel_of_pairwise_cons hl₂ (List.mem_cons_self _ _) + simp [BEq.trans hh (BEq.symm hp)] at this + theorem containsKey_eq_contains_map_fst [BEq α] [PartialEquivBEq α] (l : List ((a : α) × β a)) (k : α) : containsKey k l = (l.map Sigma.fst).contains k := by induction l with @@ -2315,59 +2343,58 @@ theorem getValueCast_insertList_not_toInsert_mem [BEq α] [LawfulBEq α] (l toIn · exact not_mem · exact distinct -theorem getKey?_insertList_not_mem [BEq α] [PartialEquivBEq α] +theorem getKey?_insertList_not_mem [BEq α] [EquivBEq α] {l toInsert : List ((a : α) × β a)} {k : α} {distinct: toInsert.Pairwise (fun a b => (a.1 == b.1) = false)} + {distinct2: DistinctKeys l} {not_mem : (toInsert.map Sigma.fst).contains k = false} : getKey? k (insertList l toInsert) = getKey? k l := by + rw [← DistinctKeys.def] at distinct rw [← containsKey_eq_contains_map_fst] at not_mem induction toInsert generalizing l with | nil => simp [insertList] | cons hd tl ih => - rw [containsKey_cons] at not_mem - simp only [Bool.or_eq_false_iff] at not_mem - unfold insertList - rw [ih] - . rw [getKey?_insertEntry] - simp [not_mem.left] + rw [containsKey_cons, Bool.or_eq_false_iff] at not_mem + rw [getKey?_of_perm distinct2.insertList (insertList_cons_perm l tl hd distinct2 distinct)] + rw [getKey?_insertEntry] + rw [ih distinct.tail] + . simp [not_mem.left] . exact not_mem.right - . simp only [pairwise_cons] at distinct; exact distinct.right + . exact distinct2 -theorem getKey?_insertList_mem [BEq α] [PartialEquivBEq α] +theorem getKey?_insertList_mem [BEq α] [EquivBEq α] {l toInsert : List ((a : α) × β a)} {k k' : α} {k_beq : k == k'} {distinct: toInsert.Pairwise (fun a b => (a.1 == b.1) = false)} + {distinct2: DistinctKeys l} {mem : k ∈ (toInsert.map Sigma.fst)} : getKey? k' (insertList l toInsert) = some k := by + rw [← DistinctKeys.def] at distinct induction toInsert generalizing l with | nil => simp at mem | cons hd tl ih => - simp only [pairwise_cons] at distinct - unfold insertList - simp only [List.map_cons, List.mem_cons, List.mem_map] at mem + rw [getKey?_of_perm distinct2.insertList (insertList_cons_perm l tl hd distinct2 distinct)] + rw [getKey?_insertEntry] + rw [List.map_cons, List.mem_cons] at mem cases mem with - | inl mem => - rw [getKey?_insertList_not_mem] - . rw [getKey?_insertEntry] - rw [← mem, k_beq] - simp - . exact distinct.right - . rw [← containsKey_eq_contains_map_fst] - rw [containsKey_eq_false_iff] - intro b b_mem - have distinct_l := distinct.left - specialize distinct_l b b_mem - apply BEq.neq_of_beq_of_neq - . rw [BEq.comm]; exact k_beq - . rw [mem]; exact distinct_l + | inl mem => simp [← mem, k_beq] | inr mem => - rw [ih] - . exact distinct.right - . simp [mem] - -theorem getKey_insertList_not_mem [BEq α] [PartialEquivBEq α] + rw [if_neg, ih distinct.tail] + . exact distinct2 + . exact mem + . have := distinct.containsKey_eq_false + rw [containsKey_eq_false_iff] at this + rw [Bool.not_eq_true] + rw [List.mem_map] at mem + rcases mem with ⟨_, pair_mem, pair_eq⟩ + apply BEq.neq_of_neq_of_beq + . apply this; exact pair_mem + . rw [pair_eq]; exact k_beq + +theorem getKey_insertList_not_mem [BEq α] [EquivBEq α] {l toInsert : List ((a : α) × β a)} {k : α} {distinct: toInsert.Pairwise (fun a b => (a.1 == b.1) = false)} + {distinct2: DistinctKeys l} {not_mem : (toInsert.map Sigma.fst).contains k = false} {h} : getKey k (insertList l toInsert) h = @@ -2377,12 +2404,14 @@ theorem getKey_insertList_not_mem [BEq α] [PartialEquivBEq α] rw [getKey?_insertList_not_mem] . rw [getKey?_eq_some_getKey] . exact distinct + . exact distinct2 . exact not_mem -theorem getKey_insertList_mem [BEq α] [PartialEquivBEq α] +theorem getKey_insertList_mem [BEq α] [EquivBEq α] {l toInsert : List ((a : α) × β a)} {k k' : α} {k_beq : k == k'} {distinct: toInsert.Pairwise (fun a b => (a.1 == b.1) = false)} + {distinct2: DistinctKeys l} {mem : k ∈ (toInsert.map Sigma.fst)} {h} : getKey k' (insertList l toInsert) h = k := by @@ -2391,23 +2420,27 @@ theorem getKey_insertList_mem [BEq α] [PartialEquivBEq α] rw [getKey?_insertList_mem] . exact k_beq . exact distinct + . exact distinct2 . exact mem -theorem getKey!_insertList_not_mem [BEq α] [PartialEquivBEq α] [Inhabited α] +theorem getKey!_insertList_not_mem [BEq α] [EquivBEq α] [Inhabited α] {l toInsert : List ((a : α) × β a)} {k : α} {distinct: toInsert.Pairwise (fun a b => (a.1 == b.1) = false)} + {distinct2: DistinctKeys l} {not_mem : (toInsert.map Sigma.fst).contains k = false} : getKey! k (insertList l toInsert) = getKey! k l := by rw [getKey!_eq_getKey?] rw [getKey?_insertList_not_mem] . rw [getKey!_eq_getKey?] . exact distinct + . exact distinct2 . exact not_mem -theorem getKey!_insertList_mem [BEq α] [PartialEquivBEq α] [Inhabited α] +theorem getKey!_insertList_mem [BEq α] [EquivBEq α] [Inhabited α] {l toInsert : List ((a : α) × β a)} {k k' : α} {k_beq : k == k'} {distinct: toInsert.Pairwise (fun a b => (a.1 == b.1) = false)} + {distinct2: DistinctKeys l} {mem : k ∈ (toInsert.map Sigma.fst)} : getKey! k' (insertList l toInsert) = k := by rw [getKey!_eq_getKey?] @@ -2415,23 +2448,27 @@ theorem getKey!_insertList_mem [BEq α] [PartialEquivBEq α] [Inhabited α] . rw [Option.get!_some] . exact k_beq . exact distinct + . exact distinct2 . exact mem -theorem getKeyD_insertList_not_mem [BEq α] [PartialEquivBEq α] +theorem getKeyD_insertList_not_mem [BEq α] [EquivBEq α] {l toInsert : List ((a : α) × β a)} {k fallback : α} {distinct: toInsert.Pairwise (fun a b => (a.1 == b.1) = false)} + {distinct2: DistinctKeys l} {not_mem : (toInsert.map Sigma.fst).contains k = false} : getKeyD k (insertList l toInsert) fallback = getKeyD k l fallback := by rw [getKeyD_eq_getKey?] rw [getKey?_insertList_not_mem] . rw [getKeyD_eq_getKey?] . exact distinct + . exact distinct2 . exact not_mem -theorem getKeyD_insertList_mem [BEq α] [PartialEquivBEq α] +theorem getKeyD_insertList_mem [BEq α] [EquivBEq α] {l toInsert : List ((a : α) × β a)} {k k' fallback : α} {k_beq : k == k'} {distinct: toInsert.Pairwise (fun a b => (a.1 == b.1) = false)} + {distinct2: DistinctKeys l} {mem : k ∈ (toInsert.map Sigma.fst)} : getKeyD k' (insertList l toInsert) fallback = k := by rw [getKeyD_eq_getKey?] @@ -2439,6 +2476,7 @@ theorem getKeyD_insertList_mem [BEq α] [PartialEquivBEq α] . rw [Option.getD_some] . exact k_beq . exact distinct + . exact distinct2 . exact mem theorem insertList_perm [BEq α] [ReflBEq α] [PartialEquivBEq α] (l toInsert: List ((a : α) × β a)) (distinct_l: DistinctKeys l) (distinct_toInsert: List.Pairwise (fun a b => (a.1 == b.1) = false) toInsert) (distinct_both: ∀ (a:α), ¬ (containsKey a l ∧ containsKey a toInsert)): From a2d70afccaa14c377d4b5a88a1962469668ded83 Mon Sep 17 00:00:00 2001 From: monsterkrampe Date: Fri, 29 Nov 2024 11:16:16 +0100 Subject: [PATCH 27/83] Rework getValueCast? according to suggestion --- .../DHashMap/Internal/List/Associative.lean | 356 ++++++------------ src/Std/Data/DHashMap/Internal/RawLemmas.lean | 99 ++--- 2 files changed, 173 insertions(+), 282 deletions(-) diff --git a/src/Std/Data/DHashMap/Internal/List/Associative.lean b/src/Std/Data/DHashMap/Internal/List/Associative.lean index 89b6e7eea457..4ca014b9f293 100644 --- a/src/Std/Data/DHashMap/Internal/List/Associative.lean +++ b/src/Std/Data/DHashMap/Internal/List/Associative.lean @@ -2094,254 +2094,130 @@ theorem containsKey_of_containsKey_insertList [BEq α] [PartialEquivBEq α] (h₂ : (toInsert.map Sigma.fst).contains k = false) : containsKey k l := by rw [containsKey_insertList, h₂] at h₁; simp at h₁; exact h₁ - - -theorem getValueCast?_insertList_start_mem [BEq α] [LawfulBEq α] (l toInsert: List ((a : α) × β a)) {k k': α} {v: β k} {distinct1: List.Pairwise (fun a b => (a.1 == b.1) = false) toInsert} {distinct2: List.Pairwise (fun a b => (a.1 == b.1) = false) l} {k_eq: k == k'} {mem: ⟨k,v⟩ ∈ l} {mem': containsKey k toInsert = false}: getValueCast? k' (insertList l toInsert) = some (cast (by congr; apply LawfulBEq.eq_of_beq k_eq) v) := by +theorem getValueCast?_insertList_not_mem [BEq α] [LawfulBEq α] + {l toInsert : List ((a : α) × β a)} {k : α} + {distinct : toInsert.Pairwise (fun a b => (a.1 == b.1) = false)} + {distinct2 : DistinctKeys l} + {not_mem : (toInsert.map Sigma.fst).contains k = false} : + getValueCast? k (insertList l toInsert) = getValueCast? k l := by + rw [← DistinctKeys.def] at distinct + rw [← containsKey_eq_contains_map_fst] at not_mem induction toInsert generalizing l with - | nil => - simp only [insertList] - apply getValueCast?_mem - · exact k_eq - · exact mem - · exact distinct2 + | nil => simp [insertList] | cons hd tl ih => - simp only [insertList] - apply ih - · exact List.Pairwise.of_cons distinct1 - · simp only[insertEntry, cond_eq_if] - split - · rw [← DistinctKeys.def] - apply DistinctKeys.replaceEntry - rw [DistinctKeys.def] - exact distinct2 - · simp only [beq_eq_false_iff_ne, ne_eq, pairwise_cons] - constructor - · rename_i h - simp only [Bool.not_eq_true] at h - rw [containsKey_eq_false_iff] at h - simp only [beq_eq_false_iff_ne, ne_eq] at h - apply h - · simp only [beq_eq_false_iff_ne, ne_eq] at distinct2 - apply distinct2 - · rw[containsKey_cons_eq_false] at mem' - rw [insertEntry_mem_of_different_keys] - · constructor - · apply mem - · simp only [beq_eq_false_iff_ne, ne_eq] - simp only [beq_eq_false_iff_ne, ne_eq] at mem' - apply Ne.symm - apply And.left mem' - · false_or_by_contra - rename_i h - rw [← h] at mem' - simp only [beq_self_eq_true, Bool.true_eq_false, false_and] at mem' - · apply distinct2 - · simp only [containsKey, Bool.or_eq_false_iff, beq_eq_false_iff_ne, ne_eq] at mem' - apply And.right mem' - -theorem getValueCast?_insertList_toInsert_mem [BEq α] [LawfulBEq α] (l toInsert: List ((a : α) × β a)) {k k': α} {k_eq: k == k'} {v: β k} {distinct: List.Pairwise (fun a b => (a.1 == b.1) = false) toInsert} {distinct2: DistinctKeys l} {mem: ⟨k,v⟩ ∈ toInsert}: getValueCast? k' (insertList l toInsert) = some (cast (by congr;apply LawfulBEq.eq_of_beq k_eq) v) := by - rw [DistinctKeys.def] at distinct2 + rw [containsKey_cons, Bool.or_eq_false_iff] at not_mem + rw [getValueCast?_of_perm distinct2.insertList (insertList_cons_perm l tl hd distinct2 distinct)] + rw [getValueCast?_insertEntry] + rw [ih distinct.tail] + . simp [not_mem.left] + . exact not_mem.right + . exact distinct2 + +theorem getValueCast?_insertList_mem [BEq α] [LawfulBEq α] + {l toInsert : List ((a : α) × β a)} + {k k' : α} {k_beq : k == k'} {v: β k} + {distinct: toInsert.Pairwise (fun a b => (a.1 == b.1) = false)} + {distinct2: DistinctKeys l} + {mem: ⟨k,v⟩ ∈ toInsert} : + getValueCast? k' (insertList l toInsert) = some (cast (by congr; exact LawfulBEq.eq_of_beq k_beq) v) := by + rw [← DistinctKeys.def] at distinct induction toInsert generalizing l with | nil => simp at mem | cons hd tl ih => - simp only [insertList, insertEntry, cond_eq_if] - split - · simp only [List.mem_cons] at mem - cases mem with - | inl mem => - apply getValueCast?_insertList_start_mem - · exact List.Pairwise.of_cons distinct - · rw [← DistinctKeys.def] - apply DistinctKeys.replaceEntry - rw [DistinctKeys.def] - exact distinct2 - · exact k_eq - · rw[replaceEntry_mem_of_mem] - · right - rw [mem] - · rename_i h - exact h - · exact distinct2 - · have k_hd: k = hd.fst := by - simp only [← mem] - rw [k_hd, containsKey_eq_false_iff] - simp only [beq_eq_false_iff_ne, ne_eq] - simp only [beq_eq_false_iff_ne, ne_eq, pairwise_cons] at distinct - apply And.left distinct - | inr mem => - apply ih - · simp only [beq_eq_false_iff_ne, ne_eq] - simp only [beq_eq_false_iff_ne, ne_eq, pairwise_cons] at distinct - apply And.right distinct - · rw [← DistinctKeys.def] - apply DistinctKeys.replaceEntry - rw [DistinctKeys.def] - exact distinct2 - · exact mem - · rename_i contains_hd - by_cases hd_k': (hd.fst ==k') = true - · apply getValueCast?_insertList_start_mem - · exact List.Pairwise.of_cons distinct - · simp only [beq_eq_false_iff_ne, ne_eq, pairwise_cons] - simp only [Bool.not_eq_true] at contains_hd - rw [containsKey_eq_false_iff] at contains_hd - simp only [beq_eq_false_iff_ne, ne_eq] at contains_hd - constructor - · apply contains_hd - · simp only [beq_eq_false_iff_ne, ne_eq] at distinct2 - apply distinct2 - · exact k_eq - · simp only [List.mem_cons] at mem - cases mem with - | inl mem => - simp [mem] - | inr mem => - simp at distinct - exfalso - apply And.left distinct ⟨k,v⟩ mem - simp only - simp only [beq_iff_eq] at hd_k' - simp only [beq_iff_eq] at k_eq - simp only [hd_k', k_eq] - · simp only [Bool.not_eq_true] at contains_hd - simp only [beq_iff_eq] at hd_k' - simp only [beq_iff_eq] at k_eq - rw [k_eq, ← hd_k'] - rw[containsKey_eq_false_iff] - simp only [beq_eq_false_iff_ne, ne_eq, pairwise_cons] at distinct - simp only [beq_eq_false_iff_ne, ne_eq] - apply And.left distinct - · apply ih - · exact List.Pairwise.of_cons distinct - · simp only [List.mem_cons] at mem - cases mem with - | inl mem => - simp only [beq_iff_eq] at hd_k' - exfalso - apply hd_k' - rw [← mem] - simp only [beq_iff_eq] at k_eq - simp [k_eq] - | inr mem => - simp only [beq_eq_false_iff_ne, ne_eq, pairwise_cons] - simp only [Bool.not_eq_true] at contains_hd - rw [containsKey_eq_false_iff] at contains_hd - simp only [beq_eq_false_iff_ne, ne_eq] at contains_hd - constructor - · exact contains_hd - · simp only [beq_eq_false_iff_ne, ne_eq] at distinct2 - apply distinct2 - · simp only [List.mem_cons] at mem - cases mem with - | inl mem => - simp only [← mem, beq_iff_eq] at hd_k' - simp only [beq_iff_eq] at k_eq - contradiction - | inr mem => - exact mem + rw [getValueCast?_of_perm distinct2.insertList (insertList_cons_perm l tl hd distinct2 distinct)] + rw [getValueCast?_insertEntry] + rw [List.mem_cons] at mem + cases mem with + | inl mem => rw [← mem]; simp [k_beq] + | inr mem => + rw [dif_neg, ih distinct.tail] + . exact distinct2 + . exact mem + . have := distinct.containsKey_eq_false + rw [containsKey_eq_false_iff] at this + rw [Bool.not_eq_true] + apply BEq.neq_of_neq_of_beq + . apply this; exact mem + . simp [k_beq] -theorem getValueCast?_insertList_not_toInsert_mem [BEq α] [LawfulBEq α] (l toInsert: List ((a : α) × β a)) {k k': α} {k_eq: k == k'} {not_mem: ¬ k ∈ List.map Sigma.fst toInsert} {distinct: DistinctKeys l}: getValueCast? k' (insertList l toInsert) = getValueCast? k' l := by - simp only [beq_iff_eq] at k_eq - rw [DistinctKeys.def] at distinct - induction toInsert generalizing l with - | nil => simp only [insertList] - | cons hd tl ih => - simp only [insertList] - rw [ih] - · rw [getValueCast?_insertEntry] - split - · rename_i h - simp only [beq_iff_eq] at h - simp only [List.map_cons, h, k_eq, List.mem_cons, List.mem_map, true_or, - not_true_eq_false] at not_mem - · rfl - · simp only [List.map_cons, List.mem_cons, List.mem_map, not_or, not_exists, - not_and] at not_mem - simp only [List.mem_map, not_exists, not_and] - apply And.right not_mem - · rw [← DistinctKeys.def] - apply DistinctKeys.insertEntry - rw [DistinctKeys.def] - exact distinct - -theorem getValueCast!_insertList_toInsert_mem [BEq α] [LawfulBEq α] (l toInsert: List ((a : α) × β a)) {k k': α} [Inhabited (β k')] {k_eq: k == k'} {v: β k} {distinct:List.Pairwise (fun a b => (a.1 == b.1) = false) toInsert} {distinct2: DistinctKeys l} {mem: ⟨k,v⟩ ∈ toInsert}: getValueCast! k' (insertList l toInsert) = (cast (by congr;apply LawfulBEq.eq_of_beq k_eq) v) := by - rw [getValueCast!_eq_getValueCast?, getValueCast?_insertList_toInsert_mem (mem:= mem) (l:=l) (toInsert:=toInsert)] - · simp - · exact k_eq - · exact distinct - · exact distinct2 +theorem getValueCast_insertList_not_mem [BEq α] [LawfulBEq α] + {l toInsert : List ((a : α) × β a)} {k : α} + {distinct : toInsert.Pairwise (fun a b => (a.1 == b.1) = false)} + {distinct2 : DistinctKeys l} + {not_mem : (toInsert.map Sigma.fst).contains k = false} + {h} : + getValueCast k (insertList l toInsert) h = getValueCast k l (containsKey_of_containsKey_insertList h not_mem) := by + rw [← Option.some_inj] + rw [← getValueCast?_eq_some_getValueCast] + rw [← getValueCast?_eq_some_getValueCast] + rw [getValueCast?_insertList_not_mem] + . exact distinct + . exact distinct2 + . exact not_mem -theorem getValueCast!_insertList_not_toInsert_mem [BEq α] [LawfulBEq α] (l toInsert: List ((a : α) × β a)) {k k': α} [Inhabited (β k')] {k_eq: k == k'} {not_mem: ¬ k ∈ List.map Sigma.fst toInsert} {distinct: DistinctKeys l}: getValueCast! k' (insertList l toInsert) = getValueCast! k' l := by - rw [getValueCast!_eq_getValueCast?,getValueCast!_eq_getValueCast? ,getValueCast?_insertList_not_toInsert_mem (not_mem:=not_mem) (distinct:=distinct)] - exact k_eq +theorem getValueCast_insertList_mem [BEq α] [LawfulBEq α] + {l toInsert : List ((a : α) × β a)} + {k k' : α} {k_beq : k == k'} {v: β k} + {distinct: toInsert.Pairwise (fun a b => (a.1 == b.1) = false)} + {distinct2: DistinctKeys l} + {mem: ⟨k,v⟩ ∈ toInsert} + {h} : + getValueCast k' (insertList l toInsert) h = cast (by congr; exact LawfulBEq.eq_of_beq k_beq) v := by + rw [← Option.some_inj] + rw [← getValueCast?_eq_some_getValueCast] + rw [getValueCast?_insertList_mem] + . exact k_beq + . exact distinct + . exact distinct2 + . exact mem -theorem getValueCastD_insertList_toInsert_mem [BEq α] [LawfulBEq α] (l toInsert: List ((a : α) × β a)) {k k': α} {k_eq: k == k'} {v: β k} {fallback : β k'} {distinct: List.Pairwise (fun a b => (a.1 == b.1) = false) toInsert} {distinct2: DistinctKeys l} {mem: ⟨k,v⟩ ∈ toInsert}: getValueCastD k' (insertList l toInsert) fallback = (cast (by congr;apply LawfulBEq.eq_of_beq k_eq) v) := by - rw [getValueCastD_eq_getValueCast?, getValueCast?_insertList_toInsert_mem (mem:= mem) (l:=l) (toInsert:=toInsert)] - · simp - · exact k_eq - · exact distinct - · exact distinct2 - -theorem getValueCastD_insertList_not_toInsert_mem [BEq α] [LawfulBEq α] (l toInsert: List ((a : α) × β a)) {k k': α} {fallback: β k'} {k_eq: k == k'} {not_mem: ¬ k ∈ List.map Sigma.fst toInsert} {distinct: DistinctKeys l}: getValueCastD k' (insertList l toInsert) fallback = getValueCastD k' l fallback := by - rw [getValueCastD_eq_getValueCast?, getValueCastD_eq_getValueCast?, getValueCast?_insertList_not_toInsert_mem (not_mem:=not_mem) (distinct:=distinct)] - exact k_eq - -theorem contains_insertList_of_mem [BEq α] [PartialEquivBEq α] (l toInsert: List ((a : α) × β a)) (k k': α) (v: β k) (k_eq: k == k') (mem: ⟨k,v⟩ ∈ toInsert): containsKey k' (insertList l toInsert) = true := by - rw [containsKey_insertList] - right - simp [List.contains_eq_any_beq] - exists ⟨k, v⟩ - constructor - · exact mem - · apply PartialEquivBEq.symm - exact k_eq - -theorem contains_insertList_of_contains_first [BEq α] [PartialEquivBEq α] {l toInsert: List ((a : α) × β a)} {k k': α} (k_eq: k == k'): containsKey k l = true → containsKey k' (insertList l toInsert) = true:= by - intro h - rw [containsKey_insertList] - left - rw [containsKey_congr (PartialEquivBEq.symm k_eq)] - exact h - -instance Inhabited_of_containsKey [BEq α] [LawfulBEq α] (l: List ((a : α) × β a)) (k k': α) (k_eq: k == k') (mem: containsKey k l = true): Inhabited (β k') := - ⟨ - match l with - | .nil => - (by simp at mem) - | .cons hd tl => - if h:hd.fst == k' - then cast (by congr; simp only [beq_iff_eq] at h; apply h) hd.snd - else - have mem': containsKey k tl = true := by - simp only [containsKey, Bool.or_eq_true, beq_iff_eq] at mem - cases mem with - | inl h' => simp only [beq_iff_eq] at h; simp only [beq_iff_eq] at k_eq; rw [k_eq] at h'; contradiction - | inr h' => exact h' - (Inhabited_of_containsKey tl k k' k_eq mem').default - ⟩ - - -theorem getValueCast_insertList_toInsert_mem [BEq α] [LawfulBEq α] (l toInsert: List ((a : α) × β a)) {k k': α} {v: β k} {k_eq: k == k'} {distinct: List.Pairwise (fun a b => (a.1 == b.1) = false) toInsert} {distinct2: DistinctKeys l} {mem: ⟨k,v⟩ ∈ toInsert}: getValueCast k' (insertList l toInsert) (contains_insertList_of_mem l toInsert k k' v k_eq mem) = cast (by congr; apply LawfulBEq.eq_of_beq k_eq) v := by - have inh: Inhabited (β k') := by - apply Inhabited_of_containsKey toInsert k k' k_eq - rw [containsKey_eq_true_iff_exists_mem] - exists ⟨k,v⟩ - simp only [mem, beq_self_eq_true, and_self] - rw [getValueCast_eq_getValueCast!, getValueCast!_insertList_toInsert_mem] - · exact k_eq - · exact distinct - · exact distinct2 - · exact mem - -theorem getValueCast_insertList_not_toInsert_mem [BEq α] [LawfulBEq α] (l toInsert: List ((a : α) × β a)) {k k': α} {k_eq: k == k'} {not_mem: ¬ k ∈ List.map Sigma.fst toInsert} {distinct: DistinctKeys l} {contains: containsKey k l}: getValueCast k' (insertList l toInsert) (contains_insertList_of_contains_first k_eq contains) = getValueCast k' l (containsKey_of_beq contains k_eq) := by - have inh: Inhabited (β k') := by - apply Inhabited_of_containsKey l k k' k_eq - exact contains - rw [getValueCast_eq_getValueCast!, getValueCast_eq_getValueCast!, getValueCast!_insertList_not_toInsert_mem] - · exact k - · exact k_eq - · exact not_mem - · exact distinct +theorem getValueCast!_insertList_not_mem [BEq α] [LawfulBEq α] + {l toInsert : List ((a : α) × β a)} {k : α} [Inhabited (β k)] + {distinct: toInsert.Pairwise (fun a b => (a.1 == b.1) = false)} + {distinct2: DistinctKeys l} + {not_mem : (toInsert.map Sigma.fst).contains k = false} : + getValueCast! k (insertList l toInsert) = getValueCast! k l := by + rw [getValueCast!_eq_getValueCast?, getValueCast!_eq_getValueCast?, getValueCast?_insertList_not_mem] + . exact distinct + . exact distinct2 + . exact not_mem + +theorem getValueCast!_insertList_mem [BEq α] [LawfulBEq α] + {l toInsert : List ((a : α) × β a)} + {k k' : α} {k_beq : k == k'} {v: β k} [Inhabited (β k')] + {distinct: toInsert.Pairwise (fun a b => (a.1 == b.1) = false)} + {distinct2: DistinctKeys l} + {mem: ⟨k,v⟩ ∈ toInsert} : + getValueCast! k' (insertList l toInsert) = cast (by congr; exact LawfulBEq.eq_of_beq k_beq) v := by + rw [getValueCast!_eq_getValueCast?, getValueCast?_insertList_mem] + . rw [Option.get!_some]; exact k_beq + . exact distinct + . exact distinct2 + . exact mem + +theorem getValueCastD_insertList_not_mem [BEq α] [LawfulBEq α] + {l toInsert : List ((a : α) × β a)} {k : α} {fallback : β k} + {distinct: toInsert.Pairwise (fun a b => (a.1 == b.1) = false)} + {distinct2: DistinctKeys l} + {not_mem : (toInsert.map Sigma.fst).contains k = false} : + getValueCastD k (insertList l toInsert) fallback = getValueCastD k l fallback := by + rw [getValueCastD_eq_getValueCast?, getValueCastD_eq_getValueCast?, getValueCast?_insertList_not_mem] + . exact distinct + . exact distinct2 + . exact not_mem + +theorem getValueCastD_insertList_mem [BEq α] [LawfulBEq α] + {l toInsert : List ((a : α) × β a)} + {k k' : α} {k_beq : k == k'} {v: β k} {fallback : β k'} + {distinct: toInsert.Pairwise (fun a b => (a.1 == b.1) = false)} + {distinct2: DistinctKeys l} + {mem: ⟨k,v⟩ ∈ toInsert} : + getValueCastD k' (insertList l toInsert) fallback = cast (by congr; exact LawfulBEq.eq_of_beq k_beq) v := by + rw [getValueCastD_eq_getValueCast?, getValueCast?_insertList_mem] + . rw [Option.getD_some]; exact k_beq + . exact distinct + . exact distinct2 + . exact mem theorem getKey?_insertList_not_mem [BEq α] [EquivBEq α] {l toInsert : List ((a : α) × β a)} {k : α} diff --git a/src/Std/Data/DHashMap/Internal/RawLemmas.lean b/src/Std/Data/DHashMap/Internal/RawLemmas.lean index ff367dc87d3d..649141988caf 100644 --- a/src/Std/Data/DHashMap/Internal/RawLemmas.lean +++ b/src/Std/Data/DHashMap/Internal/RawLemmas.lean @@ -839,7 +839,6 @@ theorem distinct_keys [EquivBEq α] [LawfulHashable α] (h : m.1.WF) : m.1.keys.Pairwise (fun a b => (a == b) = false) := by simp_to_model using (Raw.WF.out h).distinct.distinct - @[simp] theorem insertList_nil: m.insertMany [] = m := by @@ -858,63 +857,79 @@ theorem insertList_cons {l: List ((a:α) × (β a))} {k: α} {v: β k}: | nil => simp [insertListₘ] | cons hd tl => simp [insertListₘ] - theorem contains_insertList [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l: List ((a:α) × (β a))} {k: α}: (m.insertMany l).1.contains k ↔ m.contains k ∨ (l.map Sigma.fst).contains k := by simp_to_model using List.containsKey_insertList -theorem contains_insertList_of_mem_list [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l: List ((a:α) × (β a))} {k k': α} (k_eq: k == k') {v: β k}: - ⟨k,v⟩ ∈ l → (m.insertMany l).1.contains k' := by - simp_to_model using contains_insertList_of_mem - - theorem contains_of_contains_insertList [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l: List ((a:α) × (β a))} {k: α} : (m.insertMany l).1.contains k → (l.map Sigma.fst).contains k = false → m.contains k := by simp_to_model using List.containsKey_of_containsKey_insertList -theorem contains_insertList_of_contains_map [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l: List ((a:α) × (β a))} {k k': α} (k_eq: k == k'): - m.contains k = true → (m.insertMany l).1.contains k' := by - simp_to_model using contains_insertList_of_contains_first +theorem get?_insertList_not_mem [LawfulBEq α] + {l: List ((a:α) × (β a))} {k : α} + {distinct: l.Pairwise (fun a b => (a.1 == b.1) = false)} + {not_mem: (l.map (Sigma.fst)).contains k = false} + (h: m.1.WF) : + (m.insertMany l).1.get? k = m.get? k := by + simp_to_model using getValueCast?_insertList_not_mem -theorem get?_insertList_mem_list [LawfulBEq α] {l: List ((a:α) × (β a))} {k k': α} {v: β k} {distinct: List.Pairwise (fun a b => (a.1 == b.1) = false) l} {k_eq: k == k'} {mem: ⟨k,v⟩ ∈ l} (h: m.1.WF): +theorem get?_insertList_mem [LawfulBEq α] + {l: List ((a:α) × (β a))} {k k': α} {k_eq: k == k'} {v: β k} + {distinct: l.Pairwise (fun a b => (a.1 == b.1) = false)} + {mem: ⟨k,v⟩ ∈ l} + (h: m.1.WF) : (m.insertMany l).1.get? k' = some (cast (by congr; apply LawfulBEq.eq_of_beq k_eq) v) := by - simp_to_model using getValueCast?_insertList_toInsert_mem - + simp_to_model using getValueCast?_insertList_mem -theorem get?_insertList_not_mem_list [LawfulBEq α] {l: List ((a:α) × (β a))} {k k': α} {mem: ¬ k ∈ (l.map (Sigma.fst))} {k_eq: k == k'} (h: m.1.WF): - (m.insertMany l).1.get? k' = m.get? k' := by - simp_to_model using getValueCast?_insertList_not_toInsert_mem +theorem get_insertList_not_mem [LawfulBEq α] + {l: List ((a:α) × (β a))} {k : α} + {distinct: l.Pairwise (fun a b => (a.1 == b.1) = false)} + {not_mem: (l.map (Sigma.fst)).contains k = false} + {h'} + (h: m.1.WF) : + (m.insertMany l).1.get k h' = m.get k (contains_of_contains_insertList _ h h' not_mem) := by + simp_to_model using getValueCast_insertList_not_mem + +theorem get_insertList_mem [LawfulBEq α] + {l: List ((a:α) × (β a))} {k k': α} {k_eq: k == k'} {v: β k} + {distinct: l.Pairwise (fun a b => (a.1 == b.1) = false)} + {mem: ⟨k,v⟩ ∈ l} + {h'} + (h: m.1.WF) : + (m.insertMany l).1.get k' h' = cast (by congr; apply LawfulBEq.eq_of_beq k_eq) v := by + simp_to_model using getValueCast_insertList_mem + +theorem get!_insertList_not_mem [LawfulBEq α] + {l: List ((a:α) × (β a))} {k : α} [Inhabited (β k)] + {distinct: l.Pairwise (fun a b => (a.1 == b.1) = false)} + {not_mem: (l.map (Sigma.fst)).contains k = false} + (h: m.1.WF) : + (m.insertMany l).1.get! k = m.get! k := by + simp_to_model using getValueCast!_insertList_not_mem -theorem get!_insertList_mem_list [LawfulBEq α] {l: List ((a:α) × (β a))} {k k': α} {v: β k}[Inhabited (β k')] {distinct: List.Pairwise (fun a b => (a.1 == b.1) = false) l} {k_eq: k == k'} {mem: ⟨k,v⟩ ∈ l} (h: m.1.WF): +theorem get!_insertList_mem [LawfulBEq α] + {l: List ((a:α) × (β a))} {k k': α} {k_eq: k == k'} {v: β k} [Inhabited (β k')] + {distinct: l.Pairwise (fun a b => (a.1 == b.1) = false)} + {mem: ⟨k,v⟩ ∈ l} + (h: m.1.WF) : (m.insertMany l).1.get! k' = cast (by congr; apply LawfulBEq.eq_of_beq k_eq) v := by - simp_to_model using getValueCast!_insertList_toInsert_mem + simp_to_model using getValueCast!_insertList_mem -theorem get!_insertList_not_mem_list [LawfulBEq α] {l: List ((a:α) × (β a))} {k k': α} [Inhabited (β k')] {mem: ¬ k ∈ (l.map (Sigma.fst))} {k_eq: k == k'} (h: m.1.WF): - (m.insertMany l).1.get! k' = m.get! k' := by - simp_to_model using getValueCast!_insertList_not_toInsert_mem +theorem getD_insertList_not_mem [LawfulBEq α] + {l: List ((a:α) × (β a))} {k : α} {fallback : β k} + {distinct: l.Pairwise (fun a b => (a.1 == b.1) = false)} + {not_mem: (l.map (Sigma.fst)).contains k = false} + (h: m.1.WF) : + (m.insertMany l).1.getD k fallback = m.getD k fallback := by + simp_to_model using getValueCastD_insertList_not_mem -theorem getD_insertList_mem_list [LawfulBEq α] {l: List ((a:α) × (β a))} {k k': α} {v: β k} {fallback: β k'} {distinct: List.Pairwise (fun a b => (a.1 == b.1) = false) l} {k_eq: k == k'} {mem: ⟨k,v⟩ ∈ l} (h: m.1.WF): +theorem getD_insertList_mem [LawfulBEq α] + {l: List ((a:α) × (β a))} {k k': α} {k_eq: k == k'} {v: β k} {fallback : β k'} + {distinct: l.Pairwise (fun a b => (a.1 == b.1) = false)} + {mem: ⟨k,v⟩ ∈ l} + (h: m.1.WF) : (m.insertMany l).1.getD k' fallback = cast (by congr; apply LawfulBEq.eq_of_beq k_eq) v := by - simp_to_model using getValueCastD_insertList_toInsert_mem - -theorem getD_insertList_not_mem_list [LawfulBEq α] {l: List ((a:α) × (β a))} {k k': α} {fallback: β k'} {mem: ¬ k ∈ (l.map (Sigma.fst))} {k_eq: k == k'} (h: m.1.WF): - (m.insertMany l).1.getD k' fallback = m.getD k' fallback := by - simp_to_model using getValueCastD_insertList_not_toInsert_mem - - -theorem get_insertList_mem_list [LawfulBEq α] {l: List ((a:α) × (β a))} {k k': α} {v: β k} {distinct: List.Pairwise (fun a b => (a.1 == b.1) = false) l} {k_eq: k == k'} {mem: ⟨k,v⟩ ∈ l} (h: m.1.WF): - (m.insertMany l).1.get k' (contains_insertList_of_mem_list _ h k_eq mem)= cast (by congr; apply LawfulBEq.eq_of_beq k_eq ) v := by - simp_to_model using getValueCast_insertList_toInsert_mem - -theorem contains_of_beq [EquivBEq α][LawfulHashable α] {k k': α} (k_eq: k == k') (h: m.1.WF): - m.contains k = true → m.contains k' = true := by - simp_to_model - intro h' - apply containsKey_of_beq h' k_eq - -theorem get_insertList_not_mem_list [LawfulBEq α] {l: List ((a:α) × (β a))} {k k': α} {k_eq: k == k'} {mem_list: ¬ k ∈ (l.map (Sigma.fst))} (h: m.1.WF): - (h': m.contains k = true) → (m.insertMany l).1.get k' (contains_insertList_of_contains_map m h k_eq h') = m.get k' (contains_of_beq m k_eq h h') := by - simp_to_model using getValueCast_insertList_not_toInsert_mem + simp_to_model using getValueCastD_insertList_mem theorem getKey?_insertList_not_mem [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l : List ((a:α) × (β a))} {k : α} From 4fd84f3857f6d94a986d42cb1b9a75f1b621e319 Mon Sep 17 00:00:00 2001 From: monsterkrampe Date: Fri, 29 Nov 2024 11:36:12 +0100 Subject: [PATCH 28/83] Remove obsolete aux results; Fix colon spacing --- .../DHashMap/Internal/List/Associative.lean | 263 +++++------------- src/Std/Data/DHashMap/Internal/Model.lean | 6 +- src/Std/Data/DHashMap/Internal/Raw.lean | 4 +- src/Std/Data/DHashMap/Internal/RawLemmas.lean | 125 +++++---- src/Std/Data/DHashMap/Internal/WF.lean | 17 +- src/Std/Data/DHashMap/Raw.lean | 2 +- src/Std/Data/DHashMap/RawLemmas.lean | 4 +- 7 files changed, 150 insertions(+), 271 deletions(-) diff --git a/src/Std/Data/DHashMap/Internal/List/Associative.lean b/src/Std/Data/DHashMap/Internal/List/Associative.lean index 4ca014b9f293..63c173683d8f 100644 --- a/src/Std/Data/DHashMap/Internal/List/Associative.lean +++ b/src/Std/Data/DHashMap/Internal/List/Associative.lean @@ -164,7 +164,11 @@ theorem getValueCast?_cons_self [BEq α] [LawfulBEq α] {l : List ((a : α) × getValueCast? k (⟨k, v⟩ :: l) = some v := by rw [getValueCast?_cons_of_true BEq.refl, cast_eq] -theorem getValueCast?_mem [BEq α] [LawfulBEq α] {l : List ((a : α) × β a)} {k k': α} {v : β k} {k_eq: k == k'} {mem: ⟨k,v⟩ ∈ l} {distinct: List.Pairwise (fun a b => (a.1 == b.1) = false) l}: getValueCast? k' l = some (cast (by congr; apply LawfulBEq.eq_of_beq k_eq) v) := by +theorem getValueCast?_mem [BEq α] [LawfulBEq α] + {l : List ((a : α) × β a)} + {k k' : α} {v : β k} {k_eq : k == k'} {mem : ⟨k,v⟩ ∈ l} + {distinct : List.Pairwise (fun a b => (a.1 == b.1) = false) l} : + getValueCast? k' l = some (cast (by congr; apply LawfulBEq.eq_of_beq k_eq) v) := by induction l with | nil => simp at mem | cons hd tl ih => @@ -175,7 +179,7 @@ theorem getValueCast?_mem [BEq α] [LawfulBEq α] {l : List ((a : α) × β a)} rw [← mem] simp [k_eq] | inr mem => - have hd_k': (hd.fst == k') = false := by + have hd_k' : (hd.fst == k') = false := by simp only [beq_eq_false_iff_ne, ne_eq, pairwise_cons] at distinct simp only [beq_eq_false_iff_ne, ne_eq] rcases distinct with ⟨distinct, _⟩ @@ -253,7 +257,8 @@ def containsKey [BEq α] (a : α) : List ((a : α) × β a) → Bool @[simp] theorem containsKey_cons [BEq α] {l : List ((a : α) × β a)} {k a : α} {v : β k} : containsKey a (⟨k, v⟩ :: l) = (k == a || containsKey a l) := rfl -theorem containsKey_eq_false_iff [BEq α][PartialEquivBEq α] {l : List ((a : α) × β a)} {a : α}: containsKey a l = false ↔ ∀ (b:((a : α) × β a)), b ∈ l → (a == b.fst) = false := by +theorem containsKey_eq_false_iff [BEq α][PartialEquivBEq α] {l : List ((a : α) × β a)} {a : α} : + containsKey a l = false ↔ ∀ (b : ((a : α) × β a)), b ∈ l → (a == b.fst) = false := by induction l with | nil => simp | cons hd tl ih => @@ -702,90 +707,6 @@ theorem replaceEntry_of_containsKey_eq_false [BEq α] {l : List ((a : α) × β rw [containsKey_cons_eq_false] at h rw [replaceEntry_cons_of_false h.1, ih h.2] -theorem replaceEntry_mem_of_mem [BEq α] [LawfulBEq α] {l : List ((a : α) × β a)} {k: α} {v : β k} {ele: (a: α) × β a} (mem: containsKey k l) {distinct: List.Pairwise (fun a b => (a.1 == b.1) = false) l}: ele ∈ replaceEntry k v l ↔ (ele ∈ l ∧ ¬ (ele.fst == k)) ∨ (ele = ⟨k,v⟩) := by - induction l with - | nil => simp at mem - | cons hd tl ih => - simp only [replaceEntry, cond_eq_if, beq_iff_eq, List.mem_cons] - split - · simp only [List.mem_cons] - rename_i hd_k - simp only [beq_iff_eq] at hd_k - constructor - · intro h - cases h with - | inl h => right; exact h - | inr h => - left - constructor - · right - exact h - · rw [← hd_k] - simp only [beq_eq_false_iff_ne, ne_eq, pairwise_cons] at distinct - apply Ne.symm - apply (And.left distinct ele h) - · intro h - cases h with - | inl h => - rcases h with ⟨ele_mem, ele_k⟩ - cases ele_mem with - | inl h => - rw [h] at ele_k - contradiction - | inr h => - right - apply h - | inr h => - left - apply h - · simp only [List.mem_cons] - rename_i hd_k - simp only [beq_iff_eq] at hd_k - rw [ih] - · constructor - · intro h - cases h with - | inl h => - left - constructor - · left - simp only [h] - · simp only [h, hd_k, not_false_eq_true] - | inr h => - cases h with - | inl h => - simp only [beq_iff_eq] at h - left - constructor - · right - apply And.left h - · apply And.right h - | inr h => - right - exact h - · intro h - cases h with - | inl h => - rcases h with ⟨h,ele_k⟩ - cases h with - | inl h => - left - simp only [h] - | inr h => - right - left - simp [h, ele_k] - | inr h => - right - right - exact h - · rw [containsKey_cons] at mem - simp only [Bool.or_eq_true, beq_iff_eq, hd_k, false_or] at mem - exact mem - · simp only [beq_eq_false_iff_ne, ne_eq, pairwise_cons] at distinct - simp only [beq_eq_false_iff_ne, ne_eq] - apply And.right distinct - @[simp] theorem isEmpty_replaceEntry [BEq α] {l : List ((a : α) × β a)} {k : α} {v : β k} : (replaceEntry k v l).isEmpty = l.isEmpty := by @@ -1001,10 +922,9 @@ theorem containsKey_of_mem [BEq α] [ReflBEq α] {l : List ((a : α) × β a)} { theorem DistinctKeys.nil [BEq α] : DistinctKeys ([] : List ((a : α) × β a)) := ⟨by simp⟩ -theorem DistinctKeys.def [BEq α] {l : List ((a : α) × β a)}: DistinctKeys l ↔ List.Pairwise (fun a b => (a.1 == b.1) = false) l := - ⟨fun h => by simpa [keys_eq_map, List.pairwise_map] using h.distinct, - fun h => ⟨by simpa [keys_eq_map, List.pairwise_map] using h⟩⟩ - +theorem DistinctKeys.def [BEq α] {l : List ((a : α) × β a)} : DistinctKeys l ↔ List.Pairwise (fun a b => (a.1 == b.1) = false) l := + ⟨fun h => by simpa [keys_eq_map, List.pairwise_map] using h.distinct, + fun h => ⟨by simpa [keys_eq_map, List.pairwise_map] using h⟩⟩ open List @@ -1114,67 +1034,11 @@ theorem insertEntry_of_containsKey_eq_false [BEq α] {l : List ((a : α) × β a simp [insertEntry, h] theorem insertEntry_of_containsKey_eq_false_perm [BEq α] {l : List ((a : α) × β a)} {k : α} {v : β k} - (h : containsKey k l = false) : Perm (insertEntry k v l) (⟨k,v⟩::l) := by + (h : containsKey k l = false) : Perm (insertEntry k v l) (⟨k,v⟩ :: l) := by cases l with | nil => simp | cons hd tl => simp [insertEntry, h] -theorem insertEntry_mem_of_different_keys [BEq α] [LawfulBEq α] {l : List ((a : α) × β a)} {k : α} {v : β k} {a: ((a : α) × β a)} {h: ¬ ⟨k,v⟩ = a} {distinct: List.Pairwise (fun a b => (a.1 == b.1) = false) l}: ⟨k, v⟩ ∈ insertEntry a.fst a.snd l ↔ ⟨k,v⟩ ∈ l ∧ (k == a.fst) = false := by - induction l with - | nil => - simp - apply h - | cons hd tl ih => - unfold insertEntry - simp only [cond_eq_if] - split - · rename_i contains_a - rw [replaceEntry_mem_of_mem contains_a (distinct:=distinct)] - simp only [List.mem_cons, beq_iff_eq, beq_eq_false_iff_ne, ne_eq, or_iff_left_iff_imp] - intro h - contradiction - · simp only [List.mem_cons, beq_eq_false_iff_ne, ne_eq] - constructor - · intro h' - cases h' with - | inl h' => contradiction - | inr h' => - rename_i contains_a - simp only [Bool.not_eq_true] at contains_a - rw [containsKey_eq_false_iff] at contains_a - cases h' with - | inl h' => - constructor - · left - exact h' - · simp only [List.mem_cons, beq_eq_false_iff_ne, ne_eq, forall_eq_or_imp] at contains_a - apply Ne.symm - rw [← h'] at contains_a - simp only at contains_a - apply And.left contains_a - | inr h' => - constructor - · right - exact h' - · simp only [List.mem_cons, beq_eq_false_iff_ne, ne_eq, forall_eq_or_imp] at contains_a - apply Ne.symm - rcases contains_a with ⟨_, contains_a⟩ - specialize contains_a ⟨k,v⟩ - simp only at contains_a - apply contains_a h' - · intro h' - rcases h' with ⟨h', _⟩ - cases h' with - | inl h' => - right - left - apply h' - | inr h' => - right - right - apply h' - - theorem DistinctKeys.insertEntry [BEq α] [PartialEquivBEq α] {l : List ((a : α) × β a)} {k : α} {v : β k} (h : DistinctKeys l) : DistinctKeys (insertEntry k v l) := by cases h' : containsKey k l @@ -2024,7 +1888,7 @@ theorem eraseKey_append_of_containsKey_right_eq_false [BEq α] {l l' : List ((a · rw [cond_true, cond_true] /-- Internal implementation detail of the hash map -/ -def insertList [BEq α] (l toInsert: List ((a : α) × β a)) : List ((a : α) × β a) := +def insertList [BEq α] (l toInsert : List ((a : α) × β a)) : List ((a : α) × β a) := match toInsert with | .nil => l | .cons ⟨k, v⟩ toInsert => insertList (insertEntry k v l) toInsert @@ -2037,7 +1901,8 @@ theorem DistinctKeys.insertList [BEq α] [PartialEquivBEq α] {l₁ l₂ : List rw [insertList.eq_def] exact ih h.insertEntry -theorem insertList_perm_of_perm_first [BEq α] [EquivBEq α] (l1 l2 toInsert: List ((a : α) × β a)) (h: Perm l1 l2) (distinct: DistinctKeys l1): Perm (insertList l1 toInsert) (insertList l2 toInsert) := by +theorem insertList_perm_of_perm_first [BEq α] [EquivBEq α] (l1 l2 toInsert : List ((a : α) × β a)) (h : Perm l1 l2) + (distinct : DistinctKeys l1) : Perm (insertList l1 toInsert) (insertList l2 toInsert) := by induction toInsert generalizing l1 l2 with | nil => simp[insertList, h] | cons hd tl ih => @@ -2077,7 +1942,8 @@ theorem containsKey_eq_contains_map_fst [BEq α] [PartialEquivBEq α] (l : List simp only [List.map_cons, List.contains_cons] rw [BEq.comm] -theorem containsKey_insertList [BEq α] [PartialEquivBEq α] (l toInsert : List ((a : α) × β a)) (k : α): containsKey k (List.insertList l toInsert) ↔ containsKey k l ∨ (toInsert.map Sigma.fst).contains k := by +theorem containsKey_insertList [BEq α] [PartialEquivBEq α] (l toInsert : List ((a : α) × β a)) (k : α) : + containsKey k (List.insertList l toInsert) ↔ containsKey k l ∨ (toInsert.map Sigma.fst).contains k := by induction toInsert generalizing l with | nil => simp only [insertList, List.map_nil, List.elem_nil, Bool.false_eq_true, or_false] | cons hd tl ih => @@ -2115,10 +1981,10 @@ theorem getValueCast?_insertList_not_mem [BEq α] [LawfulBEq α] theorem getValueCast?_insertList_mem [BEq α] [LawfulBEq α] {l toInsert : List ((a : α) × β a)} - {k k' : α} {k_beq : k == k'} {v: β k} - {distinct: toInsert.Pairwise (fun a b => (a.1 == b.1) = false)} - {distinct2: DistinctKeys l} - {mem: ⟨k,v⟩ ∈ toInsert} : + {k k' : α} {k_beq : k == k'} {v : β k} + {distinct : toInsert.Pairwise (fun a b => (a.1 == b.1) = false)} + {distinct2 : DistinctKeys l} + {mem : ⟨k,v⟩ ∈ toInsert} : getValueCast? k' (insertList l toInsert) = some (cast (by congr; exact LawfulBEq.eq_of_beq k_beq) v) := by rw [← DistinctKeys.def] at distinct induction toInsert generalizing l with @@ -2157,10 +2023,10 @@ theorem getValueCast_insertList_not_mem [BEq α] [LawfulBEq α] theorem getValueCast_insertList_mem [BEq α] [LawfulBEq α] {l toInsert : List ((a : α) × β a)} - {k k' : α} {k_beq : k == k'} {v: β k} - {distinct: toInsert.Pairwise (fun a b => (a.1 == b.1) = false)} - {distinct2: DistinctKeys l} - {mem: ⟨k,v⟩ ∈ toInsert} + {k k' : α} {k_beq : k == k'} {v : β k} + {distinct : toInsert.Pairwise (fun a b => (a.1 == b.1) = false)} + {distinct2 : DistinctKeys l} + {mem : ⟨k,v⟩ ∈ toInsert} {h} : getValueCast k' (insertList l toInsert) h = cast (by congr; exact LawfulBEq.eq_of_beq k_beq) v := by rw [← Option.some_inj] @@ -2173,8 +2039,8 @@ theorem getValueCast_insertList_mem [BEq α] [LawfulBEq α] theorem getValueCast!_insertList_not_mem [BEq α] [LawfulBEq α] {l toInsert : List ((a : α) × β a)} {k : α} [Inhabited (β k)] - {distinct: toInsert.Pairwise (fun a b => (a.1 == b.1) = false)} - {distinct2: DistinctKeys l} + {distinct : toInsert.Pairwise (fun a b => (a.1 == b.1) = false)} + {distinct2 : DistinctKeys l} {not_mem : (toInsert.map Sigma.fst).contains k = false} : getValueCast! k (insertList l toInsert) = getValueCast! k l := by rw [getValueCast!_eq_getValueCast?, getValueCast!_eq_getValueCast?, getValueCast?_insertList_not_mem] @@ -2184,10 +2050,10 @@ theorem getValueCast!_insertList_not_mem [BEq α] [LawfulBEq α] theorem getValueCast!_insertList_mem [BEq α] [LawfulBEq α] {l toInsert : List ((a : α) × β a)} - {k k' : α} {k_beq : k == k'} {v: β k} [Inhabited (β k')] - {distinct: toInsert.Pairwise (fun a b => (a.1 == b.1) = false)} - {distinct2: DistinctKeys l} - {mem: ⟨k,v⟩ ∈ toInsert} : + {k k' : α} {k_beq : k == k'} {v : β k} [Inhabited (β k')] + {distinct : toInsert.Pairwise (fun a b => (a.1 == b.1) = false)} + {distinct2 : DistinctKeys l} + {mem : ⟨k,v⟩ ∈ toInsert} : getValueCast! k' (insertList l toInsert) = cast (by congr; exact LawfulBEq.eq_of_beq k_beq) v := by rw [getValueCast!_eq_getValueCast?, getValueCast?_insertList_mem] . rw [Option.get!_some]; exact k_beq @@ -2197,8 +2063,8 @@ theorem getValueCast!_insertList_mem [BEq α] [LawfulBEq α] theorem getValueCastD_insertList_not_mem [BEq α] [LawfulBEq α] {l toInsert : List ((a : α) × β a)} {k : α} {fallback : β k} - {distinct: toInsert.Pairwise (fun a b => (a.1 == b.1) = false)} - {distinct2: DistinctKeys l} + {distinct : toInsert.Pairwise (fun a b => (a.1 == b.1) = false)} + {distinct2 : DistinctKeys l} {not_mem : (toInsert.map Sigma.fst).contains k = false} : getValueCastD k (insertList l toInsert) fallback = getValueCastD k l fallback := by rw [getValueCastD_eq_getValueCast?, getValueCastD_eq_getValueCast?, getValueCast?_insertList_not_mem] @@ -2208,10 +2074,10 @@ theorem getValueCastD_insertList_not_mem [BEq α] [LawfulBEq α] theorem getValueCastD_insertList_mem [BEq α] [LawfulBEq α] {l toInsert : List ((a : α) × β a)} - {k k' : α} {k_beq : k == k'} {v: β k} {fallback : β k'} - {distinct: toInsert.Pairwise (fun a b => (a.1 == b.1) = false)} - {distinct2: DistinctKeys l} - {mem: ⟨k,v⟩ ∈ toInsert} : + {k k' : α} {k_beq : k == k'} {v : β k} {fallback : β k'} + {distinct : toInsert.Pairwise (fun a b => (a.1 == b.1) = false)} + {distinct2 : DistinctKeys l} + {mem : ⟨k,v⟩ ∈ toInsert} : getValueCastD k' (insertList l toInsert) fallback = cast (by congr; exact LawfulBEq.eq_of_beq k_beq) v := by rw [getValueCastD_eq_getValueCast?, getValueCast?_insertList_mem] . rw [Option.getD_some]; exact k_beq @@ -2221,8 +2087,8 @@ theorem getValueCastD_insertList_mem [BEq α] [LawfulBEq α] theorem getKey?_insertList_not_mem [BEq α] [EquivBEq α] {l toInsert : List ((a : α) × β a)} {k : α} - {distinct: toInsert.Pairwise (fun a b => (a.1 == b.1) = false)} - {distinct2: DistinctKeys l} + {distinct : toInsert.Pairwise (fun a b => (a.1 == b.1) = false)} + {distinct2 : DistinctKeys l} {not_mem : (toInsert.map Sigma.fst).contains k = false} : getKey? k (insertList l toInsert) = getKey? k l := by rw [← DistinctKeys.def] at distinct @@ -2241,8 +2107,8 @@ theorem getKey?_insertList_not_mem [BEq α] [EquivBEq α] theorem getKey?_insertList_mem [BEq α] [EquivBEq α] {l toInsert : List ((a : α) × β a)} {k k' : α} {k_beq : k == k'} - {distinct: toInsert.Pairwise (fun a b => (a.1 == b.1) = false)} - {distinct2: DistinctKeys l} + {distinct : toInsert.Pairwise (fun a b => (a.1 == b.1) = false)} + {distinct2 : DistinctKeys l} {mem : k ∈ (toInsert.map Sigma.fst)} : getKey? k' (insertList l toInsert) = some k := by rw [← DistinctKeys.def] at distinct @@ -2269,8 +2135,8 @@ theorem getKey?_insertList_mem [BEq α] [EquivBEq α] theorem getKey_insertList_not_mem [BEq α] [EquivBEq α] {l toInsert : List ((a : α) × β a)} {k : α} - {distinct: toInsert.Pairwise (fun a b => (a.1 == b.1) = false)} - {distinct2: DistinctKeys l} + {distinct : toInsert.Pairwise (fun a b => (a.1 == b.1) = false)} + {distinct2 : DistinctKeys l} {not_mem : (toInsert.map Sigma.fst).contains k = false} {h} : getKey k (insertList l toInsert) h = @@ -2286,8 +2152,8 @@ theorem getKey_insertList_not_mem [BEq α] [EquivBEq α] theorem getKey_insertList_mem [BEq α] [EquivBEq α] {l toInsert : List ((a : α) × β a)} {k k' : α} {k_beq : k == k'} - {distinct: toInsert.Pairwise (fun a b => (a.1 == b.1) = false)} - {distinct2: DistinctKeys l} + {distinct : toInsert.Pairwise (fun a b => (a.1 == b.1) = false)} + {distinct2 : DistinctKeys l} {mem : k ∈ (toInsert.map Sigma.fst)} {h} : getKey k' (insertList l toInsert) h = k := by @@ -2301,8 +2167,8 @@ theorem getKey_insertList_mem [BEq α] [EquivBEq α] theorem getKey!_insertList_not_mem [BEq α] [EquivBEq α] [Inhabited α] {l toInsert : List ((a : α) × β a)} {k : α} - {distinct: toInsert.Pairwise (fun a b => (a.1 == b.1) = false)} - {distinct2: DistinctKeys l} + {distinct : toInsert.Pairwise (fun a b => (a.1 == b.1) = false)} + {distinct2 : DistinctKeys l} {not_mem : (toInsert.map Sigma.fst).contains k = false} : getKey! k (insertList l toInsert) = getKey! k l := by rw [getKey!_eq_getKey?] @@ -2315,8 +2181,8 @@ theorem getKey!_insertList_not_mem [BEq α] [EquivBEq α] [Inhabited α] theorem getKey!_insertList_mem [BEq α] [EquivBEq α] [Inhabited α] {l toInsert : List ((a : α) × β a)} {k k' : α} {k_beq : k == k'} - {distinct: toInsert.Pairwise (fun a b => (a.1 == b.1) = false)} - {distinct2: DistinctKeys l} + {distinct : toInsert.Pairwise (fun a b => (a.1 == b.1) = false)} + {distinct2 : DistinctKeys l} {mem : k ∈ (toInsert.map Sigma.fst)} : getKey! k' (insertList l toInsert) = k := by rw [getKey!_eq_getKey?] @@ -2329,8 +2195,8 @@ theorem getKey!_insertList_mem [BEq α] [EquivBEq α] [Inhabited α] theorem getKeyD_insertList_not_mem [BEq α] [EquivBEq α] {l toInsert : List ((a : α) × β a)} {k fallback : α} - {distinct: toInsert.Pairwise (fun a b => (a.1 == b.1) = false)} - {distinct2: DistinctKeys l} + {distinct : toInsert.Pairwise (fun a b => (a.1 == b.1) = false)} + {distinct2 : DistinctKeys l} {not_mem : (toInsert.map Sigma.fst).contains k = false} : getKeyD k (insertList l toInsert) fallback = getKeyD k l fallback := by rw [getKeyD_eq_getKey?] @@ -2343,8 +2209,8 @@ theorem getKeyD_insertList_not_mem [BEq α] [EquivBEq α] theorem getKeyD_insertList_mem [BEq α] [EquivBEq α] {l toInsert : List ((a : α) × β a)} {k k' fallback : α} {k_beq : k == k'} - {distinct: toInsert.Pairwise (fun a b => (a.1 == b.1) = false)} - {distinct2: DistinctKeys l} + {distinct : toInsert.Pairwise (fun a b => (a.1 == b.1) = false)} + {distinct2 : DistinctKeys l} {mem : k ∈ (toInsert.map Sigma.fst)} : getKeyD k' (insertList l toInsert) fallback = k := by rw [getKeyD_eq_getKey?] @@ -2355,8 +2221,12 @@ theorem getKeyD_insertList_mem [BEq α] [EquivBEq α] . exact distinct2 . exact mem -theorem insertList_perm [BEq α] [ReflBEq α] [PartialEquivBEq α] (l toInsert: List ((a : α) × β a)) (distinct_l: DistinctKeys l) (distinct_toInsert: List.Pairwise (fun a b => (a.1 == b.1) = false) toInsert) (distinct_both: ∀ (a:α), ¬ (containsKey a l ∧ containsKey a toInsert)): - Perm (insertList l toInsert) (l++toInsert) := by +theorem insertList_perm [BEq α] [ReflBEq α] [PartialEquivBEq α] + (l toInsert : List ((a : α) × β a)) + (distinct_l : DistinctKeys l) + (distinct_toInsert : List.Pairwise (fun a b => (a.1 == b.1) = false) toInsert) + (distinct_both : ∀ (a : α), ¬ (containsKey a l ∧ containsKey a toInsert)) : + Perm (insertList l toInsert) (l ++ toInsert) := by rw [← DistinctKeys.def] at distinct_toInsert induction toInsert generalizing l with | nil => simp only [insertList, List.append_nil, Perm.refl] @@ -2375,9 +2245,9 @@ theorem insertList_perm [BEq α] [ReflBEq α] [PartialEquivBEq α] (l toInsert: specialize distinct_both a simp only [containsKey, Bool.or_eq_true, not_and, not_or, Bool.not_eq_true] at distinct_both simp only [containsKey_insertEntry, Bool.or_eq_true, not_and, Bool.not_eq_true] - by_cases hd_a: hd.fst == a + by_cases hd_a : hd.fst == a · rw [distinctKeys_cons_iff] at distinct_toInsert - have contains_a_tl: containsKey a tl = false := by + have contains_a_tl : containsKey a tl = false := by false_or_by_contra rename_i h simp only [Bool.not_eq_false] at h @@ -2392,20 +2262,23 @@ theorem insertList_perm [BEq α] [ReflBEq α] [PartialEquivBEq α] (l toInsert: specialize distinct_both hd.1 simp only [containsKey, BEq.refl, Bool.true_or, and_true, Bool.not_eq_true] at distinct_both simp only [distinct_both, cond_false, List.cons_append] - have h_hd: hd = ⟨hd.fst, hd.snd⟩:= rfl + have h_hd : hd = ⟨hd.fst, hd.snd⟩ := rfl rw [← h_hd] apply Perm.symm apply List.perm_middle -theorem insertList_not_isEmpty_if_start_not_isEmpty [BEq α]{l toInsert: List ((a : α) × β a)} {h:l.isEmpty = false}: +theorem insertList_not_isEmpty_if_start_not_isEmpty [BEq α] + {l toInsert : List ((a : α) × β a)} + {h : l.isEmpty = false} : (List.insertList l toInsert).isEmpty = false := by induction toInsert generalizing l with | nil => simp [insertList, h] | cons hd tl ih => simp [insertList, ih] - -theorem insertList_isEmpty [BEq α]{l toInsert: List ((a : α) × β a)}: (List.insertList l toInsert).isEmpty ↔ l.isEmpty ∧ toInsert.isEmpty := by +theorem insertList_isEmpty [BEq α] + {l toInsert : List ((a : α) × β a)} : + (List.insertList l toInsert).isEmpty ↔ l.isEmpty ∧ toInsert.isEmpty := by induction toInsert with | nil => simp [insertList] | cons hd tl ih => diff --git a/src/Std/Data/DHashMap/Internal/Model.lean b/src/Std/Data/DHashMap/Internal/Model.lean index 3c0dff49372c..a31b7c487697 100644 --- a/src/Std/Data/DHashMap/Internal/Model.lean +++ b/src/Std/Data/DHashMap/Internal/Model.lean @@ -188,7 +188,7 @@ theorem toListModel_updateAllBuckets {m : Raw₀ α β} {f : AssocList α β → omega rw [updateAllBuckets, toListModel, Array.toList_map, List.flatMap_eq_foldl, List.foldl_map, toListModel, List.flatMap_eq_foldl] - suffices ∀ (l : List (AssocList α β)) (l' : List ((a: α) × δ a)) (l'' : List ((a : α) × β a)), + suffices ∀ (l : List (AssocList α β)) (l' : List ((a : α) × δ a)) (l'' : List ((a : α) × β a)), Perm (g l'') l' → Perm (l.foldl (fun acc a => acc ++ (f a).toList) l') (g (l.foldl (fun acc a => acc ++ a.toList) l'')) by @@ -323,7 +323,7 @@ def filterₘ (m : Raw₀ α β) (f : (a : α) → β a → Bool) : Raw₀ α β ⟨withComputedSize (updateAllBuckets m.1.buckets fun l => l.filter f), by simpa using m.2⟩ /-- Internal implementation detail of the hash map -/ -def insertListₘ [BEq α] [Hashable α](m : Raw₀ α β) (l: List ((a : α) × β a)): Raw₀ α β := +def insertListₘ [BEq α] [Hashable α](m : Raw₀ α β) (l : List ((a : α) × β a)) : Raw₀ α β := match l with | .nil => m | .cons hd tl => insertListₘ (m.insert hd.1 hd.2) tl @@ -461,7 +461,7 @@ theorem map_eq_mapₘ (m : Raw₀ α β) (f : (a : α) → β a → δ a) : theorem filter_eq_filterₘ (m : Raw₀ α β) (f : (a : α) → β a → Bool) : m.filter f = m.filterₘ f := rfl -theorem insertMany_eq_insertListₘ [BEq α] [Hashable α](m : Raw₀ α β) (l: List ((a : α) × β a)): insertMany m l = insertListₘ m l := by +theorem insertMany_eq_insertListₘ [BEq α] [Hashable α](m : Raw₀ α β) (l : List ((a : α) × β a)) : insertMany m l = insertListₘ m l := by simp only [insertMany, Id.run, Id.pure_eq, Id.bind_eq, List.forIn_yield_eq_foldl] suffices ∀ (t : { m' // ∀ (P : Raw₀ α β → Prop), (∀ {m'' : Raw₀ α β} {a : α} {b : β a}, P m'' → P (m''.insert a b)) → P m → P m' }), diff --git a/src/Std/Data/DHashMap/Internal/Raw.lean b/src/Std/Data/DHashMap/Internal/Raw.lean index 2aa64a6a1371..4fcbb00d42a2 100644 --- a/src/Std/Data/DHashMap/Internal/Raw.lean +++ b/src/Std/Data/DHashMap/Internal/Raw.lean @@ -199,11 +199,11 @@ theorem filter_val [BEq α] [Hashable α] {m : Raw₀ α β} {f : (a : α) → m.val.filter f = m.filter f := by simp [Raw.filter, m.2] -theorem insertMany_eq [BEq α] [Hashable α] {m : Raw α β} (h : m.WF) {ρ : Type w} [ForIn Id ρ ((a : α) × β a)] {l : ρ}: +theorem insertMany_eq [BEq α] [Hashable α] {m : Raw α β} (h : m.WF) {ρ : Type w} [ForIn Id ρ ((a : α) × β a)] {l : ρ} : m.insertMany l = Raw₀.insertMany ⟨m, h.size_buckets_pos⟩ l := by simp[Raw.insertMany, h.size_buckets_pos] -theorem insertMany_val [BEq α][Hashable α] {m: Raw₀ α β} {ρ : Type w} [ForIn Id ρ ((a : α) × β a)] {l : ρ}: +theorem insertMany_val [BEq α][Hashable α] {m : Raw₀ α β} {ρ : Type w} [ForIn Id ρ ((a : α) × β a)] {l : ρ} : m.val.insertMany l = m.insertMany l := by simp[Raw.insertMany, m.2] diff --git a/src/Std/Data/DHashMap/Internal/RawLemmas.lean b/src/Std/Data/DHashMap/Internal/RawLemmas.lean index 649141988caf..981de485ca00 100644 --- a/src/Std/Data/DHashMap/Internal/RawLemmas.lean +++ b/src/Std/Data/DHashMap/Internal/RawLemmas.lean @@ -820,8 +820,8 @@ theorem length_keys [EquivBEq α] [LawfulHashable α] (h : m.1.WF) : simp_to_model using List.length_keys_eq_length @[simp] -theorem isEmpty_keys [EquivBEq α] [LawfulHashable α] (h: m.1.WF): - m.1.keys.isEmpty = m.1.isEmpty:= by +theorem isEmpty_keys [EquivBEq α] [LawfulHashable α] (h : m.1.WF) : + m.1.keys.isEmpty = m.1.isEmpty := by simp_to_model using List.isEmpty_keys_eq_isEmpty @[simp] @@ -840,161 +840,164 @@ theorem distinct_keys [EquivBEq α] [LawfulHashable α] (h : m.1.WF) : simp_to_model using (Raw.WF.out h).distinct.distinct @[simp] -theorem insertList_nil: +theorem insertList_nil : m.insertMany [] = m := by simp [insertMany, Id.run] @[simp] -theorem insertList_singleton [EquivBEq α] [LawfulHashable α] {k: α} {v: β k}: +theorem insertList_singleton [EquivBEq α] [LawfulHashable α] {k : α} {v : β k} : m.insertMany [⟨k,v⟩] = m.insert k v := by simp [insertMany, Id.run] @[simp] -theorem insertList_cons {l: List ((a:α) × (β a))} {k: α} {v: β k}: - (m.insertMany (⟨k,v⟩::l)).1 = ((m.insert k v).insertMany l).1 := by +theorem insertList_cons {l : List ((a : α) × (β a))} {k : α} {v : β k} : + (m.insertMany (⟨k,v⟩ :: l)).1 = ((m.insert k v).insertMany l).1 := by simp only [insertMany_eq_insertListₘ] cases l with | nil => simp [insertListₘ] | cons hd tl => simp [insertListₘ] -theorem contains_insertList [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l: List ((a:α) × (β a))} {k: α}: +theorem contains_insertList [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l : List ((a : α) × (β a))} {k : α} : (m.insertMany l).1.contains k ↔ m.contains k ∨ (l.map Sigma.fst).contains k := by simp_to_model using List.containsKey_insertList -theorem contains_of_contains_insertList [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l: List ((a:α) × (β a))} {k: α} : +theorem contains_of_contains_insertList [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l : List ((a : α) × (β a))} {k : α} : (m.insertMany l).1.contains k → (l.map Sigma.fst).contains k = false → m.contains k := by simp_to_model using List.containsKey_of_containsKey_insertList theorem get?_insertList_not_mem [LawfulBEq α] - {l: List ((a:α) × (β a))} {k : α} - {distinct: l.Pairwise (fun a b => (a.1 == b.1) = false)} - {not_mem: (l.map (Sigma.fst)).contains k = false} - (h: m.1.WF) : + {l : List ((a : α) × (β a))} {k : α} + {distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)} + {not_mem : (l.map (Sigma.fst)).contains k = false} + (h : m.1.WF) : (m.insertMany l).1.get? k = m.get? k := by simp_to_model using getValueCast?_insertList_not_mem theorem get?_insertList_mem [LawfulBEq α] - {l: List ((a:α) × (β a))} {k k': α} {k_eq: k == k'} {v: β k} - {distinct: l.Pairwise (fun a b => (a.1 == b.1) = false)} - {mem: ⟨k,v⟩ ∈ l} - (h: m.1.WF) : + {l : List ((a : α) × (β a))} {k k' : α} {k_eq : k == k'} {v : β k} + {distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)} + {mem : ⟨k,v⟩ ∈ l} + (h : m.1.WF) : (m.insertMany l).1.get? k' = some (cast (by congr; apply LawfulBEq.eq_of_beq k_eq) v) := by simp_to_model using getValueCast?_insertList_mem theorem get_insertList_not_mem [LawfulBEq α] - {l: List ((a:α) × (β a))} {k : α} - {distinct: l.Pairwise (fun a b => (a.1 == b.1) = false)} - {not_mem: (l.map (Sigma.fst)).contains k = false} + {l : List ((a : α) × (β a))} {k : α} + {distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)} + {not_mem : (l.map (Sigma.fst)).contains k = false} {h'} - (h: m.1.WF) : + (h : m.1.WF) : (m.insertMany l).1.get k h' = m.get k (contains_of_contains_insertList _ h h' not_mem) := by simp_to_model using getValueCast_insertList_not_mem theorem get_insertList_mem [LawfulBEq α] - {l: List ((a:α) × (β a))} {k k': α} {k_eq: k == k'} {v: β k} - {distinct: l.Pairwise (fun a b => (a.1 == b.1) = false)} - {mem: ⟨k,v⟩ ∈ l} + {l : List ((a : α) × (β a))} {k k' : α} {k_eq : k == k'} {v : β k} + {distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)} + {mem : ⟨k,v⟩ ∈ l} {h'} - (h: m.1.WF) : + (h : m.1.WF) : (m.insertMany l).1.get k' h' = cast (by congr; apply LawfulBEq.eq_of_beq k_eq) v := by simp_to_model using getValueCast_insertList_mem theorem get!_insertList_not_mem [LawfulBEq α] - {l: List ((a:α) × (β a))} {k : α} [Inhabited (β k)] - {distinct: l.Pairwise (fun a b => (a.1 == b.1) = false)} - {not_mem: (l.map (Sigma.fst)).contains k = false} - (h: m.1.WF) : + {l : List ((a : α) × (β a))} {k : α} [Inhabited (β k)] + {distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)} + {not_mem : (l.map (Sigma.fst)).contains k = false} + (h : m.1.WF) : (m.insertMany l).1.get! k = m.get! k := by simp_to_model using getValueCast!_insertList_not_mem theorem get!_insertList_mem [LawfulBEq α] - {l: List ((a:α) × (β a))} {k k': α} {k_eq: k == k'} {v: β k} [Inhabited (β k')] - {distinct: l.Pairwise (fun a b => (a.1 == b.1) = false)} - {mem: ⟨k,v⟩ ∈ l} - (h: m.1.WF) : + {l : List ((a : α) × (β a))} {k k' : α} {k_eq : k == k'} {v : β k} [Inhabited (β k')] + {distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)} + {mem : ⟨k,v⟩ ∈ l} + (h : m.1.WF) : (m.insertMany l).1.get! k' = cast (by congr; apply LawfulBEq.eq_of_beq k_eq) v := by simp_to_model using getValueCast!_insertList_mem theorem getD_insertList_not_mem [LawfulBEq α] - {l: List ((a:α) × (β a))} {k : α} {fallback : β k} - {distinct: l.Pairwise (fun a b => (a.1 == b.1) = false)} - {not_mem: (l.map (Sigma.fst)).contains k = false} - (h: m.1.WF) : + {l : List ((a : α) × (β a))} {k : α} {fallback : β k} + {distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)} + {not_mem : (l.map (Sigma.fst)).contains k = false} + (h : m.1.WF) : (m.insertMany l).1.getD k fallback = m.getD k fallback := by simp_to_model using getValueCastD_insertList_not_mem theorem getD_insertList_mem [LawfulBEq α] - {l: List ((a:α) × (β a))} {k k': α} {k_eq: k == k'} {v: β k} {fallback : β k'} - {distinct: l.Pairwise (fun a b => (a.1 == b.1) = false)} - {mem: ⟨k,v⟩ ∈ l} - (h: m.1.WF) : + {l : List ((a : α) × (β a))} {k k' : α} {k_eq : k == k'} {v : β k} {fallback : β k'} + {distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)} + {mem : ⟨k,v⟩ ∈ l} + (h : m.1.WF) : (m.insertMany l).1.getD k' fallback = cast (by congr; apply LawfulBEq.eq_of_beq k_eq) v := by simp_to_model using getValueCastD_insertList_mem theorem getKey?_insertList_not_mem [EquivBEq α] [LawfulHashable α] (h : m.1.WF) - {l : List ((a:α) × (β a))} {k : α} - {distinct: l.Pairwise (fun a b => (a.1 == b.1) = false)} + {l : List ((a : α) × (β a))} {k : α} + {distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)} {not_mem : (l.map Sigma.fst).contains k = false} : (m.insertMany l).1.getKey? k = m.getKey? k := by simp_to_model using List.getKey?_insertList_not_mem theorem getKey?_insertList_mem [EquivBEq α] [LawfulHashable α] (h : m.1.WF) - {l : List ((a:α) × (β a))} + {l : List ((a : α) × (β a))} {k k' : α} {k_beq : k == k'} - {distinct: l.Pairwise (fun a b => (a.1 == b.1) = false)} + {distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)} {mem : k ∈ (l.map Sigma.fst)} : (m.insertMany l).1.getKey? k' = some k := by simp_to_model using List.getKey?_insertList_mem theorem getKey_insertList_not_mem [EquivBEq α] [LawfulHashable α] (h : m.1.WF) - {l : List ((a:α) × (β a))} {k : α} - {distinct: l.Pairwise (fun a b => (a.1 == b.1) = false)} + {l : List ((a : α) × (β a))} {k : α} + {distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)} {not_mem : (l.map Sigma.fst).contains k = false} {h'} : (m.insertMany l).1.getKey k h' = m.getKey k (contains_of_contains_insertList _ h h' not_mem) := by simp_to_model using List.getKey_insertList_not_mem theorem getKey_insertList_mem [EquivBEq α] [LawfulHashable α] (h : m.1.WF) - {l : List ((a:α) × (β a))} + {l : List ((a : α) × (β a))} {k k' : α} {k_beq : k == k'} - {distinct: l.Pairwise (fun a b => (a.1 == b.1) = false)} + {distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)} {mem : k ∈ (l.map Sigma.fst)} {h'} : (m.insertMany l).1.getKey k' h' = k := by simp_to_model using List.getKey_insertList_mem theorem getKey!_insertList_not_mem [EquivBEq α] [LawfulHashable α] [Inhabited α] (h : m.1.WF) - {l : List ((a:α) × (β a))} {k : α} - {distinct: l.Pairwise (fun a b => (a.1 == b.1) = false)} + {l : List ((a : α) × (β a))} {k : α} + {distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)} {not_mem : (l.map Sigma.fst).contains k = false} : (m.insertMany l).1.getKey! k = m.getKey! k := by simp_to_model using List.getKey!_insertList_not_mem theorem getKey!_insertList_mem [EquivBEq α] [LawfulHashable α] [Inhabited α] (h : m.1.WF) - {l : List ((a:α) × (β a))} + {l : List ((a : α) × (β a))} {k k' : α} {k_beq : k == k'} - {distinct: l.Pairwise (fun a b => (a.1 == b.1) = false)} + {distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)} {mem : k ∈ (l.map Sigma.fst)} : (m.insertMany l).1.getKey! k' = k := by simp_to_model using List.getKey!_insertList_mem theorem getKeyD_insertList_not_mem [EquivBEq α] [LawfulHashable α] (h : m.1.WF) - {l : List ((a:α) × (β a))} {k fallback : α} - {distinct: l.Pairwise (fun a b => (a.1 == b.1) = false)} + {l : List ((a : α) × (β a))} {k fallback : α} + {distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)} {not_mem : (l.map Sigma.fst).contains k = false} : (m.insertMany l).1.getKeyD k fallback = m.getKeyD k fallback := by simp_to_model using List.getKeyD_insertList_not_mem theorem getKeyD_insertList_mem [EquivBEq α] [LawfulHashable α] (h : m.1.WF) - {l : List ((a:α) × (β a))} + {l : List ((a : α) × (β a))} {k k' fallback : α} {k_beq : k == k'} - {distinct: l.Pairwise (fun a b => (a.1 == b.1) = false)} + {distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)} {mem : k ∈ (l.map Sigma.fst)} : (m.insertMany l).1.getKeyD k' fallback = k := by simp_to_model using List.getKeyD_insertList_mem -theorem size_insertList [EquivBEq α] [LawfulHashable α] {l: List ((a:α) × (β a))} {distinct: List.Pairwise (fun a b => (a.1 == b.1) = false) l} (h: m.1.WF): - (∀ (a:α), ¬ (m.contains a = true ∧ List.containsKey a l = true)) → +theorem size_insertList [EquivBEq α] [LawfulHashable α] + {l : List ((a : α) × (β a))} + {distinct : List.Pairwise (fun a b => (a.1 == b.1) = false) l} + (h : m.1.WF) : + (∀ (a : α), ¬ (m.contains a = true ∧ List.containsKey a l = true)) → (m.insertMany l).1.1.size = m.1.size + l.length := by simp_to_model rw [← List.length_append] @@ -1005,11 +1008,13 @@ theorem size_insertList [EquivBEq α] [LawfulHashable α] {l: List ((a:α) × ( . exact distinct . apply distinct' -theorem insertList_notEmpty_if_m_notEmpty [EquivBEq α] [LawfulHashable α] {l: List ((a:α) × (β a))}(h: m.1.WF): +theorem insertList_notEmpty_if_m_notEmpty [EquivBEq α] [LawfulHashable α] + {l : List ((a : α) × (β a))} (h : m.1.WF) : (m.1.isEmpty = false) → (m.insertMany l).1.1.isEmpty = false := by simp_to_model using List.insertList_not_isEmpty_if_start_not_isEmpty -theorem insertList_isEmpty [EquivBEq α] [LawfulHashable α] {l: List ((a:α) × (β a))}(h: m.1.WF): +theorem insertList_isEmpty [EquivBEq α] [LawfulHashable α] + {l : List ((a : α) × (β a))} (h : m.1.WF) : (m.insertMany l).1.1.isEmpty ↔ m.1.isEmpty ∧ l.isEmpty := by simp_to_model using List.insertList_isEmpty diff --git a/src/Std/Data/DHashMap/Internal/WF.lean b/src/Std/Data/DHashMap/Internal/WF.lean index d6c324d3733f..71a03a0ca860 100644 --- a/src/Std/Data/DHashMap/Internal/WF.lean +++ b/src/Std/Data/DHashMap/Internal/WF.lean @@ -540,7 +540,7 @@ theorem Const.toListModel_getThenInsertIfNew? {β : Type v} [BEq α] [Hashable exact toListModel_insertIfNewₘ h theorem Const.wfImp_getThenInsertIfNew? {β : Type v} [BEq α] [Hashable α] [EquivBEq α] - [LawfulHashable α] {m : Raw₀ α (fun _ => β)} {a : α} {b : β} (h : Raw.WFImp m.1): + [LawfulHashable α] {m : Raw₀ α (fun _ => β)} {a : α} {b : β} (h : Raw.WFImp m.1) : Raw.WFImp (Const.getThenInsertIfNew? m a b).2.1 := by rw [getThenInsertIfNew?_eq_insertIfNewₘ] exact wfImp_insertIfNewₘ h @@ -707,7 +707,7 @@ theorem wfImp_filter [BEq α] [Hashable α] [EquivBEq α] [LawfulHashable α] {m exact wfImp_filterₘ h /-! # `insertListₘ` -/ -theorem wfImp_insertListₘ [BEq α] [Hashable α] [EquivBEq α][LawfulHashable α] {m : Raw₀ α β} {l: List ((a : α) × β a)} (h : Raw.WFImp m.1): Raw.WFImp (m.insertListₘ l).1 := by +theorem wfImp_insertListₘ [BEq α] [Hashable α] [EquivBEq α][LawfulHashable α] {m : Raw₀ α β} {l : List ((a : α) × β a)} (h : Raw.WFImp m.1) : Raw.WFImp (m.insertListₘ l).1 := by induction l generalizing m with | nil => simp only [insertListₘ] @@ -717,7 +717,7 @@ theorem wfImp_insertListₘ [BEq α] [Hashable α] [EquivBEq α][LawfulHashable apply ih apply wfImp_insert h -theorem toListModel_insertListₘ [BEq α] [Hashable α] [EquivBEq α][LawfulHashable α] {m : Raw₀ α β} {l: List ((a : α) × β a)} (h : Raw.WFImp m.1): +theorem toListModel_insertListₘ [BEq α] [Hashable α] [EquivBEq α][LawfulHashable α] {m : Raw₀ α β} {l : List ((a : α) × β a)} (h : Raw.WFImp m.1) : Perm (toListModel (insertListₘ m l).1.buckets) (List.insertList (toListModel m.1.buckets) l) := by induction l generalizing m with | nil => @@ -757,11 +757,12 @@ theorem wfImp_insertMany [BEq α] [Hashable α] [EquivBEq α] [LawfulHashable α Raw.WFImp (m.insertMany l).1.1 := Raw.WF.out ((m.insertMany l).2 _ Raw.WF.insert₀ (.wf m.2 h)) -theorem toListModel_insertMany [BEq α] [Hashable α] [EquivBEq α] [LawfulHashable α] {m : Raw₀ α β} {l: List ((a:α) × (β a))} (h : Raw.WFImp m.1): - Perm (toListModel (insertMany m l).1.1.buckets) (List.insertList (toListModel m.1.buckets) l) := by - rw[insertMany_eq_insertListₘ] - apply toListModel_insertListₘ - exact h +theorem toListModel_insertMany [BEq α] [Hashable α] [EquivBEq α] [LawfulHashable α] + {m : Raw₀ α β} {l : List ((a : α) × (β a))} (h : Raw.WFImp m.1) : + Perm (toListModel (insertMany m l).1.1.buckets) (List.insertList (toListModel m.1.buckets) l) := by + rw[insertMany_eq_insertListₘ] + apply toListModel_insertListₘ + exact h theorem Const.wfImp_insertMany {β : Type v} [BEq α] [Hashable α] [EquivBEq α] [LawfulHashable α] {ρ : Type w} [ForIn Id ρ (α × β)] {m : Raw₀ α (fun _ => β)} diff --git a/src/Std/Data/DHashMap/Raw.lean b/src/Std/Data/DHashMap/Raw.lean index dc6a4e0e2250..0f92a8c194ba 100644 --- a/src/Std/Data/DHashMap/Raw.lean +++ b/src/Std/Data/DHashMap/Raw.lean @@ -94,7 +94,7 @@ Checks whether a key is present in a map, and unconditionally inserts a value fo Equivalent to (but potentially faster than) calling `contains` followed by `insert`. -/ @[inline] def containsThenInsert [BEq α] [Hashable α] (m : Raw α β) (a : α) (b : β a) : - Bool × Raw α β:= + Bool × Raw α β := if h : 0 < m.buckets.size then let ⟨replaced, ⟨r, _⟩⟩ := Raw₀.containsThenInsert ⟨m, h⟩ a b ⟨replaced, r⟩ diff --git a/src/Std/Data/DHashMap/RawLemmas.lean b/src/Std/Data/DHashMap/RawLemmas.lean index 9dbb0e3b2d0d..8b8bb67a68db 100644 --- a/src/Std/Data/DHashMap/RawLemmas.lean +++ b/src/Std/Data/DHashMap/RawLemmas.lean @@ -765,7 +765,7 @@ theorem getKey!_eq_default_of_contains_eq_false [EquivBEq α] [LawfulHashable α m.contains a = false → m.getKey! a = default := by simp_to_raw using Raw₀.getKey!_eq_default -theorem getKey!_eq_default [EquivBEq α] [LawfulHashable α] [Inhabited α] (h : m.WF) {a : α}: +theorem getKey!_eq_default [EquivBEq α] [LawfulHashable α] [Inhabited α] (h : m.WF) {a : α} : ¬a ∈ m → m.getKey! a = default := by simpa [mem_iff_contains] using getKey!_eq_default_of_contains_eq_false h @@ -1028,7 +1028,7 @@ theorem length_keys [EquivBEq α] [LawfulHashable α] (h : m.WF) : simp_to_raw using Raw₀.length_keys ⟨m, h.size_buckets_pos⟩ h @[simp] -theorem isEmpty_keys [EquivBEq α] [LawfulHashable α] (h : m.WF): +theorem isEmpty_keys [EquivBEq α] [LawfulHashable α] (h : m.WF) : m.keys.isEmpty = m.isEmpty := by simp_to_raw using Raw₀.isEmpty_keys ⟨m, h.size_buckets_pos⟩ From 8f3eaa2200d565c16c56742bea8d2e1bd90b0463 Mon Sep 17 00:00:00 2001 From: jt0202 Date: Sat, 30 Nov 2024 16:18:35 +0100 Subject: [PATCH 29/83] constant get for insertMany --- .../DHashMap/Internal/List/Associative.lean | 208 ++++++++++++++---- src/Std/Data/DHashMap/Internal/Model.lean | 38 ++++ src/Std/Data/DHashMap/Internal/RawLemmas.lean | 120 +++++++--- src/Std/Data/DHashMap/Internal/WF.lean | 21 +- 4 files changed, 304 insertions(+), 83 deletions(-) diff --git a/src/Std/Data/DHashMap/Internal/List/Associative.lean b/src/Std/Data/DHashMap/Internal/List/Associative.lean index 63c173683d8f..c539f9c603f7 100644 --- a/src/Std/Data/DHashMap/Internal/List/Associative.lean +++ b/src/Std/Data/DHashMap/Internal/List/Associative.lean @@ -1901,7 +1901,7 @@ theorem DistinctKeys.insertList [BEq α] [PartialEquivBEq α] {l₁ l₂ : List rw [insertList.eq_def] exact ih h.insertEntry -theorem insertList_perm_of_perm_first [BEq α] [EquivBEq α] (l1 l2 toInsert : List ((a : α) × β a)) (h : Perm l1 l2) +theorem insertList_perm_of_perm_first [BEq α] [EquivBEq α] {l1 l2 toInsert : List ((a : α) × β a)} (h : Perm l1 l2) (distinct : DistinctKeys l1) : Perm (insertList l1 toInsert) (insertList l2 toInsert) := by induction toInsert generalizing l1 l2 with | nil => simp[insertList, h] @@ -1962,22 +1962,17 @@ theorem containsKey_of_containsKey_insertList [BEq α] [PartialEquivBEq α] theorem getValueCast?_insertList_not_mem [BEq α] [LawfulBEq α] {l toInsert : List ((a : α) × β a)} {k : α} - {distinct : toInsert.Pairwise (fun a b => (a.1 == b.1) = false)} - {distinct2 : DistinctKeys l} {not_mem : (toInsert.map Sigma.fst).contains k = false} : getValueCast? k (insertList l toInsert) = getValueCast? k l := by - rw [← DistinctKeys.def] at distinct rw [← containsKey_eq_contains_map_fst] at not_mem induction toInsert generalizing l with | nil => simp [insertList] | cons hd tl ih => rw [containsKey_cons, Bool.or_eq_false_iff] at not_mem - rw [getValueCast?_of_perm distinct2.insertList (insertList_cons_perm l tl hd distinct2 distinct)] + simp only [insertList] + rw [ih (And.right not_mem)] rw [getValueCast?_insertEntry] - rw [ih distinct.tail] - . simp [not_mem.left] - . exact not_mem.right - . exact distinct2 + simp [And.left not_mem] theorem getValueCast?_insertList_mem [BEq α] [LawfulBEq α] {l toInsert : List ((a : α) × β a)} @@ -2008,8 +2003,6 @@ theorem getValueCast?_insertList_mem [BEq α] [LawfulBEq α] theorem getValueCast_insertList_not_mem [BEq α] [LawfulBEq α] {l toInsert : List ((a : α) × β a)} {k : α} - {distinct : toInsert.Pairwise (fun a b => (a.1 == b.1) = false)} - {distinct2 : DistinctKeys l} {not_mem : (toInsert.map Sigma.fst).contains k = false} {h} : getValueCast k (insertList l toInsert) h = getValueCast k l (containsKey_of_containsKey_insertList h not_mem) := by @@ -2017,9 +2010,7 @@ theorem getValueCast_insertList_not_mem [BEq α] [LawfulBEq α] rw [← getValueCast?_eq_some_getValueCast] rw [← getValueCast?_eq_some_getValueCast] rw [getValueCast?_insertList_not_mem] - . exact distinct - . exact distinct2 - . exact not_mem + exact not_mem theorem getValueCast_insertList_mem [BEq α] [LawfulBEq α] {l toInsert : List ((a : α) × β a)} @@ -2039,14 +2030,10 @@ theorem getValueCast_insertList_mem [BEq α] [LawfulBEq α] theorem getValueCast!_insertList_not_mem [BEq α] [LawfulBEq α] {l toInsert : List ((a : α) × β a)} {k : α} [Inhabited (β k)] - {distinct : toInsert.Pairwise (fun a b => (a.1 == b.1) = false)} - {distinct2 : DistinctKeys l} {not_mem : (toInsert.map Sigma.fst).contains k = false} : getValueCast! k (insertList l toInsert) = getValueCast! k l := by rw [getValueCast!_eq_getValueCast?, getValueCast!_eq_getValueCast?, getValueCast?_insertList_not_mem] - . exact distinct - . exact distinct2 - . exact not_mem + exact not_mem theorem getValueCast!_insertList_mem [BEq α] [LawfulBEq α] {l toInsert : List ((a : α) × β a)} @@ -2063,14 +2050,10 @@ theorem getValueCast!_insertList_mem [BEq α] [LawfulBEq α] theorem getValueCastD_insertList_not_mem [BEq α] [LawfulBEq α] {l toInsert : List ((a : α) × β a)} {k : α} {fallback : β k} - {distinct : toInsert.Pairwise (fun a b => (a.1 == b.1) = false)} - {distinct2 : DistinctKeys l} {not_mem : (toInsert.map Sigma.fst).contains k = false} : getValueCastD k (insertList l toInsert) fallback = getValueCastD k l fallback := by rw [getValueCastD_eq_getValueCast?, getValueCastD_eq_getValueCast?, getValueCast?_insertList_not_mem] - . exact distinct - . exact distinct2 - . exact not_mem + exact not_mem theorem getValueCastD_insertList_mem [BEq α] [LawfulBEq α] {l toInsert : List ((a : α) × β a)} @@ -2087,22 +2070,19 @@ theorem getValueCastD_insertList_mem [BEq α] [LawfulBEq α] theorem getKey?_insertList_not_mem [BEq α] [EquivBEq α] {l toInsert : List ((a : α) × β a)} {k : α} - {distinct : toInsert.Pairwise (fun a b => (a.1 == b.1) = false)} - {distinct2 : DistinctKeys l} {not_mem : (toInsert.map Sigma.fst).contains k = false} : getKey? k (insertList l toInsert) = getKey? k l := by - rw [← DistinctKeys.def] at distinct rw [← containsKey_eq_contains_map_fst] at not_mem induction toInsert generalizing l with | nil => simp [insertList] | cons hd tl ih => - rw [containsKey_cons, Bool.or_eq_false_iff] at not_mem - rw [getKey?_of_perm distinct2.insertList (insertList_cons_perm l tl hd distinct2 distinct)] - rw [getKey?_insertEntry] - rw [ih distinct.tail] - . simp [not_mem.left] - . exact not_mem.right - . exact distinct2 + simp only [insertList] + rw [ih] + · rw [getKey?_insertEntry] + simp only [containsKey, Bool.or_eq_false_iff] at not_mem + simp [And.left not_mem] + · simp only [containsKey, Bool.or_eq_false_iff] at not_mem + apply And.right not_mem theorem getKey?_insertList_mem [BEq α] [EquivBEq α] {l toInsert : List ((a : α) × β a)} @@ -2135,8 +2115,6 @@ theorem getKey?_insertList_mem [BEq α] [EquivBEq α] theorem getKey_insertList_not_mem [BEq α] [EquivBEq α] {l toInsert : List ((a : α) × β a)} {k : α} - {distinct : toInsert.Pairwise (fun a b => (a.1 == b.1) = false)} - {distinct2 : DistinctKeys l} {not_mem : (toInsert.map Sigma.fst).contains k = false} {h} : getKey k (insertList l toInsert) h = @@ -2145,8 +2123,6 @@ theorem getKey_insertList_not_mem [BEq α] [EquivBEq α] rw [← getKey?_eq_some_getKey] rw [getKey?_insertList_not_mem] . rw [getKey?_eq_some_getKey] - . exact distinct - . exact distinct2 . exact not_mem theorem getKey_insertList_mem [BEq α] [EquivBEq α] @@ -2167,15 +2143,11 @@ theorem getKey_insertList_mem [BEq α] [EquivBEq α] theorem getKey!_insertList_not_mem [BEq α] [EquivBEq α] [Inhabited α] {l toInsert : List ((a : α) × β a)} {k : α} - {distinct : toInsert.Pairwise (fun a b => (a.1 == b.1) = false)} - {distinct2 : DistinctKeys l} {not_mem : (toInsert.map Sigma.fst).contains k = false} : getKey! k (insertList l toInsert) = getKey! k l := by rw [getKey!_eq_getKey?] rw [getKey?_insertList_not_mem] . rw [getKey!_eq_getKey?] - . exact distinct - . exact distinct2 . exact not_mem theorem getKey!_insertList_mem [BEq α] [EquivBEq α] [Inhabited α] @@ -2195,15 +2167,11 @@ theorem getKey!_insertList_mem [BEq α] [EquivBEq α] [Inhabited α] theorem getKeyD_insertList_not_mem [BEq α] [EquivBEq α] {l toInsert : List ((a : α) × β a)} {k fallback : α} - {distinct : toInsert.Pairwise (fun a b => (a.1 == b.1) = false)} - {distinct2 : DistinctKeys l} {not_mem : (toInsert.map Sigma.fst).contains k = false} : getKeyD k (insertList l toInsert) fallback = getKeyD k l fallback := by rw [getKeyD_eq_getKey?] rw [getKey?_insertList_not_mem] . rw [getKeyD_eq_getKey?] - . exact distinct - . exact distinct2 . exact not_mem theorem getKeyD_insertList_mem [BEq α] [EquivBEq α] @@ -2288,5 +2256,153 @@ theorem insertList_isEmpty [BEq α] apply insertList_not_isEmpty_if_start_not_isEmpty simp +section +variable {β: Type v} +theorem getValue?_insertList_not_mem [BEq α] [PartialEquivBEq α] + {l toInsert : List ((_ : α) × β)} {k : α} + {not_mem : (toInsert.map Sigma.fst).contains k = false} : + getValue? k (insertList l toInsert) = getValue? k l := by + induction toInsert generalizing l with + | nil => simp [insertList] + | cons hd tl ih => + simp only [insertList] + rw [ih] + · apply getValue?_insertEntry_of_false + simp only [List.map_cons, List.contains_cons, Bool.or_eq_false_iff] at not_mem + apply BEq.symm_false + apply (And.left not_mem) + · simp only [List.map_cons, List.contains_cons, Bool.or_eq_false_iff] at not_mem + apply And.right not_mem + +theorem getValue?_insertList_mem [BEq α] [PartialEquivBEq α] + {l toInsert : List ((_ : α) × β )} + {k k' : α} {k_beq : k == k'} {v : β} + {distinct : List.Pairwise (fun a b => (a.1 == b.1) = false) toInsert} + {mem : ⟨k,v⟩ ∈ toInsert} : + getValue? k' (insertList l toInsert) = v := by + induction toInsert generalizing l with + | nil => simp at mem + | cons hd tl ih => + simp only [insertList] + simp only [List.mem_cons] at mem + cases mem with + | inl mem => + rw [getValue?_insertList_not_mem] + simp only [← mem] + · rw [getValue?_insertEntry_of_beq] + exact k_beq + · simp only [pairwise_cons] at distinct + false_or_by_contra + rename_i h + simp only [Bool.not_eq_false, List.contains_iff_exists_mem_beq, List.mem_map] at h + rcases h with ⟨a', h, k_a'⟩ + rcases h with ⟨a, a_mem, a_fst⟩ + rcases distinct with ⟨distinct, _⟩ + specialize distinct a a_mem + rw [← mem, a_fst] at distinct + simp only at distinct + have k_a': k == a' := by + apply PartialEquivBEq.trans k_beq k_a' + rw [k_a'] at distinct + simp at distinct + | inr mem => + apply ih + · simp only [pairwise_cons] at distinct + apply And.right distinct + · exact mem + +theorem getValue!_insertList_not_mem [BEq α] [PartialEquivBEq α] [Inhabited β] + {l toInsert : List ((_ : α) × β)} {k : α} + {not_mem : (toInsert.map Sigma.fst).contains k = false} : + getValue! k (insertList l toInsert) = getValue! k l := by + simp only [getValue!_eq_getValue?] + rw [getValue?_insertList_not_mem] + exact not_mem + +theorem getValue!_insertList_mem [BEq α] [PartialEquivBEq α] [Inhabited β] + {l toInsert : List ((_ : α) × β)} {k k' : α} {v: β} {k_eq : k == k'} {distinct : List.Pairwise (fun a b => (a.1 == b.1) = false) toInsert } {mem : ⟨k,v⟩ ∈ toInsert}: + getValue! k' (insertList l toInsert) = v := by + simp only [getValue!_eq_getValue?] + rw [getValue?_insertList_mem (mem:= mem)] + · simp + · exact k_eq + · exact distinct + +theorem getValueD_insertList_not_mem [BEq α] [PartialEquivBEq α] + {l toInsert : List ((_ : α) × β)} {k : α} {fallback : β} + {not_mem : (toInsert.map Sigma.fst).contains k = false} : + getValueD k (insertList l toInsert) fallback = getValueD k l fallback := by + simp only [getValueD_eq_getValue?] + rw [getValue?_insertList_not_mem] + exact not_mem + +theorem getValueD_insertList_mem [BEq α] [PartialEquivBEq α] + {l toInsert : List ((_ : α) × β)} {k k' : α} {v fallback: β} {k_eq : k == k'} {distinct : List.Pairwise (fun a b => (a.1 == b.1) = false) toInsert } {mem : ⟨k,v⟩ ∈ toInsert}: + getValueD k' (insertList l toInsert) fallback= v := by + simp only [getValueD_eq_getValue?] + rw [getValue?_insertList_mem (mem:= mem)] + · simp + · exact k_eq + · exact distinct + +theorem getValue_insertList_not_mem [BEq α] [PartialEquivBEq α] + {l toInsert : List ((_ : α) × β)} {k : α} + {not_mem : (toInsert.map Sigma.fst).contains k = false} + {h} : + getValue k (insertList l toInsert) h = getValue k l (containsKey_of_containsKey_insertList h not_mem) := by + rw [← Option.some_inj] + rw [← getValue?_eq_some_getValue] + rw [← getValue?_eq_some_getValue] + rw [getValue?_insertList_not_mem] + exact not_mem + +theorem getValue_insertList_mem [BEq α] [PartialEquivBEq α] + {l toInsert : List ((_ : α) × β)} + {k k' : α} {k_beq : k == k'} {v : β} + {distinct : toInsert.Pairwise (fun a b => (a.1 == b.1) = false)} + {mem : ⟨k,v⟩ ∈ toInsert} + {h} : + getValue k' (insertList l toInsert) h = v := by + rw [← Option.some_inj] + rw [← getValue?_eq_some_getValue] + rw [getValue?_insertList_mem (mem:=mem)] + . exact k_beq + . exact distinct + +/-- Internal implementation detail of the hash map -/ +def insertListConst [BEq α] (l: List ((_ : α) × β))(toInsert: List (α × β)) : List ((_ : α) × β) := + match toInsert with + | .nil => l + | .cons hd tl => insertListConst (insertEntry hd.1 hd.2 l) tl + +theorem insertListConst_eq_insertList [BEq α] {l: List ((_ : α) × β)} {toInsert: List (α × β)}: + insertListConst l toInsert = insertList l (List.map (fun a => ⟨a.1, a.2⟩) toInsert) := by + induction toInsert generalizing l with + | nil => simp [insertList, insertListConst] + | cons hd tl ih => + simp [insertList, insertListConst] + apply ih + +theorem DistinctKeys.insertListConst [BEq α] [PartialEquivBEq α] {l₁ : List ((_ : α) × β)} {l₂ : List (α × β)} (h : DistinctKeys l₁) : + DistinctKeys (insertListConst l₁ l₂) := by + rw [insertListConst_eq_insertList] + apply DistinctKeys.insertList h + +theorem insertListConst_perm_of_perm_first [BEq α] [EquivBEq α] (l1 l2: List ((_ : α) × β)) (toInsert : List (α × β)) (h : Perm l1 l2) + (distinct : DistinctKeys l1) : Perm (insertListConst l1 toInsert) (insertListConst l2 toInsert) := by + simp only [insertListConst_eq_insertList] + apply insertList_perm_of_perm_first h distinct + +theorem containsKey_insertListConst [BEq α] [PartialEquivBEq α] {l : List ((_ : α) × β)} {toInsert: List (α × β)} {k : α} : + containsKey k (insertListConst l toInsert) ↔ containsKey k l ∨ (toInsert.map Prod.fst).contains k := by + rw [insertListConst_eq_insertList, containsKey_insertList] + have h: (List.map Sigma.fst (List.map (fun a => (⟨a.fst, a.snd⟩ : (_ : α ) × β )) toInsert)) = (List.map Prod.fst toInsert) := by + apply List.ext_get + · simp + · intro n h1 h2 + simp + rw [h] + +end end List diff --git a/src/Std/Data/DHashMap/Internal/Model.lean b/src/Std/Data/DHashMap/Internal/Model.lean index a31b7c487697..a5e6f27bed53 100644 --- a/src/Std/Data/DHashMap/Internal/Model.lean +++ b/src/Std/Data/DHashMap/Internal/Model.lean @@ -348,6 +348,17 @@ def Const.getDₘ [BEq α] [Hashable α] (m : Raw₀ α (fun _ => β)) (a : α) def Const.get!ₘ [BEq α] [Hashable α] [Inhabited β] (m : Raw₀ α (fun _ => β)) (a : α) : β := (Const.get?ₘ m a).get! +/-- Internal implementation detail of the hash map -/ +def Const.insertListₘ [BEq α] [Hashable α](m : Raw₀ α (fun _ => β)) (l: List (α × β)): Raw₀ α (fun _ => β) := + match l with + | .nil => m + | .cons hd tl => insertListₘ (m.insert hd.1 hd.2) tl + +/-- Internal implementation detail of the hash map -/ +def Const.insertListIfNewUnitₘ [BEq α] [Hashable α](m : Raw₀ α (fun _ => Unit)) (l: List α): Raw₀ α (fun _ => Unit) := + match l with + | .nil => m + | .cons hd tl => insertListIfNewUnitₘ (m.insertIfNew hd ()) tl end /-! # Equivalence between model functions and real implementations -/ @@ -504,6 +515,33 @@ theorem Const.getThenInsertIfNew?_eq_get?ₘ [BEq α] [Hashable α] (m : Raw₀ dsimp only [Array.ugetElem_eq_getElem, Array.uset] split <;> simp_all [-getValue?_eq_none] +theorem Const.insertMany_eq_insertListₘ [BEq α] [Hashable α] (m : Raw₀ α (fun _ => β)) (l: List (α × β)): + (Const.insertMany m l).1 = Const.insertListₘ m l := by + simp only [insertMany, Id.run, Id.pure_eq, Id.bind_eq, List.forIn_yield_eq_foldl] + suffices ∀ (t : { m' // ∀ (P : Raw₀ α (fun _ => β) → Prop), + (∀ {m'' : Raw₀ α (fun _ => β)} {a : α} {b : β}, P m'' → P (m''.insert a b)) → P m → P m' }), + (List.foldl (fun m' p => ⟨m'.val.insert p.1 p.2, fun P h₁ h₂ => h₁ (m'.2 _ h₁ h₂)⟩) t l).val = + Const.insertListₘ t.val l from this _ + intro t + induction l generalizing m with + | nil => simp[insertListₘ] + | cons hd tl ih => + simp only [List.foldl_cons,insertListₘ] + apply ih + +theorem Const.insertManyIfNewUnit_insertListIfNewUnit [BEq α] [Hashable α] (m : Raw₀ α (fun _ => Unit)) (l: List α): + (Const.insertManyIfNewUnit m l).1 = Const.insertListIfNewUnitₘ m l := by + simp only [insertManyIfNewUnit, Id.run, Id.pure_eq, Id.bind_eq, List.forIn_yield_eq_foldl] + suffices ∀ (t : { m' // ∀ (P : Raw₀ α (fun _ => Unit) → Prop), + (∀ {m'' a b}, P m'' → P (m''.insertIfNew a b)) → P m → P m'}), + (List.foldl (fun m' p => ⟨m'.val.insertIfNew p (), fun P h₁ h₂ => h₁ (m'.2 _ h₁ h₂)⟩) t l).val = + Const.insertListIfNewUnitₘ t.val l from this _ + intro t + induction l generalizing m with + | nil => simp[insertListIfNewUnitₘ] + | cons hd tl ih => + simp only [List.foldl_cons,insertListIfNewUnitₘ] + apply ih end end Raw₀ diff --git a/src/Std/Data/DHashMap/Internal/RawLemmas.lean b/src/Std/Data/DHashMap/Internal/RawLemmas.lean index 981de485ca00..47f81862dba0 100644 --- a/src/Std/Data/DHashMap/Internal/RawLemmas.lean +++ b/src/Std/Data/DHashMap/Internal/RawLemmas.lean @@ -90,7 +90,8 @@ private def queryNames : Array Name := ``Raw.pairwise_keys_iff_pairwise_keys] private def modifyNames : Array Name := - #[``toListModel_insert, ``toListModel_erase, ``toListModel_insertIfNew, ``toListModel_insertMany] + #[``toListModel_insert, ``toListModel_erase, ``toListModel_insertIfNew, ``toListModel_insertMany_list, + ``Const.toListModel_insertMany_list] private def congrNames : MacroM (Array (TSyntax `term)) := do return #[← `(_root_.List.Perm.isEmpty_eq), ← `(containsKey_of_perm), @@ -840,40 +841,38 @@ theorem distinct_keys [EquivBEq α] [LawfulHashable α] (h : m.1.WF) : simp_to_model using (Raw.WF.out h).distinct.distinct @[simp] -theorem insertList_nil : +theorem insertMany_list_nil : m.insertMany [] = m := by simp [insertMany, Id.run] @[simp] -theorem insertList_singleton [EquivBEq α] [LawfulHashable α] {k : α} {v : β k} : +theorem insertMany_list_singleton [EquivBEq α] [LawfulHashable α] {k : α} {v : β k} : m.insertMany [⟨k,v⟩] = m.insert k v := by simp [insertMany, Id.run] -@[simp] -theorem insertList_cons {l : List ((a : α) × (β a))} {k : α} {v : β k} : +theorem insertMany_list_cons {l : List ((a : α) × (β a))} {k : α} {v : β k} : (m.insertMany (⟨k,v⟩ :: l)).1 = ((m.insert k v).insertMany l).1 := by simp only [insertMany_eq_insertListₘ] cases l with | nil => simp [insertListₘ] | cons hd tl => simp [insertListₘ] -theorem contains_insertList [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l : List ((a : α) × (β a))} {k : α} : +theorem contains_insertMany_list [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l : List ((a : α) × (β a))} {k : α} : (m.insertMany l).1.contains k ↔ m.contains k ∨ (l.map Sigma.fst).contains k := by simp_to_model using List.containsKey_insertList -theorem contains_of_contains_insertList [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l : List ((a : α) × (β a))} {k : α} : +theorem contains_of_contains_insertMany_list [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l : List ((a : α) × (β a))} {k : α} : (m.insertMany l).1.contains k → (l.map Sigma.fst).contains k = false → m.contains k := by simp_to_model using List.containsKey_of_containsKey_insertList -theorem get?_insertList_not_mem [LawfulBEq α] +theorem get?_insertMany_list_not_mem [LawfulBEq α] {l : List ((a : α) × (β a))} {k : α} - {distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)} {not_mem : (l.map (Sigma.fst)).contains k = false} (h : m.1.WF) : (m.insertMany l).1.get? k = m.get? k := by simp_to_model using getValueCast?_insertList_not_mem -theorem get?_insertList_mem [LawfulBEq α] +theorem get?_insertMany_list_mem [LawfulBEq α] {l : List ((a : α) × (β a))} {k k' : α} {k_eq : k == k'} {v : β k} {distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)} {mem : ⟨k,v⟩ ∈ l} @@ -881,16 +880,15 @@ theorem get?_insertList_mem [LawfulBEq α] (m.insertMany l).1.get? k' = some (cast (by congr; apply LawfulBEq.eq_of_beq k_eq) v) := by simp_to_model using getValueCast?_insertList_mem -theorem get_insertList_not_mem [LawfulBEq α] +theorem get_insertMany_list_not_mem [LawfulBEq α] {l : List ((a : α) × (β a))} {k : α} - {distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)} {not_mem : (l.map (Sigma.fst)).contains k = false} {h'} (h : m.1.WF) : - (m.insertMany l).1.get k h' = m.get k (contains_of_contains_insertList _ h h' not_mem) := by + (m.insertMany l).1.get k h' = m.get k (contains_of_contains_insertMany_list _ h h' not_mem) := by simp_to_model using getValueCast_insertList_not_mem -theorem get_insertList_mem [LawfulBEq α] +theorem get_insertMany_list_mem [LawfulBEq α] {l : List ((a : α) × (β a))} {k k' : α} {k_eq : k == k'} {v : β k} {distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)} {mem : ⟨k,v⟩ ∈ l} @@ -899,15 +897,14 @@ theorem get_insertList_mem [LawfulBEq α] (m.insertMany l).1.get k' h' = cast (by congr; apply LawfulBEq.eq_of_beq k_eq) v := by simp_to_model using getValueCast_insertList_mem -theorem get!_insertList_not_mem [LawfulBEq α] +theorem get!_insertMany_list_not_mem [LawfulBEq α] {l : List ((a : α) × (β a))} {k : α} [Inhabited (β k)] - {distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)} {not_mem : (l.map (Sigma.fst)).contains k = false} (h : m.1.WF) : (m.insertMany l).1.get! k = m.get! k := by simp_to_model using getValueCast!_insertList_not_mem -theorem get!_insertList_mem [LawfulBEq α] +theorem get!_insertMany_list_mem [LawfulBEq α] {l : List ((a : α) × (β a))} {k k' : α} {k_eq : k == k'} {v : β k} [Inhabited (β k')] {distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)} {mem : ⟨k,v⟩ ∈ l} @@ -915,15 +912,14 @@ theorem get!_insertList_mem [LawfulBEq α] (m.insertMany l).1.get! k' = cast (by congr; apply LawfulBEq.eq_of_beq k_eq) v := by simp_to_model using getValueCast!_insertList_mem -theorem getD_insertList_not_mem [LawfulBEq α] +theorem getD_insertMany_list_not_mem [LawfulBEq α] {l : List ((a : α) × (β a))} {k : α} {fallback : β k} - {distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)} {not_mem : (l.map (Sigma.fst)).contains k = false} (h : m.1.WF) : (m.insertMany l).1.getD k fallback = m.getD k fallback := by simp_to_model using getValueCastD_insertList_not_mem -theorem getD_insertList_mem [LawfulBEq α] +theorem getD_insertMany_list_mem [LawfulBEq α] {l : List ((a : α) × (β a))} {k k' : α} {k_eq : k == k'} {v : β k} {fallback : β k'} {distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)} {mem : ⟨k,v⟩ ∈ l} @@ -931,14 +927,13 @@ theorem getD_insertList_mem [LawfulBEq α] (m.insertMany l).1.getD k' fallback = cast (by congr; apply LawfulBEq.eq_of_beq k_eq) v := by simp_to_model using getValueCastD_insertList_mem -theorem getKey?_insertList_not_mem [EquivBEq α] [LawfulHashable α] (h : m.1.WF) +theorem getKey?_insertMany_list_not_mem [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l : List ((a : α) × (β a))} {k : α} - {distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)} {not_mem : (l.map Sigma.fst).contains k = false} : (m.insertMany l).1.getKey? k = m.getKey? k := by simp_to_model using List.getKey?_insertList_not_mem -theorem getKey?_insertList_mem [EquivBEq α] [LawfulHashable α] (h : m.1.WF) +theorem getKey?_insertMany_list_mem [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l : List ((a : α) × (β a))} {k k' : α} {k_beq : k == k'} {distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)} @@ -946,15 +941,14 @@ theorem getKey?_insertList_mem [EquivBEq α] [LawfulHashable α] (h : m.1.WF) (m.insertMany l).1.getKey? k' = some k := by simp_to_model using List.getKey?_insertList_mem -theorem getKey_insertList_not_mem [EquivBEq α] [LawfulHashable α] (h : m.1.WF) +theorem getKey_insertMany_list_not_mem [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l : List ((a : α) × (β a))} {k : α} - {distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)} {not_mem : (l.map Sigma.fst).contains k = false} {h'} : - (m.insertMany l).1.getKey k h' = m.getKey k (contains_of_contains_insertList _ h h' not_mem) := by + (m.insertMany l).1.getKey k h' = m.getKey k (contains_of_contains_insertMany_list _ h h' not_mem) := by simp_to_model using List.getKey_insertList_not_mem -theorem getKey_insertList_mem [EquivBEq α] [LawfulHashable α] (h : m.1.WF) +theorem getKey_insertMany_list_mem [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l : List ((a : α) × (β a))} {k k' : α} {k_beq : k == k'} {distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)} @@ -963,14 +957,13 @@ theorem getKey_insertList_mem [EquivBEq α] [LawfulHashable α] (h : m.1.WF) (m.insertMany l).1.getKey k' h' = k := by simp_to_model using List.getKey_insertList_mem -theorem getKey!_insertList_not_mem [EquivBEq α] [LawfulHashable α] [Inhabited α] (h : m.1.WF) +theorem getKey!_insertMany_list_not_mem [EquivBEq α] [LawfulHashable α] [Inhabited α] (h : m.1.WF) {l : List ((a : α) × (β a))} {k : α} - {distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)} {not_mem : (l.map Sigma.fst).contains k = false} : (m.insertMany l).1.getKey! k = m.getKey! k := by simp_to_model using List.getKey!_insertList_not_mem -theorem getKey!_insertList_mem [EquivBEq α] [LawfulHashable α] [Inhabited α] (h : m.1.WF) +theorem getKey!_insertMany_list_mem [EquivBEq α] [LawfulHashable α] [Inhabited α] (h : m.1.WF) {l : List ((a : α) × (β a))} {k k' : α} {k_beq : k == k'} {distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)} @@ -978,14 +971,13 @@ theorem getKey!_insertList_mem [EquivBEq α] [LawfulHashable α] [Inhabited α] (m.insertMany l).1.getKey! k' = k := by simp_to_model using List.getKey!_insertList_mem -theorem getKeyD_insertList_not_mem [EquivBEq α] [LawfulHashable α] (h : m.1.WF) +theorem getKeyD_insertMany_list_not_mem [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l : List ((a : α) × (β a))} {k fallback : α} - {distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)} {not_mem : (l.map Sigma.fst).contains k = false} : (m.insertMany l).1.getKeyD k fallback = m.getKeyD k fallback := by simp_to_model using List.getKeyD_insertList_not_mem -theorem getKeyD_insertList_mem [EquivBEq α] [LawfulHashable α] (h : m.1.WF) +theorem getKeyD_insertMany_list_mem [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l : List ((a : α) × (β a))} {k k' fallback : α} {k_beq : k == k'} {distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)} @@ -993,7 +985,7 @@ theorem getKeyD_insertList_mem [EquivBEq α] [LawfulHashable α] (h : m.1.WF) (m.insertMany l).1.getKeyD k' fallback = k := by simp_to_model using List.getKeyD_insertList_mem -theorem size_insertList [EquivBEq α] [LawfulHashable α] +theorem size_insertMany_list [EquivBEq α] [LawfulHashable α] {l : List ((a : α) × (β a))} {distinct : List.Pairwise (fun a b => (a.1 == b.1) = false) l} (h : m.1.WF) : @@ -1008,16 +1000,72 @@ theorem size_insertList [EquivBEq α] [LawfulHashable α] . exact distinct . apply distinct' -theorem insertList_notEmpty_if_m_notEmpty [EquivBEq α] [LawfulHashable α] +theorem insertMany_list_notEmpty_if_map_notEmpty [EquivBEq α] [LawfulHashable α] {l : List ((a : α) × (β a))} (h : m.1.WF) : (m.1.isEmpty = false) → (m.insertMany l).1.1.isEmpty = false := by simp_to_model using List.insertList_not_isEmpty_if_start_not_isEmpty -theorem insertList_isEmpty [EquivBEq α] [LawfulHashable α] +theorem insertMany_list_isEmpty [EquivBEq α] [LawfulHashable α] {l : List ((a : α) × (β a))} (h : m.1.WF) : (m.insertMany l).1.1.isEmpty ↔ m.1.isEmpty ∧ l.isEmpty := by simp_to_model using List.insertList_isEmpty +section +variable {β: Type v} (m: Raw₀ α (fun _ => β)) + +theorem Const.get?_insertMany_list_not_mem [EquivBEq α] [LawfulHashable α] + {l : List ((_ : α) × β)} {k : α} + {not_mem : (l.map Sigma.fst).contains k = false} + (h : m.1.WF) : + get? (m.insertMany l).1 k = get? m k := by + simp_to_model using getValue?_insertList_not_mem + +theorem Const.get?_insertMany_list_mem [EquivBEq α] [LawfulHashable α] + {l : List ((_ : α) × β)} {k k' : α} {k_eq: k == k'} {v : β} {mem : ⟨k,v⟩ ∈ l} + {distinct : List.Pairwise (fun a b => (a.1 == b.1) = false ) l} (h : m.1.WF): + get? (m.insertMany l).1 k' = v := by + simp_to_model using getValue?_insertList_mem + +theorem Const.get_insertMany_list_not_mem [EquivBEq α] [LawfulHashable α] + {l : List ((_ : α) × β)} {k : α} + {not_mem : (l.map Sigma.fst).contains k = false} + (h : m.1.WF) {h'}: + get (m.insertMany l).1 k h' = get m k (contains_of_contains_insertMany_list _ h h' not_mem) := by + simp_to_model using getValue_insertList_not_mem + +theorem Const.get_insertMany_list_mem [EquivBEq α] [LawfulHashable α] + {l : List ((_ : α) × β)} {k k' : α} {k_eq: k == k'} {v : β} {mem : ⟨k,v⟩ ∈ l} + {distinct : List.Pairwise (fun a b => (a.1 == b.1) = false ) l} (h : m.1.WF) {h'}: + get (m.insertMany l).1 k' h' = v := by + simp_to_model using getValue_insertList_mem + +theorem Const.get!_insertMany_list_not_mem [EquivBEq α] [LawfulHashable α] [Inhabited β] + {l : List ((_ : α) × β)} {k : α} + {not_mem : (l.map Sigma.fst).contains k = false} + (h : m.1.WF) : + get! (m.insertMany l).1 k = get! m k := by + simp_to_model using getValue!_insertList_not_mem + +theorem Const.get!_insertMany_list_mem [EquivBEq α] [LawfulHashable α] [Inhabited β] + {l : List ((_ : α) × β)} {k k' : α} {k_eq: k == k'} {v : β} {mem : ⟨k,v⟩ ∈ l} + {distinct : List.Pairwise (fun a b => (a.1 == b.1) = false ) l} (h : m.1.WF): + get! (m.insertMany l).1 k' = v := by + simp_to_model using getValue!_insertList_mem + +theorem Const.getD_insertMany_list_not_mem [EquivBEq α] [LawfulHashable α] + {l : List ((_ : α) × β)} {k : α} {fallback : β} + {not_mem : (l.map Sigma.fst).contains k = false} + (h : m.1.WF) : + getD (m.insertMany l).1 k fallback = getD m k fallback:= by + simp_to_model using getValueD_insertList_not_mem + +theorem Const.getD_insertMany_list_mem [EquivBEq α] [LawfulHashable α] + {l : List ((_ : α) × β)} {k k' : α} {k_eq: k == k'} {v fallback: β} {mem : ⟨k,v⟩ ∈ l} + {distinct : List.Pairwise (fun a b => (a.1 == b.1) = false ) l} (h : m.1.WF): + getD (m.insertMany l).1 k' fallback = v := by + simp_to_model using getValueD_insertList_mem + +end end Raw₀ end Std.DHashMap.Internal diff --git a/src/Std/Data/DHashMap/Internal/WF.lean b/src/Std/Data/DHashMap/Internal/WF.lean index 71a03a0ca860..05f523051247 100644 --- a/src/Std/Data/DHashMap/Internal/WF.lean +++ b/src/Std/Data/DHashMap/Internal/WF.lean @@ -757,13 +757,32 @@ theorem wfImp_insertMany [BEq α] [Hashable α] [EquivBEq α] [LawfulHashable α Raw.WFImp (m.insertMany l).1.1 := Raw.WF.out ((m.insertMany l).2 _ Raw.WF.insert₀ (.wf m.2 h)) -theorem toListModel_insertMany [BEq α] [Hashable α] [EquivBEq α] [LawfulHashable α] +theorem toListModel_insertMany_list [BEq α] [Hashable α] [EquivBEq α] [LawfulHashable α] {m : Raw₀ α β} {l : List ((a : α) × (β a))} (h : Raw.WFImp m.1) : Perm (toListModel (insertMany m l).1.1.buckets) (List.insertList (toListModel m.1.buckets) l) := by rw[insertMany_eq_insertListₘ] apply toListModel_insertListₘ exact h +/-! # `Const.insertListₘ` -/ +theorem Const.toListModel_insertListₘ {β : Type v} [BEq α] [Hashable α] [EquivBEq α] [LawfulHashable α] {m : Raw₀ α (fun _ => β)}{l : List (α × β)} (h : Raw.WFImp m.1): + Perm (toListModel (Const.insertListₘ m l).1.buckets) (insertListConst (toListModel m.1.buckets) l) := by + induction l generalizing m with + | nil => simp [Const.insertListₘ, insertListConst] + | cons hd tl ih => + simp [Const.insertListₘ, insertListConst] + apply Perm.trans + apply ih (wfImp_insert h) + apply List.insertListConst_perm_of_perm_first + apply toListModel_insert h + apply (wfImp_insert h).distinct + +/-! # `Const.insertMany` -/ +theorem Const.toListModel_insertMany_list {β : Type v} [BEq α] [Hashable α] [EquivBEq α] [LawfulHashable α] {m : Raw₀ α (fun _ => β)} {l : List (α × β)} (h : Raw.WFImp m.1): + Perm (toListModel (Const.insertMany m l).1.1.buckets) (insertListConst (toListModel m.1.buckets) l) := by + rw [Const.insertMany_eq_insertListₘ] + apply toListModel_insertListₘ h + theorem Const.wfImp_insertMany {β : Type v} [BEq α] [Hashable α] [EquivBEq α] [LawfulHashable α] {ρ : Type w} [ForIn Id ρ (α × β)] {m : Raw₀ α (fun _ => β)} {l : ρ} (h : Raw.WFImp m.1) : Raw.WFImp (Const.insertMany m l).1.1 := From 446e80186e2cb3319d3f6cd9af4bdf1cc3018b68 Mon Sep 17 00:00:00 2001 From: jt0202 Date: Sat, 30 Nov 2024 22:37:35 +0100 Subject: [PATCH 30/83] Verify Const.insertMany --- .../DHashMap/Internal/List/Associative.lean | 342 +++++++++++++++++- src/Std/Data/DHashMap/Internal/RawLemmas.lean | 201 ++++++++-- 2 files changed, 500 insertions(+), 43 deletions(-) diff --git a/src/Std/Data/DHashMap/Internal/List/Associative.lean b/src/Std/Data/DHashMap/Internal/List/Associative.lean index c539f9c603f7..40ddb75c770d 100644 --- a/src/Std/Data/DHashMap/Internal/List/Associative.lean +++ b/src/Std/Data/DHashMap/Internal/List/Associative.lean @@ -38,6 +38,14 @@ theorem assoc_induction {motive : List ((a : α) × β a) → Prop} (nil : motiv | [] => nil | ⟨_, _⟩ :: _ => cons _ _ _ (assoc_induction nil cons _) +theorem keys_def {l: List ((a : α ) × β a)}: + keys l = l.map Sigma.fst := by + induction l with + | nil => simp [keys] + | cons hd tl ih => + simp only [keys, List.map_cons, List.cons.injEq, true_and] + apply ih + /-- Internal implementation detail of the hash map -/ def getEntry? [BEq α] (a : α) : List ((a : α) × β a) → Option ((a : α) × β a) | [] => none @@ -1904,7 +1912,7 @@ theorem DistinctKeys.insertList [BEq α] [PartialEquivBEq α] {l₁ l₂ : List theorem insertList_perm_of_perm_first [BEq α] [EquivBEq α] {l1 l2 toInsert : List ((a : α) × β a)} (h : Perm l1 l2) (distinct : DistinctKeys l1) : Perm (insertList l1 toInsert) (insertList l2 toInsert) := by induction toInsert generalizing l1 l2 with - | nil => simp[insertList, h] + | nil => simp [insertList, h] | cons hd tl ih => simp only [insertList] apply ih @@ -1943,17 +1951,17 @@ theorem containsKey_eq_contains_map_fst [BEq α] [PartialEquivBEq α] (l : List rw [BEq.comm] theorem containsKey_insertList [BEq α] [PartialEquivBEq α] (l toInsert : List ((a : α) × β a)) (k : α) : - containsKey k (List.insertList l toInsert) ↔ containsKey k l ∨ (toInsert.map Sigma.fst).contains k := by + containsKey k (List.insertList l toInsert) = (containsKey k l || (toInsert.map Sigma.fst).contains k) := by induction toInsert generalizing l with - | nil => simp only [insertList, List.map_nil, List.elem_nil, Bool.false_eq_true, or_false] + | nil => simp only [insertList, List.map_nil, List.elem_nil, Bool.or_false] | cons hd tl ih => unfold insertList rw [ih] rw [containsKey_insertEntry] simp only [Bool.or_eq_true, List.map_cons, List.contains_cons] rw [BEq.comm] - conv => left; left; rw [or_comm] - rw [or_assoc] + conv => left; left; rw [Bool.or_comm] + rw [Bool.or_assoc] theorem containsKey_of_containsKey_insertList [BEq α] [PartialEquivBEq α] {l toInsert : List ((a : α) × β a)} {k : α} (h₁ : containsKey k (insertList l toInsert)) @@ -2235,6 +2243,53 @@ theorem insertList_perm [BEq α] [ReflBEq α] [PartialEquivBEq α] apply Perm.symm apply List.perm_middle +theorem insertList_length [BEq α] [EquivBEq α] + {l toInsert : List ((a : α) × β a)} + {distinct_l : DistinctKeys l} + (distinct_toInsert : List.Pairwise (fun a b => (a.1 == b.1) = false) toInsert) + (distinct_both : ∀ (a : α), ¬ (containsKey a l ∧ (toInsert.map Sigma.fst).contains a)) : + (insertList l toInsert).length = l.length + toInsert.length := by + rw [← DistinctKeys.def] at distinct_toInsert + induction toInsert generalizing l with + | nil => simp [insertList] + | cons hd tl ih => + simp only [insertList, List.length_cons] + rw [ih] + · rw [length_insertEntry] + split + · rename_i h + specialize distinct_both hd.1 + simp only [not_and, Bool.not_eq_true] at distinct_both + specialize distinct_both h + simp [containsKey] at distinct_both + · rw [Nat.add_assoc, Nat.add_comm 1 _] + · apply DistinctKeys.tail distinct_toInsert + · intro a + specialize distinct_both a + simp only [List.map_cons, List.contains_cons, Bool.or_eq_true, not_and, not_or, + Bool.not_eq_true] at distinct_both + simp only [containsKey_insertEntry, Bool.or_eq_true, not_and, Bool.not_eq_true] + intro h + cases h with + | inl h => + rw [DistinctKeys.def] at distinct_toInsert + simp at distinct_toInsert + rcases distinct_toInsert with ⟨left,_⟩ + simp only [List.contains_eq_any_beq, List.any_map, List.any_eq_false, Function.comp_apply, + Bool.not_eq_true] + intro x x_mem + specialize left x x_mem + false_or_by_contra + rename_i h' + simp only [Bool.not_eq_false] at h' + have hd_x: hd.fst == x.fst := BEq.trans h h' + rw [hd_x] at left + simp only [Bool.true_eq_false] at left + | inr h => + apply And.right (distinct_both h) + · apply DistinctKeys.insertEntry distinct_l + + theorem insertList_not_isEmpty_if_start_not_isEmpty [BEq α] {l toInsert : List ((a : α) × β a)} {h : l.isEmpty = false} : @@ -2376,7 +2431,7 @@ def insertListConst [BEq α] (l: List ((_ : α) × β))(toInsert: List (α × β | .nil => l | .cons hd tl => insertListConst (insertEntry hd.1 hd.2 l) tl -theorem insertListConst_eq_insertList [BEq α] {l: List ((_ : α) × β)} {toInsert: List (α × β)}: +theorem insertListConst_eq_insertList [BEq α] {l : List ((_ : α) × β)} {toInsert : List (α × β)}: insertListConst l toInsert = insertList l (List.map (fun a => ⟨a.1, a.2⟩) toInsert) := by induction toInsert generalizing l with | nil => simp [insertList, insertListConst] @@ -2394,15 +2449,270 @@ theorem insertListConst_perm_of_perm_first [BEq α] [EquivBEq α] (l1 l2: List ( simp only [insertListConst_eq_insertList] apply insertList_perm_of_perm_first h distinct -theorem containsKey_insertListConst [BEq α] [PartialEquivBEq α] {l : List ((_ : α) × β)} {toInsert: List (α × β)} {k : α} : - containsKey k (insertListConst l toInsert) ↔ containsKey k l ∨ (toInsert.map Prod.fst).contains k := by - rw [insertListConst_eq_insertList, containsKey_insertList] - have h: (List.map Sigma.fst (List.map (fun a => (⟨a.fst, a.snd⟩ : (_ : α ) × β )) toInsert)) = (List.map Prod.fst toInsert) := by - apply List.ext_get - · simp - · intro n h1 h2 - simp - rw [h] +theorem list_conversion_insertListConst {toInsert : List (α × β)}: (List.map Sigma.fst (List.map (fun a => (⟨a.fst, a.snd⟩ : (_ : α ) × β )) toInsert)) = (List.map Prod.fst toInsert) := by + apply List.ext_get + · simp + · intro n h1 h2 + simp -end +theorem list_conversion_insertListConst_contains [BEq α] [PartialEquivBEq α] {toInsert : List (α × β)} {a : α}: + containsKey a (List.map (fun a => (⟨a.fst, a.snd⟩ : ((_ : α ) × β) )) toInsert) = true ↔ (toInsert.map Prod.fst).contains a = true := by + rw [List.contains_iff_exists_mem_beq] + rw [containsKey_iff_exists] + simp only [keys_def, list_conversion_insertListConst] + +theorem pairwise_list_conversion_insertListConst [BEq α] {toInsert : List (α × β)} (distinct : toInsert.Pairwise (fun a b => (a.1 == b.1) = false) ): + List.Pairwise (fun a b => (a.1 == b.1) = false) (List.map (fun a => (⟨a.1, a.2⟩ : ((_ : α ) × β)) ) toInsert) := by + induction toInsert with + | nil => simp + | cons hd tl ih => + simp only [List.map_cons, pairwise_cons] + simp only [List.map_cons, pairwise_cons] at distinct + rcases distinct with ⟨left,right⟩ + constructor + · intro a' a_mem + simp only [List.mem_map, Prod.exists] at a_mem + rcases a_mem with ⟨a,b,ab_tl, ab_a⟩ + rw [← ab_a] + simp only + specialize left (a,b) ab_tl + apply left + · apply ih right + +theorem containsKey_insertListConst [BEq α] [PartialEquivBEq α] {l : List ((_ : α) × β)} {toInsert : List (α × β)} {k : α} : + containsKey k (insertListConst l toInsert) = (containsKey k l || (toInsert.map Prod.fst).contains k) := by + rw [insertListConst_eq_insertList, containsKey_insertList,list_conversion_insertListConst] + +theorem containsKey_of_containsKey_insertListConst [BEq α] [PartialEquivBEq α] {l : List ((_ : α) × β)} {toInsert : List (α × β)} {k : α} (h₁ : containsKey k (insertListConst l toInsert)) + (h₂ : (toInsert.map Prod.fst).contains k = false) : containsKey k l := by + rw [insertListConst_eq_insertList] at h₁ + apply containsKey_of_containsKey_insertList h₁ + rw [list_conversion_insertListConst] + exact h₂ + +theorem getKey?_insertListConst_not_mem [BEq α] [EquivBEq α] + {l : List ((_ : α) × β)} {toInsert : List (α × β)} {k : α} + {not_mem : (toInsert.map Prod.fst).contains k = false} : + getKey? k (insertListConst l toInsert) = getKey? k l := by + rw [insertListConst_eq_insertList] + apply getKey?_insertList_not_mem + rw [list_conversion_insertListConst] + exact not_mem + +theorem getKey?_insertListConst_mem [BEq α] [EquivBEq α] + {l : List ((_ : α) × β)} {toInsert : List (α × β)} + {k k' : α} {k_beq : k == k'} + {distinct : toInsert.Pairwise (fun a b => (a.1 == b.1) = false)} + {distinct2 : DistinctKeys l} + {mem : k ∈ (toInsert.map Prod.fst)} : + getKey? k' (insertListConst l toInsert) = some k := by + rw [insertListConst_eq_insertList] + apply getKey?_insertList_mem + · exact k_beq + · apply pairwise_list_conversion_insertListConst distinct + · exact distinct2 + · rw [list_conversion_insertListConst] + assumption + +theorem getKey_insertListConst_not_mem [BEq α] [EquivBEq α] + {l : List ((_ : α) × β)} {toInsert : List (α × β)} {k : α} + {not_mem : (toInsert.map Prod.fst).contains k = false} + {h} : + getKey k (insertListConst l toInsert) h = + getKey k l (containsKey_of_containsKey_insertListConst h not_mem) := by + rw [← Option.some_inj] + rw [← getKey?_eq_some_getKey] + rw [getKey?_insertListConst_not_mem] + . rw [getKey?_eq_some_getKey] + . exact not_mem + +theorem getKey_insertListConst_mem [BEq α] [EquivBEq α] + {l : List ((_ : α) × β)} {toInsert : List (α × β)} + {k k' : α} {k_beq : k == k'} + {distinct : toInsert.Pairwise (fun a b => (a.1 == b.1) = false)} + {distinct2 : DistinctKeys l} + {mem : k ∈ (toInsert.map Prod.fst)} + {h} : + getKey k' (insertListConst l toInsert) h = k := by + rw [← Option.some_inj] + rw [← getKey?_eq_some_getKey] + rw [getKey?_insertListConst_mem] + . exact k_beq + . exact distinct + . exact distinct2 + . exact mem + +theorem getKey!_insertListConst_not_mem [BEq α] [EquivBEq α] [Inhabited α] + {l : List ((_ : α) × β)} {toInsert : List (α × β)} {k : α} + {not_mem : (toInsert.map Prod.fst).contains k = false} : + getKey! k (insertListConst l toInsert) = getKey! k l := by + rw [insertListConst_eq_insertList] + apply getKey!_insertList_not_mem + rw [list_conversion_insertListConst] + exact not_mem + +theorem getKey!_insertListConst_mem [BEq α] [EquivBEq α] [Inhabited α] + {l : List ((_ : α) × β)} {toInsert : List (α × β)} + {k k' : α} {k_beq : k == k'} + {distinct : toInsert.Pairwise (fun a b => (a.1 == b.1) = false)} + {distinct2 : DistinctKeys l} + {mem : k ∈ (toInsert.map Prod.fst)} : + getKey! k' (insertListConst l toInsert) = k := by + rw [insertListConst_eq_insertList] + apply getKey!_insertList_mem + · exact k_beq + · apply pairwise_list_conversion_insertListConst + exact distinct + · exact distinct2 + · rw [list_conversion_insertListConst] + exact mem + +theorem getKeyD_insertListConst_not_mem [BEq α] [EquivBEq α] + {l : List ((_ : α) × β)} {toInsert : List (α × β)} {k fallback : α} + {not_mem : (toInsert.map Prod.fst).contains k = false} : + getKeyD k (insertListConst l toInsert) fallback = getKeyD k l fallback := by + rw [insertListConst_eq_insertList] + apply getKeyD_insertList_not_mem + rw [list_conversion_insertListConst] + exact not_mem + +theorem getKeyD_insertListConst_mem [BEq α] [EquivBEq α] + {l : List ((_ : α) × β)} {toInsert : List (α × β)} + {k k' fallback : α} {k_beq : k == k'} + {distinct : toInsert.Pairwise (fun a b => (a.1 == b.1) = false)} + {distinct2 : DistinctKeys l} + {mem : k ∈ (toInsert.map Prod.fst)} : + getKeyD k' (insertListConst l toInsert) fallback = k := by + rw [insertListConst_eq_insertList] + apply getKeyD_insertList_mem + · exact k_beq + · apply pairwise_list_conversion_insertListConst + exact distinct + · exact distinct2 + · rw [list_conversion_insertListConst] + exact mem + +theorem insertListConst_length [BEq α] [EquivBEq α] + {l : List ((_ : α) × β)} {toInsert : List (α × β)} + {distinct_l : DistinctKeys l} + (distinct_toInsert : List.Pairwise (fun a b => (a.1 == b.1) = false) toInsert) + (distinct_both : ∀ (a : α), ¬ (containsKey a l ∧ (toInsert.map Prod.fst).contains a)) : + (insertListConst l toInsert).length = l.length + toInsert.length := by + rw [insertListConst_eq_insertList] + rw [insertList_length] + · simp + · apply pairwise_list_conversion_insertListConst distinct_toInsert + · intro a + simp only [not_and, Bool.not_eq_true] + intro h + specialize distinct_both a + simp only [not_and, Bool.not_eq_true] at distinct_both + specialize distinct_both h + rw [List.contains_eq_any_beq] + simp only [List.map_map, List.any_map, List.any_eq_false, Function.comp_apply, Bool.not_eq_true, + Prod.forall] + rw [List.contains_eq_any_beq] at distinct_both + simp only [List.any_map, List.any_eq_false, Function.comp_apply, Bool.not_eq_true, + Prod.forall] at distinct_both + apply distinct_both + · exact distinct_l + +theorem insertListConst_not_isEmpty_if_start_not_isEmpty [BEq α] + {l : List ((_ : α) × β)} {toInsert : List (α × β)} + {h : l.isEmpty = false} : + (List.insertListConst l toInsert).isEmpty = false := by + induction toInsert generalizing l with + | nil => simp [insertListConst, h] + | cons hd tl ih => + simp [insertListConst, ih] + +theorem insertListConst_isEmpty [BEq α] + {l : List ((_ : α) × β)} {toInsert : List (α × β)} : + (List.insertListConst l toInsert).isEmpty ↔ l.isEmpty ∧ toInsert.isEmpty := by + rw [insertListConst_eq_insertList] + rw [insertList_isEmpty] + simp only [List.isEmpty_eq_true, List.map_eq_nil_iff] + +theorem getValue?_insertListConst_not_mem [BEq α] [PartialEquivBEq α] + {l : List ((_ : α) × β)} {toInsert : List (α × β)} {k : α} + {not_mem : (toInsert.map Prod.fst).contains k = false} : + getValue? k (insertListConst l toInsert) = getValue? k l := by + rw [insertListConst_eq_insertList] + rw [getValue?_insertList_not_mem] + rw [list_conversion_insertListConst] + exact not_mem + +theorem getValue?_insertListConst_mem [BEq α] [PartialEquivBEq α] + {l : List ((_ : α) × β)} {toInsert : List (α × β)} + {k k' : α} {k_beq : k == k'} {v : β} + {distinct : List.Pairwise (fun a b => (a.1 == b.1) = false) toInsert} + {mem : ⟨k,v⟩ ∈ toInsert} : + getValue? k' (insertListConst l toInsert) = v := by + rw [insertListConst_eq_insertList] + apply getValue?_insertList_mem (k_beq:=k_beq) + · apply pairwise_list_conversion_insertListConst distinct + · simp only [List.mem_map, Prod.exists] + exists k + exists v + +theorem getValue!_insertListConst_not_mem [BEq α] [PartialEquivBEq α] [Inhabited β] + {l : List ((_ : α) × β)} {toInsert : List (α × β)} {k : α} + {not_mem : (toInsert.map Prod.fst).contains k = false} : + getValue! k (insertListConst l toInsert) = getValue! k l := by + rw [insertListConst_eq_insertList] + rw [getValue!_insertList_not_mem] + rw [list_conversion_insertListConst] + exact not_mem + +theorem getValue!_insertListConst_mem [BEq α] [PartialEquivBEq α] [Inhabited β] + {l : List ((_ : α) × β)} {toInsert : List (α × β)} {k k' : α} {v: β} {k_eq : k == k'} {distinct : List.Pairwise (fun a b => (a.1 == b.1) = false) toInsert } {mem : ⟨k,v⟩ ∈ toInsert}: + getValue! k' (insertListConst l toInsert) = v := by + rw [insertListConst_eq_insertList, getValue!_insertList_mem (k_eq:=k_eq)] + · apply pairwise_list_conversion_insertListConst distinct + · simp only [List.mem_map, Prod.exists] + exists k + exists v + +theorem getValueD_insertListConst_not_mem [BEq α] [PartialEquivBEq α] + {l : List ((_ : α) × β)} {toInsert : List (α × β)} {k : α} {fallback : β} + {not_mem : (toInsert.map Prod.fst).contains k = false} : + getValueD k (insertListConst l toInsert) fallback = getValueD k l fallback := by + rw [insertListConst_eq_insertList, getValueD_insertList_not_mem] + rw [list_conversion_insertListConst] + exact not_mem + +theorem getValueD_insertListConst_mem [BEq α] [PartialEquivBEq α] + {l : List ((_ : α) × β)} {toInsert : List (α × β)} {k k' : α} {v fallback: β} {k_eq : k == k'} {distinct : List.Pairwise (fun a b => (a.1 == b.1) = false) toInsert } {mem : ⟨k,v⟩ ∈ toInsert}: + getValueD k' (insertListConst l toInsert) fallback= v := by + rw [insertListConst_eq_insertList, getValueD_insertList_mem (k_eq:= k_eq)] + · apply pairwise_list_conversion_insertListConst distinct + · simp + exists k + exists v + +theorem getValue_insertListConst_not_mem [BEq α] [PartialEquivBEq α] + {l : List ((_ : α) × β)} {toInsert : List (α × β)} {k : α} + {not_mem : (toInsert.map Prod.fst).contains k = false} + {h} : + getValue k (insertListConst l toInsert) h = getValue k l (containsKey_of_containsKey_insertListConst h not_mem) := by + rw [← Option.some_inj] + rw [← getValue?_eq_some_getValue] + rw [← getValue?_eq_some_getValue] + rw [getValue?_insertListConst_not_mem] + exact not_mem + +theorem getValue_insertListConst_mem [BEq α] [PartialEquivBEq α] + {l : List ((_ : α) × β)} {toInsert : List (α × β)} + {k k' : α} {k_beq : k == k'} {v : β} + {distinct : toInsert.Pairwise (fun a b => (a.1 == b.1) = false)} + {mem : ⟨k,v⟩ ∈ toInsert} + {h} : + getValue k' (insertListConst l toInsert) h = v := by + rw [← Option.some_inj] + rw [← getValue?_eq_some_getValue] + rw [getValue?_insertListConst_mem (mem:=mem)] + . exact k_beq + . exact distinct + + end end List diff --git a/src/Std/Data/DHashMap/Internal/RawLemmas.lean b/src/Std/Data/DHashMap/Internal/RawLemmas.lean index 47f81862dba0..efd2225391ca 100644 --- a/src/Std/Data/DHashMap/Internal/RawLemmas.lean +++ b/src/Std/Data/DHashMap/Internal/RawLemmas.lean @@ -70,7 +70,7 @@ variable [BEq α] [Hashable α] /-- Internal implementation detail of the hash map -/ scoped macro "wf_trivial" : tactic => `(tactic| repeat (first - | apply Raw₀.wfImp_insert | apply Raw₀.wfImp_insertIfNew | apply Raw₀.wfImp_insertMany | apply Raw₀.wfImp_erase + | apply Raw₀.wfImp_insert | apply Raw₀.wfImp_insertIfNew | apply Raw₀.wfImp_insertMany | apply Raw₀.Const.wfImp_insertMany| apply Raw₀.wfImp_erase | apply Raw.WF.out | assumption | apply Raw₀.wfImp_empty | apply Raw.WFImp.distinct | apply Raw.WF.empty₀)) @@ -858,7 +858,7 @@ theorem insertMany_list_cons {l : List ((a : α) × (β a))} {k : α} {v : β k} | cons hd tl => simp [insertListₘ] theorem contains_insertMany_list [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l : List ((a : α) × (β a))} {k : α} : - (m.insertMany l).1.contains k ↔ m.contains k ∨ (l.map Sigma.fst).contains k := by + (m.insertMany l).1.contains k = (m.contains k || (l.map Sigma.fst).contains k) := by simp_to_model using List.containsKey_insertList theorem contains_of_contains_insertMany_list [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l : List ((a : α) × (β a))} {k : α} : @@ -989,16 +989,9 @@ theorem size_insertMany_list [EquivBEq α] [LawfulHashable α] {l : List ((a : α) × (β a))} {distinct : List.Pairwise (fun a b => (a.1 == b.1) = false) l} (h : m.1.WF) : - (∀ (a : α), ¬ (m.contains a = true ∧ List.containsKey a l = true)) → + (∀ (a : α), ¬ (m.contains a = true ∧ (l.map Sigma.fst).contains a = true)) → (m.insertMany l).1.1.size = m.1.size + l.length := by - simp_to_model - rw [← List.length_append] - intro distinct' - apply List.Perm.length_eq - apply List.insertList_perm - . apply (Raw.WF.out h).distinct - . exact distinct - . apply distinct' + simp_to_model using List.insertList_length theorem insertMany_list_notEmpty_if_map_notEmpty [EquivBEq α] [LawfulHashable α] {l : List ((a : α) × (β a))} (h : m.1.WF) : @@ -1013,58 +1006,212 @@ theorem insertMany_list_isEmpty [EquivBEq α] [LawfulHashable α] section variable {β: Type v} (m: Raw₀ α (fun _ => β)) -theorem Const.get?_insertMany_list_not_mem [EquivBEq α] [LawfulHashable α] +theorem get?_insertMany_list_not_mem_const [EquivBEq α] [LawfulHashable α] {l : List ((_ : α) × β)} {k : α} {not_mem : (l.map Sigma.fst).contains k = false} (h : m.1.WF) : - get? (m.insertMany l).1 k = get? m k := by + Const.get? (m.insertMany l).1 k = Const.get? m k := by simp_to_model using getValue?_insertList_not_mem -theorem Const.get?_insertMany_list_mem [EquivBEq α] [LawfulHashable α] +theorem get?_insertMany_list_mem_const [EquivBEq α] [LawfulHashable α] {l : List ((_ : α) × β)} {k k' : α} {k_eq: k == k'} {v : β} {mem : ⟨k,v⟩ ∈ l} {distinct : List.Pairwise (fun a b => (a.1 == b.1) = false ) l} (h : m.1.WF): - get? (m.insertMany l).1 k' = v := by + Const.get? (m.insertMany l).1 k' = v := by simp_to_model using getValue?_insertList_mem -theorem Const.get_insertMany_list_not_mem [EquivBEq α] [LawfulHashable α] +theorem get_insertMany_list_not_mem_const [EquivBEq α] [LawfulHashable α] {l : List ((_ : α) × β)} {k : α} {not_mem : (l.map Sigma.fst).contains k = false} (h : m.1.WF) {h'}: - get (m.insertMany l).1 k h' = get m k (contains_of_contains_insertMany_list _ h h' not_mem) := by + Const.get (m.insertMany l).1 k h' = Const.get m k (contains_of_contains_insertMany_list _ h h' not_mem) := by simp_to_model using getValue_insertList_not_mem -theorem Const.get_insertMany_list_mem [EquivBEq α] [LawfulHashable α] +theorem get_insertMany_list_mem_const [EquivBEq α] [LawfulHashable α] {l : List ((_ : α) × β)} {k k' : α} {k_eq: k == k'} {v : β} {mem : ⟨k,v⟩ ∈ l} {distinct : List.Pairwise (fun a b => (a.1 == b.1) = false ) l} (h : m.1.WF) {h'}: - get (m.insertMany l).1 k' h' = v := by + Const.get (m.insertMany l).1 k' h' = v := by simp_to_model using getValue_insertList_mem -theorem Const.get!_insertMany_list_not_mem [EquivBEq α] [LawfulHashable α] [Inhabited β] +theorem get!_insertMany_list_not_mem_const [EquivBEq α] [LawfulHashable α] [Inhabited β] {l : List ((_ : α) × β)} {k : α} {not_mem : (l.map Sigma.fst).contains k = false} (h : m.1.WF) : - get! (m.insertMany l).1 k = get! m k := by + Const.get! (m.insertMany l).1 k = Const.get! m k := by simp_to_model using getValue!_insertList_not_mem -theorem Const.get!_insertMany_list_mem [EquivBEq α] [LawfulHashable α] [Inhabited β] +theorem get!_insertMany_list_mem_const [EquivBEq α] [LawfulHashable α] [Inhabited β] {l : List ((_ : α) × β)} {k k' : α} {k_eq: k == k'} {v : β} {mem : ⟨k,v⟩ ∈ l} {distinct : List.Pairwise (fun a b => (a.1 == b.1) = false ) l} (h : m.1.WF): - get! (m.insertMany l).1 k' = v := by + Const.get! (m.insertMany l).1 k' = v := by simp_to_model using getValue!_insertList_mem -theorem Const.getD_insertMany_list_not_mem [EquivBEq α] [LawfulHashable α] +theorem getD_insertMany_list_not_mem_const [EquivBEq α] [LawfulHashable α] {l : List ((_ : α) × β)} {k : α} {fallback : β} {not_mem : (l.map Sigma.fst).contains k = false} (h : m.1.WF) : - getD (m.insertMany l).1 k fallback = getD m k fallback:= by + Const.getD (m.insertMany l).1 k fallback = Const.getD m k fallback:= by simp_to_model using getValueD_insertList_not_mem -theorem Const.getD_insertMany_list_mem [EquivBEq α] [LawfulHashable α] +theorem getD_insertMany_list_mem_const [EquivBEq α] [LawfulHashable α] {l : List ((_ : α) × β)} {k k' : α} {k_eq: k == k'} {v fallback: β} {mem : ⟨k,v⟩ ∈ l} {distinct : List.Pairwise (fun a b => (a.1 == b.1) = false ) l} (h : m.1.WF): - getD (m.insertMany l).1 k' fallback = v := by + Const.getD (m.insertMany l).1 k' fallback = v := by simp_to_model using getValueD_insertList_mem +namespace Const +@[simp] +theorem insertMany_list_nil : + insertMany m [] = m := by + simp [insertMany, Id.run] + +@[simp] +theorem insertMany_list_singleton [EquivBEq α] [LawfulHashable α] {k : α} {v : β} : + insertMany m [⟨k,v⟩] = m.insert k v := by + simp [insertMany, Id.run] + +theorem insertMany_list_cons {l : List (α × β)} {k : α} {v : β} : + (insertMany m (⟨k,v⟩ :: l)).1 = (insertMany (m.insert k v) l).1 := by + simp only [insertMany_eq_insertListₘ] + cases l with + | nil => simp [insertListₘ] + | cons hd tl => simp [insertListₘ] + +theorem contains_insertMany_list [EquivBEq α] [LawfulHashable α] {l : List (α × β)} {k : α} (h : m.1.WF) : + (Const.insertMany m l).1.contains k = (m.contains k || (l.map Prod.fst).contains k) := by + simp_to_model using List.containsKey_insertListConst + +theorem contains_of_contains_insertMany_list [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l : List ( α × β )} {k : α} : + (insertMany m l).1.contains k → (l.map Prod.fst).contains k = false → m.contains k := by + simp_to_model using List.containsKey_of_containsKey_insertListConst + +theorem getKey?_insertMany_list_not_mem [EquivBEq α] [LawfulHashable α] (h : m.1.WF) + {l : List (α × β)} {k : α} + {not_mem : (l.map Prod.fst).contains k = false} : + (insertMany m l).1.getKey? k = m.getKey? k := by + simp_to_model using List.getKey?_insertListConst_not_mem + +theorem getKey?_insertMany_list_mem [EquivBEq α] [LawfulHashable α] (h : m.1.WF) + {l : List (α × β)} + {k k' : α} {k_beq : k == k'} + {distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)} + {mem : k ∈ (l.map Prod.fst)} : + (insertMany m l).1.getKey? k' = some k := by + simp_to_model using List.getKey?_insertListConst_mem + +theorem getKey_insertMany_list_not_mem [EquivBEq α] [LawfulHashable α] (h : m.1.WF) + {l : List (α × β)} {k : α} + {not_mem : (l.map Prod.fst).contains k = false} + {h'} : + (insertMany m l).1.getKey k h' = m.getKey k (contains_of_contains_insertMany_list _ h h' not_mem) := by + simp_to_model using List.getKey_insertListConst_not_mem + +theorem getKey_insertMany_list_mem [EquivBEq α] [LawfulHashable α] (h : m.1.WF) + {l : List (α × β)} + {k k' : α} {k_beq : k == k'} + {distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)} + {mem : k ∈ (l.map Prod.fst)} + {h'} : + (insertMany m l).1.getKey k' h' = k := by + simp_to_model using List.getKey_insertListConst_mem + +theorem getKey!_insertMany_list_not_mem [EquivBEq α] [LawfulHashable α] [Inhabited α] (h : m.1.WF) + {l : List (α × β)} {k : α} + {not_mem : (l.map Prod.fst).contains k = false} : + (insertMany m l).1.getKey! k = m.getKey! k := by + simp_to_model using List.getKey!_insertListConst_not_mem + +theorem getKey!_insertMany_list_mem [EquivBEq α] [LawfulHashable α] [Inhabited α] (h : m.1.WF) + {l : List (α × β)} + {k k' : α} {k_beq : k == k'} + {distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)} + {mem : k ∈ (l.map Prod.fst)} : + (insertMany m l).1.getKey! k' = k := by + simp_to_model using List.getKey!_insertListConst_mem + +theorem getKeyD_insertMany_list_not_mem [EquivBEq α] [LawfulHashable α] (h : m.1.WF) + {l : List (α × β)} {k fallback : α} + {not_mem : (l.map Prod.fst).contains k = false} : + (insertMany m l).1.getKeyD k fallback = m.getKeyD k fallback := by + simp_to_model using List.getKeyD_insertListConst_not_mem + +theorem getKeyD_insertMany_list_mem [EquivBEq α] [LawfulHashable α] (h : m.1.WF) + {l : List (α × β)} + {k k' fallback : α} {k_beq : k == k'} + {distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)} + {mem : k ∈ (l.map Prod.fst)} : + (insertMany m l).1.getKeyD k' fallback = k := by + simp_to_model using List.getKeyD_insertListConst_mem + +theorem size_insertMany_list [EquivBEq α] [LawfulHashable α] + {l : List (α × β)} + {distinct : List.Pairwise (fun a b => (a.1 == b.1) = false) l} + (h : m.1.WF) : + (∀ (a : α), ¬ (m.contains a = true ∧ (l.map Prod.fst).contains a = true)) → + (insertMany m l).1.1.size = m.1.size + l.length := by + simp_to_model using List.insertListConst_length + +theorem insertMany_list_notEmpty_if_map_notEmpty [EquivBEq α] [LawfulHashable α] + {l : List (α × β)} (h : m.1.WF) : + (m.1.isEmpty = false) → (insertMany m l).1.1.isEmpty = false := by + simp_to_model using List.insertListConst_not_isEmpty_if_start_not_isEmpty + +theorem insertMany_list_isEmpty [EquivBEq α] [LawfulHashable α] + {l : List (α × β)} (h : m.1.WF) : + (insertMany m l).1.1.isEmpty ↔ m.1.isEmpty ∧ l.isEmpty := by + simp_to_model using List.insertListConst_isEmpty + +theorem get?_insertMany_list_not_mem_const [EquivBEq α] [LawfulHashable α] + {l : List (α × β)} {k : α} + {not_mem : (l.map Prod.fst).contains k = false} + (h : m.1.WF) : + get? (insertMany m l).1 k = get? m k := by + simp_to_model using getValue?_insertListConst_not_mem + +theorem get?_insertMany_list_mem_const [EquivBEq α] [LawfulHashable α] + {l : List (α × β)} {k k' : α} {k_eq: k == k'} {v : β} {mem : ⟨k,v⟩ ∈ l} + {distinct : List.Pairwise (fun a b => (a.1 == b.1) = false ) l} (h : m.1.WF): + get? (insertMany m l).1 k' = v := by + simp_to_model using getValue?_insertListConst_mem + +theorem get_insertMany_list_not_mem_const [EquivBEq α] [LawfulHashable α] + {l : List (α × β)} {k : α} + {not_mem : (l.map Prod.fst).contains k = false} + (h : m.1.WF) {h'}: + get (insertMany m l).1 k h' = get m k (contains_of_contains_insertMany_list _ h h' not_mem) := by + simp_to_model using getValue_insertListConst_not_mem + +theorem get_insertMany_list_mem_const [EquivBEq α] [LawfulHashable α] + {l : List (α × β)} {k k' : α} {k_eq: k == k'} {v : β} {mem : ⟨k,v⟩ ∈ l} + {distinct : List.Pairwise (fun a b => (a.1 == b.1) = false ) l} (h : m.1.WF) {h'}: + get (insertMany m l).1 k' h' = v := by + simp_to_model using getValue_insertListConst_mem + +theorem get!_insertMany_list_not_mem_const [EquivBEq α] [LawfulHashable α] [Inhabited β] + {l : List (α × β)} {k : α} + {not_mem : (l.map Prod.fst).contains k = false} + (h : m.1.WF) : + get! (insertMany m l).1 k = get! m k := by + simp_to_model using getValue!_insertListConst_not_mem + +theorem get!_insertMany_list_mem_const [EquivBEq α] [LawfulHashable α] [Inhabited β] + {l : List (α × β)} {k k' : α} {k_eq: k == k'} {v : β} {mem : ⟨k,v⟩ ∈ l} + {distinct : List.Pairwise (fun a b => (a.1 == b.1) = false ) l} (h : m.1.WF): + get! (insertMany m l).1 k' = v := by + simp_to_model using getValue!_insertListConst_mem + +theorem getD_insertMany_list_not_mem_const [EquivBEq α] [LawfulHashable α] + {l : List (α × β)} {k : α} {fallback : β} + {not_mem : (l.map Prod.fst).contains k = false} + (h : m.1.WF) : + getD (insertMany m l).1 k fallback = getD m k fallback:= by + simp_to_model using getValueD_insertListConst_not_mem + +theorem getD_insertMany_list_mem_const [EquivBEq α] [LawfulHashable α] + {l : List (α × β)} {k k' : α} {k_eq: k == k'} {v fallback: β} {mem : ⟨k,v⟩ ∈ l} + {distinct : List.Pairwise (fun a b => (a.1 == b.1) = false ) l} (h : m.1.WF): + getD (insertMany m l).1 k' fallback = v := by + simp_to_model using getValueD_insertListConst_mem +end Const end end Raw₀ From 49e4a9dec444bdc2153e1f88870d9d4a1dbb6ad8 Mon Sep 17 00:00:00 2001 From: jt0202 Date: Sun, 1 Dec 2024 20:11:13 +0100 Subject: [PATCH 31/83] Verify insertManyIfNewUnit --- .../DHashMap/Internal/List/Associative.lean | 339 +++++++++++++++++- src/Std/Data/DHashMap/Internal/Model.lean | 2 +- src/Std/Data/DHashMap/Internal/RawLemmas.lean | 128 ++++++- src/Std/Data/DHashMap/Internal/WF.lean | 19 + 4 files changed, 484 insertions(+), 4 deletions(-) diff --git a/src/Std/Data/DHashMap/Internal/List/Associative.lean b/src/Std/Data/DHashMap/Internal/List/Associative.lean index 40ddb75c770d..2ffb77e16242 100644 --- a/src/Std/Data/DHashMap/Internal/List/Associative.lean +++ b/src/Std/Data/DHashMap/Internal/List/Associative.lean @@ -1264,6 +1264,15 @@ theorem insertEntryIfNew_of_containsKey_eq_false [BEq α] {l : List ((a : α) × {v : β k} (h : containsKey k l = false) : insertEntryIfNew k v l = ⟨k, v⟩ :: l := by simp_all [insertEntryIfNew] +theorem DistinctKeys.insertEntryIfNew [BEq α] [PartialEquivBEq α] {k : α} {v : β k} {l : List ((a : α) × β a)} (h: DistinctKeys l): + DistinctKeys (insertEntryIfNew k v l) := by + simp only [Std.DHashMap.Internal.List.insertEntryIfNew, cond_eq_if] + split + · exact h + · rw [distinctKeys_cons_iff] + rename_i h' + simp [h, h'] + @[simp] theorem isEmpty_insertEntryIfNew [BEq α] {l : List ((a : α) × β a)} {k : α} {v : β k} : (insertEntryIfNew k v l).isEmpty = false := by @@ -2714,5 +2723,333 @@ theorem getValue_insertListConst_mem [BEq α] [PartialEquivBEq α] . exact k_beq . exact distinct - end +/-- Internal implementation detail of the hash map -/ +def insertListIfNewUnit [BEq α] (l: List ((_ : α) × Unit))(toInsert: List α) : List ((_ : α) × Unit) := + match toInsert with + | .nil => l + | .cons hd tl => insertListIfNewUnit (insertEntryIfNew hd () l) tl + +theorem insertListIfNewUnit_perm_of_perm_first [BEq α] [EquivBEq α] {l1 l2 : List ((_ : α) × Unit)} {toInsert : List α } (h : Perm l1 l2) + (distinct : DistinctKeys l1) : Perm (insertListIfNewUnit l1 toInsert) (insertListIfNewUnit l2 toInsert) := by + induction toInsert generalizing l1 l2 with + | nil => simp [insertListIfNewUnit, h] + | cons hd tl ih => + simp only [insertListIfNewUnit] + apply ih + · simp only [insertEntryIfNew, cond_eq_if] + have contains_eq : containsKey hd l1 = containsKey hd l2 := containsKey_of_perm h + rw [contains_eq] + by_cases contains_hd: containsKey hd l2 = true + · simp only [contains_hd, ↓reduceIte] + exact h + · simp only [contains_hd, Bool.false_eq_true, ↓reduceIte, List.perm_cons] + exact h + · apply DistinctKeys.insertEntryIfNew distinct + +theorem DistinctKeys.insertListIfNewUnit [BEq α] [PartialEquivBEq α] (l : List ((_ : α) × Unit))(toInsert : List α) (distinct: DistinctKeys l): + DistinctKeys (insertListIfNewUnit l toInsert) := by + induction toInsert generalizing l with + | nil => simp [List.insertListIfNewUnit, distinct] + | cons hd tl ih => + simp only [List.insertListIfNewUnit] + apply ih + apply insertEntryIfNew + exact distinct + +theorem containsKey_insertListIfNewUnit [BEq α] [PartialEquivBEq α] {l : List ((_ : α) × Unit)} {toInsert : List α} {k : α} : + containsKey k (insertListIfNewUnit l toInsert) = (containsKey k l || toInsert.contains k) := by + induction toInsert generalizing l with + | nil => simp [insertListIfNewUnit] + | cons hd tl ih => + simp only [insertListIfNewUnit, List.contains_cons] + rw [ih, containsKey_insertEntryIfNew] + rw [Bool.or_comm (hd == k), Bool.or_assoc, BEq.comm (a:=hd)] + +theorem containsKey_of_containsKey_insertListIfNewUnit [BEq α] [PartialEquivBEq α] {l : List ((_ : α) × Unit)} {toInsert : List α} {k : α} + (h₂ : toInsert.contains k = false) : containsKey k (insertListIfNewUnit l toInsert) → containsKey k l := by + intro h₁ + rw [containsKey_insertListIfNewUnit] at h₁ + simp only [Bool.or_eq_true] at h₁ + cases h₁ with + | inl h => exact h + | inr h => + simp only [h, Bool.true_eq_false] at h₂ + +theorem getKey?_insertListIfNewUnit_not_mem [BEq α] [EquivBEq α] + {l : List ((_ : α) × Unit)} {toInsert : List α} {k : α} + {not_mem : toInsert.contains k = false} : + getKey? k (insertListIfNewUnit l toInsert) = getKey? k l := by + induction toInsert generalizing l with + | nil => simp [insertListIfNewUnit] + | cons hd tl ih => + simp only [insertListIfNewUnit] + rw [ih] + · rw [getKey?_insertEntryIfNew] + split + · rename_i hd_k + simp at not_mem + exfalso + rcases hd_k with ⟨h,_⟩ + rcases not_mem with ⟨h',_⟩ + rw [BEq.symm h] at h' + simp at h' + · rfl + · simp only [List.contains_cons, Bool.or_eq_false_iff] at not_mem + apply And.right not_mem + +theorem getKey?_insertListIfNewUnit_mem [BEq α] [EquivBEq α] + {l : List ((_ : α) × Unit)} {toInsert : List α} + {k k' : α} {k_beq : k == k'} + {distinct : toInsert.Pairwise (fun a b => (a == b) = false)} + {mem : k ∈ toInsert} {mem' : containsKey k l = false}: + getKey? k' (insertListIfNewUnit l toInsert) = some k := by + induction toInsert generalizing l with + | nil => simp at mem + | cons hd tl ih => + simp only [insertListIfNewUnit] + simp only [List.mem_cons] at mem + cases mem with + | inl mem => + rw [getKey?_insertListIfNewUnit_not_mem, ← mem] + · simp only [insertEntryIfNew, mem', cond_eq_if, Bool.false_eq_true, ↓reduceIte, + getKey?_cons, ite_eq_left_iff, Bool.not_eq_true] + intro h + rw [h] at k_beq + simp only [Bool.false_eq_true] at k_beq + · simp only [pairwise_cons] at distinct + simp only [List.contains_eq_any_beq, List.any_eq_false, Bool.not_eq_true] + intro a h + apply BEq.neq_of_beq_of_neq + apply BEq.symm k_beq + rw [mem] + apply (And.left distinct) a h + | inr mem => + apply ih + · simp only [pairwise_cons] at distinct + apply And.right distinct + · exact mem + · rw [containsKey_insertEntryIfNew] + simp only [Bool.or_eq_false_iff] + constructor + · simp only [pairwise_cons] at distinct + apply (And.left distinct) k mem + · exact mem' + +theorem getKey?_insertListIfNewUnit_mem_both [BEq α] [EquivBEq α] + {l : List ((_ : α) × Unit)} {toInsert : List α} + {k : α} + {distinct : toInsert.Pairwise (fun a b => (a == b) = false)} + {mem : toInsert.contains k} {mem' : containsKey k l = true}: + getKey? k (insertListIfNewUnit l toInsert) = getKey? k l := by + induction toInsert generalizing l with + | nil => simp at mem + | cons hd tl ih => + simp at mem + cases mem with + | inl mem => + simp only [insertListIfNewUnit] + rw [getKey?_insertListIfNewUnit_not_mem, getKey?_insertEntryIfNew] + · simp only [BEq.symm mem, true_and, ite_eq_right_iff] + have h':= containsKey_of_beq mem' mem + intro h + rw [h] at h' + simp only [Bool.false_eq_true] at h' + · simp only [pairwise_cons] at distinct + simp only [List.contains_eq_any_beq, List.any_eq_false, Bool.not_eq_true] + intro x x_mem + apply BEq.neq_of_beq_of_neq mem + apply And.left distinct x x_mem + | inr mem => + simp only [insertListIfNewUnit] + simp only [pairwise_cons] at distinct + rw [ih (distinct := And.right distinct) (mem := mem)] + · rw [getKey?_insertEntryIfNew] + split + · rename_i h + have h':= containsKey_of_beq mem' (BEq.symm (And.left h)) + rw [And.right h] at h' + simp only [Bool.false_eq_true] at h' + · rfl + · rw [containsKey_insertEntryIfNew] + simp only [Bool.or_eq_true] + right + apply mem' + +theorem getKey_insertListIfNewUnit_not_mem [BEq α] [EquivBEq α] + {l : List ((_ : α) × Unit)} {toInsert : List α} {k : α} + {not_mem : toInsert.contains k = false} + {h} : + getKey k (insertListIfNewUnit l toInsert) h = + getKey k l (containsKey_of_containsKey_insertListIfNewUnit not_mem h) := by + rw [← Option.some_inj] + rw [← getKey?_eq_some_getKey] + rw [getKey?_insertListIfNewUnit_not_mem] + . rw [getKey?_eq_some_getKey] + . exact not_mem + +theorem getKey_insertListIfNewUnit_mem [BEq α] [EquivBEq α] + {l : List ((_ : α) × Unit)} {toInsert : List α} + {k k' : α} {k_beq : k == k'} + {distinct : toInsert.Pairwise (fun a b => (a == b) = false)} + {mem : k ∈ toInsert} + {h} : containsKey k l = false → + getKey k' (insertListIfNewUnit l toInsert) h = k := by + intro mem' + rw [← Option.some_inj] + rw [← getKey?_eq_some_getKey] + rw [getKey?_insertListIfNewUnit_mem] + . exact k_beq + . exact distinct + . exact mem + . exact mem' + +theorem getKey_insertListIfNewUnit_mem_both [BEq α] [EquivBEq α] + {l : List ((_ : α) × Unit)} {toInsert : List α} + {k : α} + {distinct : toInsert.Pairwise (fun a b => (a == b) = false)} + {mem : toInsert.contains k} {mem' : containsKey k l = true} {h}: + getKey k (insertListIfNewUnit l toInsert) h = getKey k l mem' := by + rw [← Option.some_inj] + rw [← getKey?_eq_some_getKey] + rw [← getKey?_eq_some_getKey] + rw [getKey?_insertListIfNewUnit_mem_both] + · exact distinct + · exact mem + · exact mem' + +theorem getKey!_insertListIfNewUnit_not_mem [BEq α] [EquivBEq α] [Inhabited α] + {l : List ((_ : α) × Unit)} {toInsert : List α} {k : α} + {not_mem : toInsert.contains k = false} : + getKey! k (insertListIfNewUnit l toInsert) = getKey! k l := by + rw [getKey!_eq_getKey?] + rw [getKey?_insertListIfNewUnit_not_mem] + . rw [getKey!_eq_getKey?] + . exact not_mem + +theorem getKey!_insertListIfNewUnit_mem [BEq α] [EquivBEq α] [Inhabited α] + {l : List ((_ : α) × Unit)} {toInsert : List α} + {k k' : α} {k_beq : k == k'} + {distinct : toInsert.Pairwise (fun a b => (a == b) = false)} + {mem : k ∈ toInsert} {mem': containsKey k l = false} : + getKey! k' (insertListIfNewUnit l toInsert) = k := by + rw [getKey!_eq_getKey?] + rw [getKey?_insertListIfNewUnit_mem] + . rw [Option.get!_some] + . exact k_beq + . exact distinct + . exact mem + . exact mem' + +theorem getKey!_insertListIfNewUnit_mem_both [BEq α] [EquivBEq α] [Inhabited α] + {l : List ((_ : α) × Unit)} {toInsert : List α} + {k : α} + {distinct : toInsert.Pairwise (fun a b => (a == b) = false)} + {mem : toInsert.contains k} {mem' : containsKey k l = true} : + getKey! k (insertListIfNewUnit l toInsert) = getKey! k l := by + rw [getKey!_eq_getKey?] + rw [getKey?_insertListIfNewUnit_mem_both] + . rw [getKey!_eq_getKey?] + · exact distinct + · exact mem + · exact mem' + +theorem getKeyD_insertListIfNewUnit_not_mem [BEq α] [EquivBEq α] + {l : List ((_ : α) × Unit)} {toInsert : List α} {k fallback : α} + {not_mem : toInsert.contains k = false} : + getKeyD k (insertListIfNewUnit l toInsert) fallback = getKeyD k l fallback := by + rw [getKeyD_eq_getKey?] + rw [getKey?_insertListIfNewUnit_not_mem] + . rw [getKeyD_eq_getKey?] + . exact not_mem + +theorem getKeyD_insertListIfNewUnit_mem [BEq α] [EquivBEq α] + {l : List ((_ : α) × Unit)} {toInsert : List α} + {k k' fallback : α} {k_beq : k == k'} + {distinct : toInsert.Pairwise (fun a b => (a == b) = false)} + {mem : k ∈ toInsert } {mem': containsKey k l = false} : + getKeyD k' (insertListIfNewUnit l toInsert) fallback = k := by + rw [getKeyD_eq_getKey?] + rw [getKey?_insertListIfNewUnit_mem] + . rw [Option.getD_some] + . exact k_beq + . exact distinct + . exact mem + . exact mem' + +theorem getKeyD_insertListIfNewUnit_mem_both [BEq α] [EquivBEq α] + {l : List ((_ : α) × Unit)} {toInsert : List α} + {k fallback : α} + {distinct : toInsert.Pairwise (fun a b => (a == b) = false)} + {mem : toInsert.contains k} {mem' : containsKey k l = true} : + getKeyD k (insertListIfNewUnit l toInsert) fallback = getKeyD k l fallback := by + rw [getKeyD_eq_getKey?] + rw [getKey?_insertListIfNewUnit_mem_both] + . rw [getKeyD_eq_getKey?] + · exact distinct + · exact mem + · exact mem' + +theorem insertListIfNewUnit_length [BEq α] [EquivBEq α] + {l : List ((_ : α) × Unit)} {toInsert : List α} + {distinct_l : DistinctKeys l} + (distinct_toInsert : List.Pairwise (fun a b => (a == b) = false) toInsert) + (distinct_both : ∀ (a : α), ¬ (containsKey a l ∧ toInsert.contains a)) : + (insertListIfNewUnit l toInsert).length = l.length + toInsert.length := by + induction toInsert generalizing l with + | nil => simp [insertListIfNewUnit] + | cons hd tl ih => + simp only [insertListIfNewUnit, List.length_cons] + rw [ih] + · rw [length_insertEntryIfNew] + specialize distinct_both hd + simp only [List.contains_cons, BEq.refl, Bool.true_or, and_true, + Bool.not_eq_true] at distinct_both + simp only [distinct_both, Bool.false_eq_true, ↓reduceIte] + rw [Nat.add_assoc, Nat.add_comm 1 _] + · simp only [pairwise_cons] at distinct_toInsert + apply And.right distinct_toInsert + · intro a + simp only [not_and, Bool.not_eq_true] + simp only [List.contains_cons, Bool.or_eq_true, not_and, not_or, + Bool.not_eq_true] at distinct_both + rw [containsKey_insertEntryIfNew] + simp only [Bool.or_eq_true] + intro h + cases h with + | inl h => + simp only [pairwise_cons] at distinct_toInsert + rw [List.contains_eq_any_beq] + simp only [List.any_eq_false, Bool.not_eq_true] + intro x x_mem + rcases distinct_toInsert with ⟨left,_⟩ + specialize left x x_mem + apply BEq.neq_of_beq_of_neq + apply BEq.symm h + exact left + | inr h => + specialize distinct_both a h + apply And.right distinct_both + · apply DistinctKeys.insertEntryIfNew distinct_l + +theorem insertListIfNewUnit_not_isEmpty_if_start_not_isEmpty [BEq α] + {l : List ((_ : α) × Unit)} {toInsert : List α} + {h : l.isEmpty = false} : + (List.insertListIfNewUnit l toInsert).isEmpty = false := by + induction toInsert generalizing l with + | nil => simp [insertListIfNewUnit, h] + | cons hd tl ih => + simp [insertListIfNewUnit, ih] + +theorem insertListIfNewUnit_isEmpty [BEq α] + {l : List ((_ : α) × Unit)} {toInsert : List α} : + (List.insertListIfNewUnit l toInsert).isEmpty ↔ l.isEmpty ∧ toInsert.isEmpty := by + induction toInsert with + | nil => simp [insertListIfNewUnit] + | cons hd tl ih => + simp only [insertListIfNewUnit, List.isEmpty_cons, Bool.false_eq_true, and_false, + iff_false] + apply ne_true_of_eq_false + apply insertListIfNewUnit_not_isEmpty_if_start_not_isEmpty + simp +end end List diff --git a/src/Std/Data/DHashMap/Internal/Model.lean b/src/Std/Data/DHashMap/Internal/Model.lean index a5e6f27bed53..b221595ae2ff 100644 --- a/src/Std/Data/DHashMap/Internal/Model.lean +++ b/src/Std/Data/DHashMap/Internal/Model.lean @@ -529,7 +529,7 @@ theorem Const.insertMany_eq_insertListₘ [BEq α] [Hashable α] (m : Raw₀ α simp only [List.foldl_cons,insertListₘ] apply ih -theorem Const.insertManyIfNewUnit_insertListIfNewUnit [BEq α] [Hashable α] (m : Raw₀ α (fun _ => Unit)) (l: List α): +theorem Const.insertManyIfNewUnit_eq_insertListIfNewUnit [BEq α] [Hashable α] (m : Raw₀ α (fun _ => Unit)) (l: List α): (Const.insertManyIfNewUnit m l).1 = Const.insertListIfNewUnitₘ m l := by simp only [insertManyIfNewUnit, Id.run, Id.pure_eq, Id.bind_eq, List.forIn_yield_eq_foldl] suffices ∀ (t : { m' // ∀ (P : Raw₀ α (fun _ => Unit) → Prop), diff --git a/src/Std/Data/DHashMap/Internal/RawLemmas.lean b/src/Std/Data/DHashMap/Internal/RawLemmas.lean index efd2225391ca..a50052e73437 100644 --- a/src/Std/Data/DHashMap/Internal/RawLemmas.lean +++ b/src/Std/Data/DHashMap/Internal/RawLemmas.lean @@ -70,7 +70,7 @@ variable [BEq α] [Hashable α] /-- Internal implementation detail of the hash map -/ scoped macro "wf_trivial" : tactic => `(tactic| repeat (first - | apply Raw₀.wfImp_insert | apply Raw₀.wfImp_insertIfNew | apply Raw₀.wfImp_insertMany | apply Raw₀.Const.wfImp_insertMany| apply Raw₀.wfImp_erase + | apply Raw₀.wfImp_insert | apply Raw₀.wfImp_insertIfNew | apply Raw₀.wfImp_insertMany | apply Raw₀.Const.wfImp_insertMany| apply Raw₀.Const.wfImp_insertManyIfNewUnit | apply Raw₀.wfImp_erase | apply Raw.WF.out | assumption | apply Raw₀.wfImp_empty | apply Raw.WFImp.distinct | apply Raw.WF.empty₀)) @@ -91,7 +91,7 @@ private def queryNames : Array Name := private def modifyNames : Array Name := #[``toListModel_insert, ``toListModel_erase, ``toListModel_insertIfNew, ``toListModel_insertMany_list, - ``Const.toListModel_insertMany_list] + ``Const.toListModel_insertMany_list, ``Const.toListModel_insertManyIfNewUnit_list] private def congrNames : MacroM (Array (TSyntax `term)) := do return #[← `(_root_.List.Perm.isEmpty_eq), ← `(containsKey_of_perm), @@ -1211,6 +1211,130 @@ theorem getD_insertMany_list_mem_const [EquivBEq α] [LawfulHashable α] {distinct : List.Pairwise (fun a b => (a.1 == b.1) = false ) l} (h : m.1.WF): getD (insertMany m l).1 k' fallback = v := by simp_to_model using getValueD_insertListConst_mem + + +variable (m: Raw₀ α (fun _ => Unit)) + +@[simp] +theorem insertManyIfNewUnit_list_nil : + insertManyIfNewUnit m [] = m := by + simp [insertManyIfNewUnit, Id.run] + +@[simp] +theorem insertManyIfNewUnit_list_singleton [EquivBEq α] [LawfulHashable α] {k : α} : + insertManyIfNewUnit m [k] = m.insertIfNew k () := by + simp [insertManyIfNewUnit, Id.run] + +theorem contains_insertManyIfNewUnit_list [EquivBEq α] [LawfulHashable α] {l : List α} {k : α} (h : m.1.WF): + (insertManyIfNewUnit m l).1.contains k = (m.contains k || l.contains k) := by + simp_to_model using containsKey_insertListIfNewUnit + +theorem contains_of_contains_insertManyIfNewUnit_list [EquivBEq α] [LawfulHashable α] {l : List α} {k : α} (h : m.1.WF) + (h₂ : l.contains k = false) : (insertManyIfNewUnit m l).1.contains k = true → m.contains k := by + simp_to_model using containsKey_of_containsKey_insertListIfNewUnit + +theorem getKey?_insertManyIfNewUnit_list_not_mem [EquivBEq α] [LawfulHashable α] + {l : List α} {k : α} + {not_mem : l.contains k = false} (h : m.1.WF) : + getKey? (insertManyIfNewUnit m l).1 k = getKey? m k := by + simp_to_model using getKey?_insertListIfNewUnit_not_mem + +theorem getKey?_insertManyIfNewUnit_list_mem [EquivBEq α] [LawfulHashable α] + {l : List α} {k k' : α} {k_beq : k == k'} + {distinct : l.Pairwise (fun a b => (a == b) = false)} + {mem : k ∈ l} (h : m.1.WF) : m.contains k = false → + getKey? (insertManyIfNewUnit m l).1 k' = some k := by + simp_to_model using getKey?_insertListIfNewUnit_mem + +theorem getKey?_insertManyIfNewUnit_list_mem_both [EquivBEq α] [LawfulHashable α] + {l : List α} {k : α} + {distinct : l.Pairwise (fun a b => (a == b) = false)} + {mem : l.contains k} (h : m.1.WF): + m.contains k = true → getKey? (insertManyIfNewUnit m l).1 k = getKey? m k := by + simp_to_model using getKey?_insertListIfNewUnit_mem_both + +theorem getKey_insertManyIfNewUnit_list_not_mem [EquivBEq α] [LawfulHashable α] + {l : List α} {k : α} + {not_mem : l.contains k = false} + {h'} (h: m.1.WF): + getKey (insertManyIfNewUnit m l).1 k h' = + getKey m k (contains_of_contains_insertManyIfNewUnit_list m h not_mem h') := by + simp_to_model using getKey_insertListIfNewUnit_not_mem + +theorem getKey_insertManyIfNewUnit_list_mem [EquivBEq α] [LawfulHashable α] + {l : List α} + {k k' : α} {k_beq : k == k'} + {distinct : l.Pairwise (fun a b => (a == b) = false)} + {mem : k ∈ l}{h'} (h: m.1.WF): + m.contains k = false → + getKey (insertManyIfNewUnit m l).1 k' h' = k := by + simp_to_model using getKey_insertListIfNewUnit_mem + +theorem getKey_insertManyIfNewUnit_list_mem_both [EquivBEq α] [LawfulHashable α] + {l : List α} {k : α} + {distinct : l.Pairwise (fun a b => (a == b) = false)} + {mem : l.contains k} {mem' : m.contains k = true} {h'} (h : m.1.WF): + getKey (insertManyIfNewUnit m l).1 k h' = getKey m k mem' := by + simp_to_model using getKey_insertListIfNewUnit_mem_both + +theorem getKey!_insertManyIfNewUnit_list_not_mem [EquivBEq α] [LawfulHashable α] [Inhabited α] + {l : List α} {k : α} + {not_mem : l.contains k = false} (h : m.1.WF) : + getKey! (insertManyIfNewUnit m l).1 k = getKey! m k := by + simp_to_model using getKey!_insertListIfNewUnit_not_mem + + +theorem getKey!_insertManyIfNewUnit_list_mem [EquivBEq α] [LawfulHashable α] [Inhabited α] + {l : List α} {k k' : α} {k_beq : k == k'} + {distinct : l.Pairwise (fun a b => (a == b) = false)} + {mem : k ∈ l} (h : m.1.WF) : contains m k = false → + getKey! (insertManyIfNewUnit m l).1 k' = k := by + simp_to_model using getKey!_insertListIfNewUnit_mem + +theorem getKey!_insertManyIfNewUnit_list_mem_both [EquivBEq α] [LawfulHashable α] [Inhabited α] + {l : List α} {k : α} + {distinct : l.Pairwise (fun a b => (a == b) = false)} + {mem : l.contains k} (h : m.1.WF) : m.contains k = true → + getKey! (insertManyIfNewUnit m l).1 k = getKey! m k := by + simp_to_model using getKey!_insertListIfNewUnit_mem_both + +theorem getKeyD_insertManyIfNewUnit_list_not_mem [EquivBEq α] [LawfulHashable α] + {l : List α} {k fallback : α} + {not_mem : l.contains k = false} (h : m.1.WF): + getKeyD (insertManyIfNewUnit m l).1 k fallback = getKeyD m k fallback := by + simp_to_model using getKeyD_insertListIfNewUnit_not_mem + +theorem getKeyD_insertManyIfNewUnit_list_mem [EquivBEq α] [LawfulHashable α] + {l : List α} {k k' fallback : α} {k_beq : k == k'} + {distinct : l.Pairwise (fun a b => (a == b) = false)} + {mem : k ∈ l } (h : m.1.WF) : m.contains k = false → + getKeyD (insertManyIfNewUnit m l).1 k' fallback = k := by + simp_to_model using getKeyD_insertListIfNewUnit_mem + +theorem getKeyD_insertManyIfNewUnit_list_mem_both [EquivBEq α] [LawfulHashable α] + {l : List α} {k fallback : α} + {distinct : l.Pairwise (fun a b => (a == b) = false)} + {mem : l.contains k} (h : m.1.WF) : m.contains k = true → + getKeyD (insertManyIfNewUnit m l).1 k fallback = getKeyD m k fallback := by + simp_to_model using getKeyD_insertListIfNewUnit_mem_both + +theorem insertManyIfNewUnit_list_length [EquivBEq α] [LawfulHashable α] + {l : List α} + (distinct : List.Pairwise (fun a b => (a == b) = false) l) + (h : m.1.WF) : (∀ (a : α), ¬ (m.contains a = true ∧ l.contains a = true)) → + (insertManyIfNewUnit m l).1.1.size = m.1.size + l.length := by + simp_to_model using insertListIfNewUnit_length + +theorem insertManyIfNewUnit_list_not_isEmpty_if_start_not_isEmpty [EquivBEq α] [LawfulHashable α] + {l : List α} (h : m.1.WF) : m.1.isEmpty = false → + (insertManyIfNewUnit m l).1.1.isEmpty = false := by + simp_to_model using insertListIfNewUnit_not_isEmpty_if_start_not_isEmpty + +theorem insertManyIfNewUnit_list_isEmpty [EquivBEq α] [LawfulHashable α] + {l : List α} (h : m.1.WF) : + (insertManyIfNewUnit m l).1.1.isEmpty ↔ m.1.isEmpty ∧ l.isEmpty := by + simp_to_model using insertListIfNewUnit_isEmpty + end Const end end Raw₀ diff --git a/src/Std/Data/DHashMap/Internal/WF.lean b/src/Std/Data/DHashMap/Internal/WF.lean index 05f523051247..e70b20fda488 100644 --- a/src/Std/Data/DHashMap/Internal/WF.lean +++ b/src/Std/Data/DHashMap/Internal/WF.lean @@ -788,6 +788,25 @@ theorem Const.wfImp_insertMany {β : Type v} [BEq α] [Hashable α] [EquivBEq α {l : ρ} (h : Raw.WFImp m.1) : Raw.WFImp (Const.insertMany m l).1.1 := Raw.WF.out ((Const.insertMany m l).2 _ Raw.WF.insert₀ (.wf m.2 h)) +/-! # `Const.insertListIfNewUnitₘ` -/ +theorem Const.toListModel_insertListIfNewUnitₘ [BEq α] [Hashable α] [EquivBEq α] [LawfulHashable α] {m : Raw₀ α (fun _ => Unit)} {l : List α} (h : Raw.WFImp m.1): + Perm (toListModel (Const.insertListIfNewUnitₘ m l).1.buckets) (List.insertListIfNewUnit (toListModel m.1.buckets) l) := by + induction l generalizing m with + | nil => simp [insertListIfNewUnitₘ, List.insertListIfNewUnit] + | cons hd tl ih => + simp only [insertListIfNewUnitₘ, insertListIfNewUnit] + apply Perm.trans + apply ih (wfImp_insertIfNew h) + apply List.insertListIfNewUnit_perm_of_perm_first + apply toListModel_insertIfNew h + apply (wfImp_insertIfNew h).distinct + +/-! # `Const.insertManyIfNewUnit` -/ +theorem Const.toListModel_insertManyIfNewUnit_list [BEq α] [Hashable α] [EquivBEq α] [LawfulHashable α] {m : Raw₀ α (fun _ => Unit)} {l : List α} (h : Raw.WFImp m.1): + Perm (toListModel (Const.insertManyIfNewUnit m l).1.1.buckets) (List.insertListIfNewUnit (toListModel m.1.buckets) l) := by + rw [Const.insertManyIfNewUnit_eq_insertListIfNewUnit] + apply toListModel_insertListIfNewUnitₘ h + theorem Const.wfImp_insertManyIfNewUnit [BEq α] [Hashable α] [EquivBEq α] [LawfulHashable α] {ρ : Type w} [ForIn Id ρ α] {m : Raw₀ α (fun _ => Unit)} {l : ρ} (h : Raw.WFImp m.1) : Raw.WFImp (Const.insertManyIfNewUnit m l).1.1 := From 8340caf43d2a5fa380ef103e8a5fabd45d4fca7f Mon Sep 17 00:00:00 2001 From: jt0202 Date: Mon, 2 Dec 2024 16:01:48 +0100 Subject: [PATCH 32/83] get results for insertManyIfNewUnit --- .../DHashMap/Internal/List/Associative.lean | 50 +++++++++++++++++++ src/Std/Data/DHashMap/Internal/RawLemmas.lean | 23 ++++++++- 2 files changed, 71 insertions(+), 2 deletions(-) diff --git a/src/Std/Data/DHashMap/Internal/List/Associative.lean b/src/Std/Data/DHashMap/Internal/List/Associative.lean index 2ffb77e16242..5cd62bae40d3 100644 --- a/src/Std/Data/DHashMap/Internal/List/Associative.lean +++ b/src/Std/Data/DHashMap/Internal/List/Associative.lean @@ -3051,5 +3051,55 @@ theorem insertListIfNewUnit_isEmpty [BEq α] apply ne_true_of_eq_false apply insertListIfNewUnit_not_isEmpty_if_start_not_isEmpty simp + +theorem getValue?_list_unit [BEq α] {l : List ((_ : α) × Unit)} {k : α}: + getValue? k l = if containsKey k l = true then some () else none := by + induction l with + | nil => simp + | cons hd tl ih => + simp only [getValue?, containsKey, Bool.or_eq_true, Bool.cond_eq_ite_iff] + by_cases hd_k: (hd.fst == k) = true + · simp [hd_k] + · simp [hd_k, ih] + +theorem getValue_list_unit [BEq α] {l : List ((_ : α) × Unit)} {k : α} {h}: + getValue k l h = () := by + induction l with + | nil => simp + | cons hd tl ih => simp + +theorem getValue!_list_unit [BEq α] {l : List ((_ : α) × Unit)} {k : α} : + getValue! k l = () := by + induction l with + | nil => simp + | cons hd tl ih => simp + +theorem getValueD_list_unit [BEq α] {l : List ((_ : α) × Unit)} {k : α} {fallback : Unit} : + getValueD k l fallback = () := by + induction l with + | nil => simp + | cons hd tl ih => simp + +theorem getValue?_insertListIfNewUnit [BEq α] [PartialEquivBEq α] + {l : List ((_ : α) × Unit)} {toInsert : List α} {k : α}: + getValue? k (insertListIfNewUnit l toInsert) = + if (containsKey k l || toInsert.contains k) then some () else none := by + simp [← containsKey_insertListIfNewUnit, getValue?_list_unit] + +theorem getValue_insertListIfNewUnit [BEq α] [PartialEquivBEq α] + {l : List ((_ : α) × Unit)} {toInsert : List α} {k : α} {h}: + getValue k (insertListIfNewUnit l toInsert) h = () := by + rw [getValue_list_unit] + +theorem getValue!_insertListIfNewUnit [BEq α] [PartialEquivBEq α] + {l : List ((_ : α) × Unit)} {toInsert : List α} {k : α} : + getValue! k (insertListIfNewUnit l toInsert) = () := by + rw [getValue!_list_unit] + +theorem getValueD_insertListIfNewUnit [BEq α] [PartialEquivBEq α] + {l : List ((_ : α) × Unit)} {toInsert : List α} {k : α} {fallback : Unit}: + getValueD k (insertListIfNewUnit l toInsert) fallback = () := by + rw [getValueD_list_unit] + end end List diff --git a/src/Std/Data/DHashMap/Internal/RawLemmas.lean b/src/Std/Data/DHashMap/Internal/RawLemmas.lean index a50052e73437..fd574828cfec 100644 --- a/src/Std/Data/DHashMap/Internal/RawLemmas.lean +++ b/src/Std/Data/DHashMap/Internal/RawLemmas.lean @@ -1212,7 +1212,6 @@ theorem getD_insertMany_list_mem_const [EquivBEq α] [LawfulHashable α] getD (insertMany m l).1 k' fallback = v := by simp_to_model using getValueD_insertListConst_mem - variable (m: Raw₀ α (fun _ => Unit)) @[simp] @@ -1283,7 +1282,6 @@ theorem getKey!_insertManyIfNewUnit_list_not_mem [EquivBEq α] [LawfulHashable getKey! (insertManyIfNewUnit m l).1 k = getKey! m k := by simp_to_model using getKey!_insertListIfNewUnit_not_mem - theorem getKey!_insertManyIfNewUnit_list_mem [EquivBEq α] [LawfulHashable α] [Inhabited α] {l : List α} {k k' : α} {k_beq : k == k'} {distinct : l.Pairwise (fun a b => (a == b) = false)} @@ -1335,6 +1333,27 @@ theorem insertManyIfNewUnit_list_isEmpty [EquivBEq α] [LawfulHashable α] (insertManyIfNewUnit m l).1.1.isEmpty ↔ m.1.isEmpty ∧ l.isEmpty := by simp_to_model using insertListIfNewUnit_isEmpty +theorem get?_insertManyIfNewUnit_list [EquivBEq α] [LawfulHashable α] + {l : List α} {k : α} (h : m.1.WF) : + get? (insertManyIfNewUnit m l).1 k = + if ( m.contains k || l.contains k) then some () else none := by + simp_to_model using getValue?_insertListIfNewUnit + +theorem getValue_insertManyIfNewUnit_list + {l : List α} {k : α} {h} : + get (insertManyIfNewUnit m l).1 k h = () := by + simp + +theorem getValue!_insertManyIfNewUnit_list + {l : List α} {k : α} : + get! (insertManyIfNewUnit m l).1 k = () := by + simp + +theorem getValueD_insertManyIfNewUnit_list + {l : List α} {k : α} {fallback : Unit} : + getD (insertManyIfNewUnit m l).1 k fallback = () := by + simp + end Const end end Raw₀ From 45e3bc0e25c43d897440026efc6b8f4b1736b9cf Mon Sep 17 00:00:00 2001 From: jt0202 Date: Mon, 2 Dec 2024 23:01:56 +0100 Subject: [PATCH 33/83] Style fixes --- .../DHashMap/Internal/List/Associative.lean | 690 ++++++++++-------- src/Std/Data/DHashMap/Internal/Model.lean | 2 +- src/Std/Data/DHashMap/Internal/RawLemmas.lean | 660 +++++++++-------- src/Std/Data/DHashMap/Internal/WF.lean | 55 +- src/Std/Data/DHashMap/Raw.lean | 1 - 5 files changed, 729 insertions(+), 679 deletions(-) diff --git a/src/Std/Data/DHashMap/Internal/List/Associative.lean b/src/Std/Data/DHashMap/Internal/List/Associative.lean index 5cd62bae40d3..d52543f3d951 100644 --- a/src/Std/Data/DHashMap/Internal/List/Associative.lean +++ b/src/Std/Data/DHashMap/Internal/List/Associative.lean @@ -172,11 +172,11 @@ theorem getValueCast?_cons_self [BEq α] [LawfulBEq α] {l : List ((a : α) × getValueCast? k (⟨k, v⟩ :: l) = some v := by rw [getValueCast?_cons_of_true BEq.refl, cast_eq] -theorem getValueCast?_mem [BEq α] [LawfulBEq α] +theorem getValueCast?_of_mem [BEq α] [LawfulBEq α] {l : List ((a : α) × β a)} - {k k' : α} {v : β k} {k_eq : k == k'} {mem : ⟨k,v⟩ ∈ l} - {distinct : List.Pairwise (fun a b => (a.1 == b.1) = false) l} : - getValueCast? k' l = some (cast (by congr; apply LawfulBEq.eq_of_beq k_eq) v) := by + {k : α} (k' : α ) (v : β k) (k_beq : k == k') (mem : ⟨k, v⟩ ∈ l) + (distinct : List.Pairwise (fun a b => (a.1 == b.1) = false) l) : + getValueCast? k' l = some (cast (by congr; apply LawfulBEq.eq_of_beq k_beq) v) := by induction l with | nil => simp at mem | cons hd tl ih => @@ -185,25 +185,22 @@ theorem getValueCast?_mem [BEq α] [LawfulBEq α] cases mem with | inl mem => rw [← mem] - simp [k_eq] + simp [k_beq] | inr mem => have hd_k' : (hd.fst == k') = false := by simp only [beq_eq_false_iff_ne, ne_eq, pairwise_cons] at distinct simp only [beq_eq_false_iff_ne, ne_eq] rcases distinct with ⟨distinct, _⟩ - specialize distinct ⟨k,v⟩ mem + specialize distinct ⟨k, v⟩ mem simp only at distinct - simp only [beq_iff_eq] at k_eq - rw [k_eq] at distinct + simp only [beq_iff_eq] at k_beq + rw [k_beq] at distinct apply distinct simp only [hd_k', Bool.false_eq_true, ↓reduceDIte] apply ih · exact mem · exact List.Pairwise.of_cons distinct - - - theorem getValue?_eq_getValueCast? [BEq α] [LawfulBEq α] {β : Type v} {l : List ((_ : α) × β)} {a : α} : getValue? a l = getValueCast? a l := by induction l using assoc_induction <;> simp_all [getValueCast?_cons, getValue?_cons] @@ -930,7 +927,8 @@ theorem containsKey_of_mem [BEq α] [ReflBEq α] {l : List ((a : α) × β a)} { theorem DistinctKeys.nil [BEq α] : DistinctKeys ([] : List ((a : α) × β a)) := ⟨by simp⟩ -theorem DistinctKeys.def [BEq α] {l : List ((a : α) × β a)} : DistinctKeys l ↔ List.Pairwise (fun a b => (a.1 == b.1) = false) l := +theorem DistinctKeys.def [BEq α] {l : List ((a : α) × β a)} : + DistinctKeys l ↔ List.Pairwise (fun a b => (a.1 == b.1) = false) l := ⟨fun h => by simpa [keys_eq_map, List.pairwise_map] using h.distinct, fun h => ⟨by simpa [keys_eq_map, List.pairwise_map] using h⟩⟩ @@ -960,7 +958,6 @@ theorem DistinctKeys.of_keys_eq [BEq α] {l : List ((a : α) × β a)} {l' : Lis (h : keys l = keys l') : DistinctKeys l → DistinctKeys l' := distinctKeys_of_sublist_keys (h ▸ Sublist.refl _) - theorem containsKey_iff_exists [BEq α] [PartialEquivBEq α] {l : List ((a : α) × β a)} {a : α} : containsKey a l ↔ ∃ a' ∈ keys l, a == a' := by rw [containsKey_eq_keys_contains, List.contains_iff_exists_mem_beq] @@ -1041,8 +1038,9 @@ theorem insertEntry_of_containsKey_eq_false [BEq α] {l : List ((a : α) × β a (h : containsKey k l = false) : insertEntry k v l = ⟨k, v⟩ :: l := by simp [insertEntry, h] -theorem insertEntry_of_containsKey_eq_false_perm [BEq α] {l : List ((a : α) × β a)} {k : α} {v : β k} - (h : containsKey k l = false) : Perm (insertEntry k v l) (⟨k,v⟩ :: l) := by +theorem perm_insertEntry_of_containsKey_eq_false [BEq α] + {l : List ((a : α) × β a)} {k : α} {v : β k} + (h : containsKey k l = false) : Perm (insertEntry k v l) (⟨k, v⟩ :: l) := by cases l with | nil => simp | cons hd tl => simp [insertEntry, h] @@ -1264,7 +1262,8 @@ theorem insertEntryIfNew_of_containsKey_eq_false [BEq α] {l : List ((a : α) × {v : β k} (h : containsKey k l = false) : insertEntryIfNew k v l = ⟨k, v⟩ :: l := by simp_all [insertEntryIfNew] -theorem DistinctKeys.insertEntryIfNew [BEq α] [PartialEquivBEq α] {k : α} {v : β k} {l : List ((a : α) × β a)} (h: DistinctKeys l): +theorem DistinctKeys.insertEntryIfNew [BEq α] [PartialEquivBEq α] {k : α} {v : β k} + {l : List ((a : α) × β a)} (h: DistinctKeys l): DistinctKeys (insertEntryIfNew k v l) := by simp only [Std.DHashMap.Internal.List.insertEntryIfNew, cond_eq_if] split @@ -1910,7 +1909,8 @@ def insertList [BEq α] (l toInsert : List ((a : α) × β a)) : List ((a : α) | .nil => l | .cons ⟨k, v⟩ toInsert => insertList (insertEntry k v l) toInsert -theorem DistinctKeys.insertList [BEq α] [PartialEquivBEq α] {l₁ l₂ : List ((a : α) × β a)} (h : DistinctKeys l₁) : +theorem DistinctKeys.insertList [BEq α] [PartialEquivBEq α] {l₁ l₂ : List ((a : α) × β a)} + (h : DistinctKeys l₁) : DistinctKeys (insertList l₁ l₂) := by induction l₂ using assoc_induction generalizing l₁ · simpa [insertList] @@ -1918,8 +1918,9 @@ theorem DistinctKeys.insertList [BEq α] [PartialEquivBEq α] {l₁ l₂ : List rw [insertList.eq_def] exact ih h.insertEntry -theorem insertList_perm_of_perm_first [BEq α] [EquivBEq α] {l1 l2 toInsert : List ((a : α) × β a)} (h : Perm l1 l2) - (distinct : DistinctKeys l1) : Perm (insertList l1 toInsert) (insertList l2 toInsert) := by +theorem insertList_perm_of_perm_first [BEq α] [EquivBEq α] {l1 l2 toInsert : List ((a : α) × β a)} + (h : Perm l1 l2) (distinct : DistinctKeys l1) : + Perm (insertList l1 toInsert) (insertList l2 toInsert) := by induction toInsert generalizing l1 l2 with | nil => simp [insertList, h] | cons hd tl ih => @@ -1930,18 +1931,18 @@ theorem insertList_perm_of_perm_first [BEq α] [EquivBEq α] {l1 l2 toInsert : L exact h apply DistinctKeys.insertEntry distinct -theorem insertList_cons_perm [BEq α] [EquivBEq α] (l₁ l₂ : List ((a : α) × β a)) - (p : (a : α) × β a) (hl₁ : DistinctKeys l₁) (hl₂ : DistinctKeys (p :: l₂)) : +theorem insertList_cons_perm [BEq α] [EquivBEq α] {l₁ l₂ : List ((a : α) × β a)} + {p : (a : α) × β a} (hl₁ : DistinctKeys l₁) (hl₂ : DistinctKeys (p :: l₂)) : (insertList l₁ (p :: l₂)).Perm (insertEntry p.1 p.2 (insertList l₁ l₂)) := by induction l₂ generalizing p l₁ · simp [insertList] · rename_i h t ih rw [insertList] - refine (ih _ _ hl₁.insertEntry hl₂.tail).trans + refine (ih hl₁.insertEntry hl₂.tail).trans ((insertEntry_of_perm hl₁.insertList - (ih _ _ hl₁ (distinctKeys_of_sublist (by simp) hl₂))).trans + (ih hl₁ (distinctKeys_of_sublist (by simp) hl₂))).trans (List.Perm.trans ?_ (insertEntry_of_perm hl₁.insertList.insertEntry - (ih _ _ hl₁ (distinctKeys_of_sublist (by simp) hl₂)).symm))) + (ih hl₁ (distinctKeys_of_sublist (by simp) hl₂)).symm))) apply getEntry?_ext hl₁.insertList.insertEntry.insertEntry hl₁.insertList.insertEntry.insertEntry (fun k => ?_) simp only [getEntry?_insertEntry] @@ -1950,8 +1951,8 @@ theorem insertList_cons_perm [BEq α] [EquivBEq α] (l₁ l₂ : List ((a : α) have := List.rel_of_pairwise_cons hl₂ (List.mem_cons_self _ _) simp [BEq.trans hh (BEq.symm hp)] at this -theorem containsKey_eq_contains_map_fst [BEq α] [PartialEquivBEq α] (l : List ((a : α) × β a)) (k : α) : containsKey k l = -(l.map Sigma.fst).contains k := by +theorem containsKey_eq_contains_map_fst [BEq α] [PartialEquivBEq α] {l : List ((a : α) × β a)} + {k : α} : containsKey k l = (l.map Sigma.fst).contains k := by induction l with | nil => simp | cons hd tl ih => @@ -1959,8 +1960,9 @@ theorem containsKey_eq_contains_map_fst [BEq α] [PartialEquivBEq α] (l : List simp only [List.map_cons, List.contains_cons] rw [BEq.comm] -theorem containsKey_insertList [BEq α] [PartialEquivBEq α] (l toInsert : List ((a : α) × β a)) (k : α) : - containsKey k (List.insertList l toInsert) = (containsKey k l || (toInsert.map Sigma.fst).contains k) := by +theorem containsKey_insertList [BEq α] [PartialEquivBEq α] {l toInsert : List ((a : α) × β a)} + {k : α} : containsKey k (List.insertList l toInsert) = + (containsKey k l || (toInsert.map Sigma.fst).contains k) := by induction toInsert generalizing l with | nil => simp only [insertList, List.map_nil, List.elem_nil, Bool.or_false] | cons hd tl ih => @@ -1977,32 +1979,33 @@ theorem containsKey_of_containsKey_insertList [BEq α] [PartialEquivBEq α] (h₂ : (toInsert.map Sigma.fst).contains k = false) : containsKey k l := by rw [containsKey_insertList, h₂] at h₁; simp at h₁; exact h₁ -theorem getValueCast?_insertList_not_mem [BEq α] [LawfulBEq α] +theorem getValueCast?_insertList_of_contains_eq_false [BEq α] [LawfulBEq α] {l toInsert : List ((a : α) × β a)} {k : α} - {not_mem : (toInsert.map Sigma.fst).contains k = false} : + (not_contains : (toInsert.map Sigma.fst).contains k = false) : getValueCast? k (insertList l toInsert) = getValueCast? k l := by - rw [← containsKey_eq_contains_map_fst] at not_mem + rw [← containsKey_eq_contains_map_fst] at not_contains induction toInsert generalizing l with | nil => simp [insertList] | cons hd tl ih => - rw [containsKey_cons, Bool.or_eq_false_iff] at not_mem + rw [containsKey_cons, Bool.or_eq_false_iff] at not_contains simp only [insertList] - rw [ih (And.right not_mem)] + rw [ih (And.right not_contains)] rw [getValueCast?_insertEntry] - simp [And.left not_mem] + simp [And.left not_contains] -theorem getValueCast?_insertList_mem [BEq α] [LawfulBEq α] +theorem getValueCast?_insertList_of_mem [BEq α] [LawfulBEq α] {l toInsert : List ((a : α) × β a)} - {k k' : α} {k_beq : k == k'} {v : β k} - {distinct : toInsert.Pairwise (fun a b => (a.1 == b.1) = false)} - {distinct2 : DistinctKeys l} - {mem : ⟨k,v⟩ ∈ toInsert} : - getValueCast? k' (insertList l toInsert) = some (cast (by congr; exact LawfulBEq.eq_of_beq k_beq) v) := by + {k k' : α} (k_beq : k == k') (v : β k) + (distinct : toInsert.Pairwise (fun a b => (a.1 == b.1) = false)) + (distinct2 : DistinctKeys l) + (mem : ⟨k, v⟩ ∈ toInsert) : + getValueCast? k' (insertList l toInsert) = + some (cast (by congr; exact LawfulBEq.eq_of_beq k_beq) v) := by rw [← DistinctKeys.def] at distinct induction toInsert generalizing l with | nil => simp at mem | cons hd tl ih => - rw [getValueCast?_of_perm distinct2.insertList (insertList_cons_perm l tl hd distinct2 distinct)] + rw [getValueCast?_of_perm distinct2.insertList (insertList_cons_perm distinct2 distinct)] rw [getValueCast?_insertEntry] rw [List.mem_cons] at mem cases mem with @@ -2018,101 +2021,107 @@ theorem getValueCast?_insertList_mem [BEq α] [LawfulBEq α] . apply this; exact mem . simp [k_beq] -theorem getValueCast_insertList_not_mem [BEq α] [LawfulBEq α] +theorem getValueCast_insertList_of_contains_eq_false [BEq α] [LawfulBEq α] {l toInsert : List ((a : α) × β a)} {k : α} - {not_mem : (toInsert.map Sigma.fst).contains k = false} + (not_contains : (toInsert.map Sigma.fst).contains k = false) {h} : - getValueCast k (insertList l toInsert) h = getValueCast k l (containsKey_of_containsKey_insertList h not_mem) := by + getValueCast k (insertList l toInsert) h = + getValueCast k l (containsKey_of_containsKey_insertList h not_contains) := by rw [← Option.some_inj] rw [← getValueCast?_eq_some_getValueCast] rw [← getValueCast?_eq_some_getValueCast] - rw [getValueCast?_insertList_not_mem] - exact not_mem + rw [getValueCast?_insertList_of_contains_eq_false] + exact not_contains -theorem getValueCast_insertList_mem [BEq α] [LawfulBEq α] +theorem getValueCast_insertList_of_mem [BEq α] [LawfulBEq α] {l toInsert : List ((a : α) × β a)} - {k k' : α} {k_beq : k == k'} {v : β k} - {distinct : toInsert.Pairwise (fun a b => (a.1 == b.1) = false)} - {distinct2 : DistinctKeys l} - {mem : ⟨k,v⟩ ∈ toInsert} + {k k' : α} (k_beq : k == k') (v : β k) + (distinct : toInsert.Pairwise (fun a b => (a.1 == b.1) = false)) + (distinct2 : DistinctKeys l) + (mem : ⟨k, v⟩ ∈ toInsert) {h} : - getValueCast k' (insertList l toInsert) h = cast (by congr; exact LawfulBEq.eq_of_beq k_beq) v := by + getValueCast k' (insertList l toInsert) h = + cast (by congr; exact LawfulBEq.eq_of_beq k_beq) v := by rw [← Option.some_inj] rw [← getValueCast?_eq_some_getValueCast] - rw [getValueCast?_insertList_mem] + rw [getValueCast?_insertList_of_mem] . exact k_beq . exact distinct . exact distinct2 . exact mem -theorem getValueCast!_insertList_not_mem [BEq α] [LawfulBEq α] +theorem getValueCast!_insertList_of_contains_eq_false [BEq α] [LawfulBEq α] {l toInsert : List ((a : α) × β a)} {k : α} [Inhabited (β k)] - {not_mem : (toInsert.map Sigma.fst).contains k = false} : + (not_contains : (toInsert.map Sigma.fst).contains k = false) : getValueCast! k (insertList l toInsert) = getValueCast! k l := by - rw [getValueCast!_eq_getValueCast?, getValueCast!_eq_getValueCast?, getValueCast?_insertList_not_mem] - exact not_mem + rw [getValueCast!_eq_getValueCast?, getValueCast!_eq_getValueCast?, + getValueCast?_insertList_of_contains_eq_false] + exact not_contains -theorem getValueCast!_insertList_mem [BEq α] [LawfulBEq α] +theorem getValueCast!_insertList_of_mem [BEq α] [LawfulBEq α] {l toInsert : List ((a : α) × β a)} - {k k' : α} {k_beq : k == k'} {v : β k} [Inhabited (β k')] - {distinct : toInsert.Pairwise (fun a b => (a.1 == b.1) = false)} - {distinct2 : DistinctKeys l} - {mem : ⟨k,v⟩ ∈ toInsert} : - getValueCast! k' (insertList l toInsert) = cast (by congr; exact LawfulBEq.eq_of_beq k_beq) v := by - rw [getValueCast!_eq_getValueCast?, getValueCast?_insertList_mem] + {k k' : α} (k_beq : k == k') (v : β k) [Inhabited (β k')] + (distinct : toInsert.Pairwise (fun a b => (a.1 == b.1) = false)) + (distinct2 : DistinctKeys l) + (mem : ⟨k, v⟩ ∈ toInsert) : + getValueCast! k' (insertList l toInsert) = + cast (by congr; exact LawfulBEq.eq_of_beq k_beq) v := by + rw [getValueCast!_eq_getValueCast?, getValueCast?_insertList_of_mem] . rw [Option.get!_some]; exact k_beq . exact distinct . exact distinct2 . exact mem -theorem getValueCastD_insertList_not_mem [BEq α] [LawfulBEq α] +theorem getValueCastD_insertList_of_contains_eq_false [BEq α] [LawfulBEq α] {l toInsert : List ((a : α) × β a)} {k : α} {fallback : β k} - {not_mem : (toInsert.map Sigma.fst).contains k = false} : + (not_mem : (toInsert.map Sigma.fst).contains k = false) : getValueCastD k (insertList l toInsert) fallback = getValueCastD k l fallback := by - rw [getValueCastD_eq_getValueCast?, getValueCastD_eq_getValueCast?, getValueCast?_insertList_not_mem] + rw [getValueCastD_eq_getValueCast?, getValueCastD_eq_getValueCast?, + getValueCast?_insertList_of_contains_eq_false] exact not_mem -theorem getValueCastD_insertList_mem [BEq α] [LawfulBEq α] +theorem getValueCastD_insertList_of_mem [BEq α] [LawfulBEq α] {l toInsert : List ((a : α) × β a)} - {k k' : α} {k_beq : k == k'} {v : β k} {fallback : β k'} - {distinct : toInsert.Pairwise (fun a b => (a.1 == b.1) = false)} - {distinct2 : DistinctKeys l} - {mem : ⟨k,v⟩ ∈ toInsert} : - getValueCastD k' (insertList l toInsert) fallback = cast (by congr; exact LawfulBEq.eq_of_beq k_beq) v := by - rw [getValueCastD_eq_getValueCast?, getValueCast?_insertList_mem] + {k k' : α} (k_beq : k == k') {v : β k} {fallback : β k'} + (distinct : toInsert.Pairwise (fun a b => (a.1 == b.1) = false)) + (distinct2 : DistinctKeys l) + (mem : ⟨k, v⟩ ∈ toInsert) : + getValueCastD k' (insertList l toInsert) fallback = + cast (by congr; exact LawfulBEq.eq_of_beq k_beq) v := by + rw [getValueCastD_eq_getValueCast?, getValueCast?_insertList_of_mem] . rw [Option.getD_some]; exact k_beq . exact distinct . exact distinct2 . exact mem -theorem getKey?_insertList_not_mem [BEq α] [EquivBEq α] +theorem getKey?_insertList_of_contains_eq_false [BEq α] [EquivBEq α] {l toInsert : List ((a : α) × β a)} {k : α} - {not_mem : (toInsert.map Sigma.fst).contains k = false} : + (not_contains : (toInsert.map Sigma.fst).contains k = false) : getKey? k (insertList l toInsert) = getKey? k l := by - rw [← containsKey_eq_contains_map_fst] at not_mem + rw [← containsKey_eq_contains_map_fst] at not_contains induction toInsert generalizing l with | nil => simp [insertList] | cons hd tl ih => simp only [insertList] rw [ih] · rw [getKey?_insertEntry] - simp only [containsKey, Bool.or_eq_false_iff] at not_mem - simp [And.left not_mem] - · simp only [containsKey, Bool.or_eq_false_iff] at not_mem - apply And.right not_mem + simp only [containsKey, Bool.or_eq_false_iff] at not_contains + simp [And.left not_contains] + · simp only [containsKey, Bool.or_eq_false_iff] at not_contains + apply And.right not_contains -theorem getKey?_insertList_mem [BEq α] [EquivBEq α] +theorem getKey?_insertList_of_mem [BEq α] [EquivBEq α] {l toInsert : List ((a : α) × β a)} - {k k' : α} {k_beq : k == k'} - {distinct : toInsert.Pairwise (fun a b => (a.1 == b.1) = false)} - {distinct2 : DistinctKeys l} - {mem : k ∈ (toInsert.map Sigma.fst)} : + {k k' : α} (k_beq : k == k') + (distinct : toInsert.Pairwise (fun a b => (a.1 == b.1) = false)) + (distinct2 : DistinctKeys l) + (mem : k ∈ (toInsert.map Sigma.fst)) : getKey? k' (insertList l toInsert) = some k := by rw [← DistinctKeys.def] at distinct induction toInsert generalizing l with | nil => simp at mem | cons hd tl ih => - rw [getKey?_of_perm distinct2.insertList (insertList_cons_perm l tl hd distinct2 distinct)] + rw [getKey?_of_perm distinct2.insertList (insertList_cons_perm distinct2 distinct)] rw [getKey?_insertEntry] rw [List.map_cons, List.mem_cons] at mem cases mem with @@ -2130,84 +2139,84 @@ theorem getKey?_insertList_mem [BEq α] [EquivBEq α] . apply this; exact pair_mem . rw [pair_eq]; exact k_beq -theorem getKey_insertList_not_mem [BEq α] [EquivBEq α] +theorem getKey_insertList_of_contains_eq_false [BEq α] [EquivBEq α] {l toInsert : List ((a : α) × β a)} {k : α} - {not_mem : (toInsert.map Sigma.fst).contains k = false} + (not_contains : (toInsert.map Sigma.fst).contains k = false) {h} : getKey k (insertList l toInsert) h = - getKey k l (containsKey_of_containsKey_insertList h not_mem) := by + getKey k l (containsKey_of_containsKey_insertList h not_contains) := by rw [← Option.some_inj] rw [← getKey?_eq_some_getKey] - rw [getKey?_insertList_not_mem] + rw [getKey?_insertList_of_contains_eq_false] . rw [getKey?_eq_some_getKey] - . exact not_mem + . exact not_contains -theorem getKey_insertList_mem [BEq α] [EquivBEq α] +theorem getKey_insertList_of_mem [BEq α] [EquivBEq α] {l toInsert : List ((a : α) × β a)} - {k k' : α} {k_beq : k == k'} - {distinct : toInsert.Pairwise (fun a b => (a.1 == b.1) = false)} - {distinct2 : DistinctKeys l} - {mem : k ∈ (toInsert.map Sigma.fst)} + {k k' : α} (k_beq : k == k') + (distinct : toInsert.Pairwise (fun a b => (a.1 == b.1) = false)) + (distinct2 : DistinctKeys l) + (mem : k ∈ (toInsert.map Sigma.fst)) {h} : getKey k' (insertList l toInsert) h = k := by rw [← Option.some_inj] rw [← getKey?_eq_some_getKey] - rw [getKey?_insertList_mem] + rw [getKey?_insertList_of_mem] . exact k_beq . exact distinct . exact distinct2 . exact mem -theorem getKey!_insertList_not_mem [BEq α] [EquivBEq α] [Inhabited α] +theorem getKey!_insertList_of_contains_eq_false [BEq α] [EquivBEq α] [Inhabited α] {l toInsert : List ((a : α) × β a)} {k : α} - {not_mem : (toInsert.map Sigma.fst).contains k = false} : + (contains_false : (toInsert.map Sigma.fst).contains k = false) : getKey! k (insertList l toInsert) = getKey! k l := by rw [getKey!_eq_getKey?] - rw [getKey?_insertList_not_mem] + rw [getKey?_insertList_of_contains_eq_false] . rw [getKey!_eq_getKey?] - . exact not_mem + . exact contains_false -theorem getKey!_insertList_mem [BEq α] [EquivBEq α] [Inhabited α] +theorem getKey!_insertList_of_mem [BEq α] [EquivBEq α] [Inhabited α] {l toInsert : List ((a : α) × β a)} - {k k' : α} {k_beq : k == k'} - {distinct : toInsert.Pairwise (fun a b => (a.1 == b.1) = false)} - {distinct2 : DistinctKeys l} - {mem : k ∈ (toInsert.map Sigma.fst)} : + {k k' : α} (k_beq : k == k') + (distinct : toInsert.Pairwise (fun a b => (a.1 == b.1) = false)) + (distinct2 : DistinctKeys l) + (mem : k ∈ (toInsert.map Sigma.fst)) : getKey! k' (insertList l toInsert) = k := by rw [getKey!_eq_getKey?] - rw [getKey?_insertList_mem] + rw [getKey?_insertList_of_mem] . rw [Option.get!_some] . exact k_beq . exact distinct . exact distinct2 . exact mem -theorem getKeyD_insertList_not_mem [BEq α] [EquivBEq α] +theorem getKeyD_insertList_of_contains_eq_false [BEq α] [EquivBEq α] {l toInsert : List ((a : α) × β a)} {k fallback : α} - {not_mem : (toInsert.map Sigma.fst).contains k = false} : + (h : (toInsert.map Sigma.fst).contains k = false) : getKeyD k (insertList l toInsert) fallback = getKeyD k l fallback := by rw [getKeyD_eq_getKey?] - rw [getKey?_insertList_not_mem] + rw [getKey?_insertList_of_contains_eq_false] . rw [getKeyD_eq_getKey?] - . exact not_mem + . exact h -theorem getKeyD_insertList_mem [BEq α] [EquivBEq α] +theorem getKeyD_insertList_of_mem [BEq α] [EquivBEq α] {l toInsert : List ((a : α) × β a)} - {k k' fallback : α} {k_beq : k == k'} - {distinct : toInsert.Pairwise (fun a b => (a.1 == b.1) = false)} - {distinct2 : DistinctKeys l} - {mem : k ∈ (toInsert.map Sigma.fst)} : + {k k' fallback : α} (k_beq : k == k') + (distinct : toInsert.Pairwise (fun a b => (a.1 == b.1) = false)) + (distinct2 : DistinctKeys l) + (mem : k ∈ (toInsert.map Sigma.fst)) : getKeyD k' (insertList l toInsert) fallback = k := by rw [getKeyD_eq_getKey?] - rw [getKey?_insertList_mem] + rw [getKey?_insertList_of_mem] . rw [Option.getD_some] . exact k_beq . exact distinct . exact distinct2 . exact mem -theorem insertList_perm [BEq α] [ReflBEq α] [PartialEquivBEq α] - (l toInsert : List ((a : α) × β a)) +theorem perm_insertList [BEq α] [ReflBEq α] [PartialEquivBEq α] + {l toInsert : List ((a : α) × β a)} (distinct_l : DistinctKeys l) (distinct_toInsert : List.Pairwise (fun a b => (a.1 == b.1) = false) toInsert) (distinct_both : ∀ (a : α), ¬ (containsKey a l ∧ containsKey a toInsert)) : @@ -2217,7 +2226,6 @@ theorem insertList_perm [BEq α] [ReflBEq α] [PartialEquivBEq α] | nil => simp only [insertList, List.append_nil, Perm.refl] | cons hd tl ih => simp only [insertList] - specialize ih (insertEntry hd.fst hd.snd l) apply Perm.trans · apply ih · simp only [insertEntry] @@ -2252,9 +2260,9 @@ theorem insertList_perm [BEq α] [ReflBEq α] [PartialEquivBEq α] apply Perm.symm apply List.perm_middle -theorem insertList_length [BEq α] [EquivBEq α] +theorem length_insertList [BEq α] [EquivBEq α] {l toInsert : List ((a : α) × β a)} - {distinct_l : DistinctKeys l} + (distinct_l : DistinctKeys l) (distinct_toInsert : List.Pairwise (fun a b => (a.1 == b.1) = false) toInsert) (distinct_both : ∀ (a : α), ¬ (containsKey a l ∧ (toInsert.map Sigma.fst).contains a)) : (insertList l toInsert).length = l.length + toInsert.length := by @@ -2272,6 +2280,7 @@ theorem insertList_length [BEq α] [EquivBEq α] specialize distinct_both h simp [containsKey] at distinct_both · rw [Nat.add_assoc, Nat.add_comm 1 _] + · apply DistinctKeys.insertEntry distinct_l · apply DistinctKeys.tail distinct_toInsert · intro a specialize distinct_both a @@ -2296,19 +2305,18 @@ theorem insertList_length [BEq α] [EquivBEq α] simp only [Bool.true_eq_false] at left | inr h => apply And.right (distinct_both h) - · apply DistinctKeys.insertEntry distinct_l -theorem insertList_not_isEmpty_if_start_not_isEmpty [BEq α] +theorem isEmpty_insertList_eq_false_of_isEmpty_eq_false [BEq α] {l toInsert : List ((a : α) × β a)} - {h : l.isEmpty = false} : + (h : l.isEmpty = false) : (List.insertList l toInsert).isEmpty = false := by induction toInsert generalizing l with | nil => simp [insertList, h] | cons hd tl ih => simp [insertList, ih] -theorem insertList_isEmpty [BEq α] +theorem isEmpty_insertList [BEq α] {l toInsert : List ((a : α) × β a)} : (List.insertList l toInsert).isEmpty ↔ l.isEmpty ∧ toInsert.isEmpty := by induction toInsert with @@ -2317,15 +2325,15 @@ theorem insertList_isEmpty [BEq α] simp only [insertList, List.isEmpty_cons, Bool.false_eq_true, and_false, iff_false] apply ne_true_of_eq_false - apply insertList_not_isEmpty_if_start_not_isEmpty + apply isEmpty_insertList_eq_false_of_isEmpty_eq_false simp section variable {β: Type v} -theorem getValue?_insertList_not_mem [BEq α] [PartialEquivBEq α] +theorem getValue?_insertList_of_contains_eq_false [BEq α] [PartialEquivBEq α] {l toInsert : List ((_ : α) × β)} {k : α} - {not_mem : (toInsert.map Sigma.fst).contains k = false} : + (h : (toInsert.map Sigma.fst).contains k = false) : getValue? k (insertList l toInsert) = getValue? k l := by induction toInsert generalizing l with | nil => simp [insertList] @@ -2333,17 +2341,17 @@ theorem getValue?_insertList_not_mem [BEq α] [PartialEquivBEq α] simp only [insertList] rw [ih] · apply getValue?_insertEntry_of_false - simp only [List.map_cons, List.contains_cons, Bool.or_eq_false_iff] at not_mem + simp only [List.map_cons, List.contains_cons, Bool.or_eq_false_iff] at h apply BEq.symm_false - apply (And.left not_mem) - · simp only [List.map_cons, List.contains_cons, Bool.or_eq_false_iff] at not_mem - apply And.right not_mem + apply (And.left h) + · simp only [List.map_cons, List.contains_cons, Bool.or_eq_false_iff] at h + apply And.right h -theorem getValue?_insertList_mem [BEq α] [PartialEquivBEq α] +theorem getValue?_insertList_of_mem [BEq α] [PartialEquivBEq α] {l toInsert : List ((_ : α) × β )} - {k k' : α} {k_beq : k == k'} {v : β} - {distinct : List.Pairwise (fun a b => (a.1 == b.1) = false) toInsert} - {mem : ⟨k,v⟩ ∈ toInsert} : + {k k' : α} (k_beq : k == k') {v : β} + (distinct : List.Pairwise (fun a b => (a.1 == b.1) = false) toInsert) + (mem : ⟨k, v⟩ ∈ toInsert) : getValue? k' (insertList l toInsert) = v := by induction toInsert generalizing l with | nil => simp at mem @@ -2352,7 +2360,7 @@ theorem getValue?_insertList_mem [BEq α] [PartialEquivBEq α] simp only [List.mem_cons] at mem cases mem with | inl mem => - rw [getValue?_insertList_not_mem] + rw [getValue?_insertList_of_contains_eq_false] simp only [← mem] · rw [getValue?_insertEntry_of_beq] exact k_beq @@ -2376,61 +2384,64 @@ theorem getValue?_insertList_mem [BEq α] [PartialEquivBEq α] apply And.right distinct · exact mem -theorem getValue!_insertList_not_mem [BEq α] [PartialEquivBEq α] [Inhabited β] +theorem getValue!_insertList_of_contains_eq_false [BEq α] [PartialEquivBEq α] [Inhabited β] {l toInsert : List ((_ : α) × β)} {k : α} - {not_mem : (toInsert.map Sigma.fst).contains k = false} : + (h : (toInsert.map Sigma.fst).contains k = false) : getValue! k (insertList l toInsert) = getValue! k l := by simp only [getValue!_eq_getValue?] - rw [getValue?_insertList_not_mem] - exact not_mem + rw [getValue?_insertList_of_contains_eq_false] + exact h -theorem getValue!_insertList_mem [BEq α] [PartialEquivBEq α] [Inhabited β] - {l toInsert : List ((_ : α) × β)} {k k' : α} {v: β} {k_eq : k == k'} {distinct : List.Pairwise (fun a b => (a.1 == b.1) = false) toInsert } {mem : ⟨k,v⟩ ∈ toInsert}: +theorem getValue!_insertList_of_mem [BEq α] [PartialEquivBEq α] [Inhabited β] + {l toInsert : List ((_ : α) × β)} {k k' : α} {v: β} (k_beq : k == k') + (distinct : List.Pairwise (fun a b => (a.1 == b.1) = false) toInsert) (mem : ⟨k, v⟩ ∈ toInsert): getValue! k' (insertList l toInsert) = v := by simp only [getValue!_eq_getValue?] - rw [getValue?_insertList_mem (mem:= mem)] + rw [getValue?_insertList_of_mem (mem:= mem)] · simp - · exact k_eq + · exact k_beq · exact distinct -theorem getValueD_insertList_not_mem [BEq α] [PartialEquivBEq α] +theorem getValueD_insertList_of_contains_eq_false [BEq α] [PartialEquivBEq α] {l toInsert : List ((_ : α) × β)} {k : α} {fallback : β} - {not_mem : (toInsert.map Sigma.fst).contains k = false} : + (h : (toInsert.map Sigma.fst).contains k = false) : getValueD k (insertList l toInsert) fallback = getValueD k l fallback := by simp only [getValueD_eq_getValue?] - rw [getValue?_insertList_not_mem] - exact not_mem + rw [getValue?_insertList_of_contains_eq_false] + exact h -theorem getValueD_insertList_mem [BEq α] [PartialEquivBEq α] - {l toInsert : List ((_ : α) × β)} {k k' : α} {v fallback: β} {k_eq : k == k'} {distinct : List.Pairwise (fun a b => (a.1 == b.1) = false) toInsert } {mem : ⟨k,v⟩ ∈ toInsert}: +theorem getValueD_insertList_of_mem [BEq α] [PartialEquivBEq α] + {l toInsert : List ((_ : α) × β)} {k k' : α} {v fallback: β} (k_beq : k == k') + (distinct : List.Pairwise (fun a b => (a.1 == b.1) = false) toInsert) (mem : ⟨k, v⟩ ∈ toInsert): getValueD k' (insertList l toInsert) fallback= v := by simp only [getValueD_eq_getValue?] - rw [getValue?_insertList_mem (mem:= mem)] + rw [getValue?_insertList_of_mem (mem:= mem)] · simp - · exact k_eq + · exact k_beq · exact distinct -theorem getValue_insertList_not_mem [BEq α] [PartialEquivBEq α] +theorem getValue_insertList_of_contains_eq_false [BEq α] [PartialEquivBEq α] {l toInsert : List ((_ : α) × β)} {k : α} - {not_mem : (toInsert.map Sigma.fst).contains k = false} + (h' : (toInsert.map Sigma.fst).contains k = false) {h} : - getValue k (insertList l toInsert) h = getValue k l (containsKey_of_containsKey_insertList h not_mem) := by + getValue k (insertList l toInsert) h = + getValue k l (containsKey_of_containsKey_insertList h h') := by rw [← Option.some_inj] rw [← getValue?_eq_some_getValue] rw [← getValue?_eq_some_getValue] - rw [getValue?_insertList_not_mem] - exact not_mem + rw [getValue?_insertList_of_contains_eq_false] + exact h' -theorem getValue_insertList_mem [BEq α] [PartialEquivBEq α] +theorem getValue_insertList_of_mem [BEq α] [PartialEquivBEq α] {l toInsert : List ((_ : α) × β)} - {k k' : α} {k_beq : k == k'} {v : β} - {distinct : toInsert.Pairwise (fun a b => (a.1 == b.1) = false)} - {mem : ⟨k,v⟩ ∈ toInsert} + {k k' : α} (k_beq : k == k') {v : β} + (distinct : toInsert.Pairwise (fun a b => (a.1 == b.1) = false)) + (mem : ⟨k, v⟩ ∈ toInsert) {h} : getValue k' (insertList l toInsert) h = v := by rw [← Option.some_inj] rw [← getValue?_eq_some_getValue] - rw [getValue?_insertList_mem (mem:=mem)] + rw [getValue?_insertList_of_mem (mem:=mem)] . exact k_beq . exact distinct @@ -2448,30 +2459,38 @@ theorem insertListConst_eq_insertList [BEq α] {l : List ((_ : α) × β)} {toIn simp [insertList, insertListConst] apply ih -theorem DistinctKeys.insertListConst [BEq α] [PartialEquivBEq α] {l₁ : List ((_ : α) × β)} {l₂ : List (α × β)} (h : DistinctKeys l₁) : +theorem DistinctKeys.insertListConst [BEq α] [PartialEquivBEq α] {l₁ : List ((_ : α) × β)} + {l₂ : List (α × β)} (h : DistinctKeys l₁) : DistinctKeys (insertListConst l₁ l₂) := by rw [insertListConst_eq_insertList] apply DistinctKeys.insertList h -theorem insertListConst_perm_of_perm_first [BEq α] [EquivBEq α] (l1 l2: List ((_ : α) × β)) (toInsert : List (α × β)) (h : Perm l1 l2) - (distinct : DistinctKeys l1) : Perm (insertListConst l1 toInsert) (insertListConst l2 toInsert) := by +theorem insertListConst_perm_of_perm_first [BEq α] [EquivBEq α] {l1 l2: List ((_ : α) × β)} + {toInsert : List (α × β)} (h : Perm l1 l2) (distinct : DistinctKeys l1) : + Perm (insertListConst l1 toInsert) (insertListConst l2 toInsert) := by simp only [insertListConst_eq_insertList] apply insertList_perm_of_perm_first h distinct -theorem list_conversion_insertListConst {toInsert : List (α × β)}: (List.map Sigma.fst (List.map (fun a => (⟨a.fst, a.snd⟩ : (_ : α ) × β )) toInsert)) = (List.map Prod.fst toInsert) := by +theorem list_conversion_insertListConst {toInsert : List (α × β)}: + (List.map Sigma.fst (List.map (fun a => (⟨a.fst, a.snd⟩ : (_ : α ) × β )) toInsert)) = + (List.map Prod.fst toInsert) := by apply List.ext_get · simp · intro n h1 h2 simp -theorem list_conversion_insertListConst_contains [BEq α] [PartialEquivBEq α] {toInsert : List (α × β)} {a : α}: - containsKey a (List.map (fun a => (⟨a.fst, a.snd⟩ : ((_ : α ) × β) )) toInsert) = true ↔ (toInsert.map Prod.fst).contains a = true := by +theorem list_conversion_insertListConst_containsKey_true_iff [BEq α] [PartialEquivBEq α] + {toInsert : List (α × β)} {a : α}: + containsKey a (List.map (fun a => (⟨a.fst, a.snd⟩ : ((_ : α ) × β) )) toInsert) = true ↔ + (toInsert.map Prod.fst).contains a = true := by rw [List.contains_iff_exists_mem_beq] rw [containsKey_iff_exists] simp only [keys_def, list_conversion_insertListConst] -theorem pairwise_list_conversion_insertListConst [BEq α] {toInsert : List (α × β)} (distinct : toInsert.Pairwise (fun a b => (a.1 == b.1) = false) ): - List.Pairwise (fun a b => (a.1 == b.1) = false) (List.map (fun a => (⟨a.1, a.2⟩ : ((_ : α ) × β)) ) toInsert) := by +theorem pairwise_list_conversion_insertListConst [BEq α] {toInsert : List (α × β)} + (distinct : toInsert.Pairwise (fun a b => (a.1 == b.1) = false) ): + List.Pairwise (fun a b => (a.1 == b.1) = false) + (List.map (fun a => (⟨a.1, a.2⟩ : ((_ : α ) × β)) ) toInsert) := by induction toInsert with | nil => simp | cons hd tl ih => @@ -2488,87 +2507,91 @@ theorem pairwise_list_conversion_insertListConst [BEq α] {toInsert : List (α apply left · apply ih right -theorem containsKey_insertListConst [BEq α] [PartialEquivBEq α] {l : List ((_ : α) × β)} {toInsert : List (α × β)} {k : α} : - containsKey k (insertListConst l toInsert) = (containsKey k l || (toInsert.map Prod.fst).contains k) := by +theorem containsKey_insertListConst [BEq α] [PartialEquivBEq α] + {l : List ((_ : α) × β)} {toInsert : List (α × β)} {k : α} : + containsKey k (insertListConst l toInsert) = + (containsKey k l || (toInsert.map Prod.fst).contains k) := by rw [insertListConst_eq_insertList, containsKey_insertList,list_conversion_insertListConst] -theorem containsKey_of_containsKey_insertListConst [BEq α] [PartialEquivBEq α] {l : List ((_ : α) × β)} {toInsert : List (α × β)} {k : α} (h₁ : containsKey k (insertListConst l toInsert)) +theorem containsKey_of_containsKey_insertListConst [BEq α] [PartialEquivBEq α] + {l : List ((_ : α) × β)} {toInsert : List (α × β)} {k : α} + (h₁ : containsKey k (insertListConst l toInsert)) (h₂ : (toInsert.map Prod.fst).contains k = false) : containsKey k l := by rw [insertListConst_eq_insertList] at h₁ apply containsKey_of_containsKey_insertList h₁ rw [list_conversion_insertListConst] exact h₂ -theorem getKey?_insertListConst_not_mem [BEq α] [EquivBEq α] +theorem getKey?_insertListConst_of_contains_eq_false [BEq α] [EquivBEq α] {l : List ((_ : α) × β)} {toInsert : List (α × β)} {k : α} - {not_mem : (toInsert.map Prod.fst).contains k = false} : + (h : (toInsert.map Prod.fst).contains k = false) : getKey? k (insertListConst l toInsert) = getKey? k l := by rw [insertListConst_eq_insertList] - apply getKey?_insertList_not_mem + apply getKey?_insertList_of_contains_eq_false rw [list_conversion_insertListConst] - exact not_mem + exact h -theorem getKey?_insertListConst_mem [BEq α] [EquivBEq α] +theorem getKey?_insertListConst_of_mem [BEq α] [EquivBEq α] {l : List ((_ : α) × β)} {toInsert : List (α × β)} - {k k' : α} {k_beq : k == k'} - {distinct : toInsert.Pairwise (fun a b => (a.1 == b.1) = false)} - {distinct2 : DistinctKeys l} - {mem : k ∈ (toInsert.map Prod.fst)} : + {k k' : α} (k_beq : k == k') + (distinct : toInsert.Pairwise (fun a b => (a.1 == b.1) = false)) + (distinct2 : DistinctKeys l) + (mem : k ∈ (toInsert.map Prod.fst)) : getKey? k' (insertListConst l toInsert) = some k := by rw [insertListConst_eq_insertList] - apply getKey?_insertList_mem + apply getKey?_insertList_of_mem · exact k_beq · apply pairwise_list_conversion_insertListConst distinct · exact distinct2 · rw [list_conversion_insertListConst] assumption -theorem getKey_insertListConst_not_mem [BEq α] [EquivBEq α] +theorem getKey_insertListConst_of_contains_eq_false [BEq α] [EquivBEq α] {l : List ((_ : α) × β)} {toInsert : List (α × β)} {k : α} - {not_mem : (toInsert.map Prod.fst).contains k = false} + (h' : (toInsert.map Prod.fst).contains k = false) {h} : getKey k (insertListConst l toInsert) h = - getKey k l (containsKey_of_containsKey_insertListConst h not_mem) := by + getKey k l (containsKey_of_containsKey_insertListConst h h') := by rw [← Option.some_inj] rw [← getKey?_eq_some_getKey] - rw [getKey?_insertListConst_not_mem] + rw [getKey?_insertListConst_of_contains_eq_false] . rw [getKey?_eq_some_getKey] - . exact not_mem + . exact h' -theorem getKey_insertListConst_mem [BEq α] [EquivBEq α] +theorem getKey_insertListConst_of_mem [BEq α] [EquivBEq α] {l : List ((_ : α) × β)} {toInsert : List (α × β)} - {k k' : α} {k_beq : k == k'} - {distinct : toInsert.Pairwise (fun a b => (a.1 == b.1) = false)} - {distinct2 : DistinctKeys l} - {mem : k ∈ (toInsert.map Prod.fst)} + {k k' : α} (k_beq : k == k') + (distinct : toInsert.Pairwise (fun a b => (a.1 == b.1) = false)) + (distinct2 : DistinctKeys l) + (mem : k ∈ (toInsert.map Prod.fst)) {h} : getKey k' (insertListConst l toInsert) h = k := by rw [← Option.some_inj] rw [← getKey?_eq_some_getKey] - rw [getKey?_insertListConst_mem] + rw [getKey?_insertListConst_of_mem] . exact k_beq . exact distinct . exact distinct2 . exact mem -theorem getKey!_insertListConst_not_mem [BEq α] [EquivBEq α] [Inhabited α] +theorem getKey!_insertListConst_of_contains_eq_false [BEq α] [EquivBEq α] [Inhabited α] {l : List ((_ : α) × β)} {toInsert : List (α × β)} {k : α} - {not_mem : (toInsert.map Prod.fst).contains k = false} : + (h : (toInsert.map Prod.fst).contains k = false) : getKey! k (insertListConst l toInsert) = getKey! k l := by rw [insertListConst_eq_insertList] - apply getKey!_insertList_not_mem + apply getKey!_insertList_of_contains_eq_false rw [list_conversion_insertListConst] - exact not_mem + exact h -theorem getKey!_insertListConst_mem [BEq α] [EquivBEq α] [Inhabited α] +theorem getKey!_insertListConst_of_mem [BEq α] [EquivBEq α] [Inhabited α] {l : List ((_ : α) × β)} {toInsert : List (α × β)} - {k k' : α} {k_beq : k == k'} - {distinct : toInsert.Pairwise (fun a b => (a.1 == b.1) = false)} - {distinct2 : DistinctKeys l} - {mem : k ∈ (toInsert.map Prod.fst)} : + {k k' : α} (k_beq : k == k') + (distinct : toInsert.Pairwise (fun a b => (a.1 == b.1) = false)) + (distinct2 : DistinctKeys l) + (mem : k ∈ (toInsert.map Prod.fst)) : getKey! k' (insertListConst l toInsert) = k := by rw [insertListConst_eq_insertList] - apply getKey!_insertList_mem + apply getKey!_insertList_of_mem · exact k_beq · apply pairwise_list_conversion_insertListConst exact distinct @@ -2576,24 +2599,24 @@ theorem getKey!_insertListConst_mem [BEq α] [EquivBEq α] [Inhabited α] · rw [list_conversion_insertListConst] exact mem -theorem getKeyD_insertListConst_not_mem [BEq α] [EquivBEq α] +theorem getKeyD_insertListConst_of_contains_eq_false [BEq α] [EquivBEq α] {l : List ((_ : α) × β)} {toInsert : List (α × β)} {k fallback : α} - {not_mem : (toInsert.map Prod.fst).contains k = false} : + (h : (toInsert.map Prod.fst).contains k = false) : getKeyD k (insertListConst l toInsert) fallback = getKeyD k l fallback := by rw [insertListConst_eq_insertList] - apply getKeyD_insertList_not_mem + apply getKeyD_insertList_of_contains_eq_false rw [list_conversion_insertListConst] - exact not_mem + exact h -theorem getKeyD_insertListConst_mem [BEq α] [EquivBEq α] +theorem getKeyD_insertListConst_of_mem [BEq α] [EquivBEq α] {l : List ((_ : α) × β)} {toInsert : List (α × β)} - {k k' fallback : α} {k_beq : k == k'} - {distinct : toInsert.Pairwise (fun a b => (a.1 == b.1) = false)} - {distinct2 : DistinctKeys l} - {mem : k ∈ (toInsert.map Prod.fst)} : + {k k' fallback : α} (k_beq : k == k') + (distinct : toInsert.Pairwise (fun a b => (a.1 == b.1) = false)) + (distinct2 : DistinctKeys l) + (mem : k ∈ (toInsert.map Prod.fst)) : getKeyD k' (insertListConst l toInsert) fallback = k := by rw [insertListConst_eq_insertList] - apply getKeyD_insertList_mem + apply getKeyD_insertList_of_mem · exact k_beq · apply pairwise_list_conversion_insertListConst exact distinct @@ -2601,15 +2624,16 @@ theorem getKeyD_insertListConst_mem [BEq α] [EquivBEq α] · rw [list_conversion_insertListConst] exact mem -theorem insertListConst_length [BEq α] [EquivBEq α] +theorem length_insertListConst [BEq α] [EquivBEq α] {l : List ((_ : α) × β)} {toInsert : List (α × β)} - {distinct_l : DistinctKeys l} + (distinct_l : DistinctKeys l) (distinct_toInsert : List.Pairwise (fun a b => (a.1 == b.1) = false) toInsert) (distinct_both : ∀ (a : α), ¬ (containsKey a l ∧ (toInsert.map Prod.fst).contains a)) : (insertListConst l toInsert).length = l.length + toInsert.length := by rw [insertListConst_eq_insertList] - rw [insertList_length] + rw [length_insertList] · simp + · exact distinct_l · apply pairwise_list_conversion_insertListConst distinct_toInsert · intro a simp only [not_and, Bool.not_eq_true] @@ -2624,113 +2648,117 @@ theorem insertListConst_length [BEq α] [EquivBEq α] simp only [List.any_map, List.any_eq_false, Function.comp_apply, Bool.not_eq_true, Prod.forall] at distinct_both apply distinct_both - · exact distinct_l -theorem insertListConst_not_isEmpty_if_start_not_isEmpty [BEq α] +theorem isEmpty_insertListConst_eq_false_of_isEmpty_eq_false [BEq α] {l : List ((_ : α) × β)} {toInsert : List (α × β)} - {h : l.isEmpty = false} : + (h : l.isEmpty = false) : (List.insertListConst l toInsert).isEmpty = false := by induction toInsert generalizing l with | nil => simp [insertListConst, h] | cons hd tl ih => simp [insertListConst, ih] -theorem insertListConst_isEmpty [BEq α] +theorem isEmpty_insertListConst [BEq α] {l : List ((_ : α) × β)} {toInsert : List (α × β)} : (List.insertListConst l toInsert).isEmpty ↔ l.isEmpty ∧ toInsert.isEmpty := by rw [insertListConst_eq_insertList] - rw [insertList_isEmpty] + rw [isEmpty_insertList] simp only [List.isEmpty_eq_true, List.map_eq_nil_iff] -theorem getValue?_insertListConst_not_mem [BEq α] [PartialEquivBEq α] +theorem getValue?_insertListConst_of_contains_eq_false [BEq α] [PartialEquivBEq α] {l : List ((_ : α) × β)} {toInsert : List (α × β)} {k : α} - {not_mem : (toInsert.map Prod.fst).contains k = false} : + (h : (toInsert.map Prod.fst).contains k = false) : getValue? k (insertListConst l toInsert) = getValue? k l := by rw [insertListConst_eq_insertList] - rw [getValue?_insertList_not_mem] + rw [getValue?_insertList_of_contains_eq_false] rw [list_conversion_insertListConst] - exact not_mem + exact h -theorem getValue?_insertListConst_mem [BEq α] [PartialEquivBEq α] +theorem getValue?_insertListConst_of_mem [BEq α] [PartialEquivBEq α] {l : List ((_ : α) × β)} {toInsert : List (α × β)} - {k k' : α} {k_beq : k == k'} {v : β} - {distinct : List.Pairwise (fun a b => (a.1 == b.1) = false) toInsert} - {mem : ⟨k,v⟩ ∈ toInsert} : + {k k' : α} (k_beq : k == k') {v : β} + (distinct : List.Pairwise (fun a b => (a.1 == b.1) = false) toInsert) + (mem : ⟨k, v⟩ ∈ toInsert) : getValue? k' (insertListConst l toInsert) = v := by rw [insertListConst_eq_insertList] - apply getValue?_insertList_mem (k_beq:=k_beq) + apply getValue?_insertList_of_mem (k_beq:=k_beq) · apply pairwise_list_conversion_insertListConst distinct · simp only [List.mem_map, Prod.exists] exists k exists v -theorem getValue!_insertListConst_not_mem [BEq α] [PartialEquivBEq α] [Inhabited β] +theorem getValue!_insertListConst_of_contains_eq_false [BEq α] [PartialEquivBEq α] [Inhabited β] {l : List ((_ : α) × β)} {toInsert : List (α × β)} {k : α} - {not_mem : (toInsert.map Prod.fst).contains k = false} : + (h : (toInsert.map Prod.fst).contains k = false) : getValue! k (insertListConst l toInsert) = getValue! k l := by rw [insertListConst_eq_insertList] - rw [getValue!_insertList_not_mem] + rw [getValue!_insertList_of_contains_eq_false] rw [list_conversion_insertListConst] - exact not_mem + exact h -theorem getValue!_insertListConst_mem [BEq α] [PartialEquivBEq α] [Inhabited β] - {l : List ((_ : α) × β)} {toInsert : List (α × β)} {k k' : α} {v: β} {k_eq : k == k'} {distinct : List.Pairwise (fun a b => (a.1 == b.1) = false) toInsert } {mem : ⟨k,v⟩ ∈ toInsert}: +theorem getValue!_insertListConst_of_mem [BEq α] [PartialEquivBEq α] [Inhabited β] + {l : List ((_ : α) × β)} {toInsert : List (α × β)} {k k' : α} {v: β} (k_beq : k == k') + (distinct : List.Pairwise (fun a b => (a.1 == b.1) = false) toInsert) (mem : ⟨k, v⟩ ∈ toInsert): getValue! k' (insertListConst l toInsert) = v := by - rw [insertListConst_eq_insertList, getValue!_insertList_mem (k_eq:=k_eq)] + rw [insertListConst_eq_insertList, getValue!_insertList_of_mem (k_beq:=k_beq)] · apply pairwise_list_conversion_insertListConst distinct · simp only [List.mem_map, Prod.exists] exists k exists v -theorem getValueD_insertListConst_not_mem [BEq α] [PartialEquivBEq α] +theorem getValueD_insertListConst_of_contains_eq_false [BEq α] [PartialEquivBEq α] {l : List ((_ : α) × β)} {toInsert : List (α × β)} {k : α} {fallback : β} - {not_mem : (toInsert.map Prod.fst).contains k = false} : + (not_mem : (toInsert.map Prod.fst).contains k = false) : getValueD k (insertListConst l toInsert) fallback = getValueD k l fallback := by - rw [insertListConst_eq_insertList, getValueD_insertList_not_mem] + rw [insertListConst_eq_insertList, getValueD_insertList_of_contains_eq_false] rw [list_conversion_insertListConst] exact not_mem -theorem getValueD_insertListConst_mem [BEq α] [PartialEquivBEq α] - {l : List ((_ : α) × β)} {toInsert : List (α × β)} {k k' : α} {v fallback: β} {k_eq : k == k'} {distinct : List.Pairwise (fun a b => (a.1 == b.1) = false) toInsert } {mem : ⟨k,v⟩ ∈ toInsert}: +theorem getValueD_insertListConst_of_mem [BEq α] [PartialEquivBEq α] + {l : List ((_ : α) × β)} {toInsert : List (α × β)} {k k' : α} {v fallback: β} (k_beq : k == k') + (distinct : List.Pairwise (fun a b => (a.1 == b.1) = false) toInsert) (mem : ⟨k, v⟩ ∈ toInsert): getValueD k' (insertListConst l toInsert) fallback= v := by - rw [insertListConst_eq_insertList, getValueD_insertList_mem (k_eq:= k_eq)] + rw [insertListConst_eq_insertList, getValueD_insertList_of_mem (k_beq:= k_beq)] · apply pairwise_list_conversion_insertListConst distinct · simp exists k exists v -theorem getValue_insertListConst_not_mem [BEq α] [PartialEquivBEq α] +theorem getValue_insertListConst_of_contains_eq_false [BEq α] [PartialEquivBEq α] {l : List ((_ : α) × β)} {toInsert : List (α × β)} {k : α} - {not_mem : (toInsert.map Prod.fst).contains k = false} - {h} : - getValue k (insertListConst l toInsert) h = getValue k l (containsKey_of_containsKey_insertListConst h not_mem) := by + {h : (toInsert.map Prod.fst).contains k = false} + {h'} : + getValue k (insertListConst l toInsert) h' = + getValue k l (containsKey_of_containsKey_insertListConst h' h) := by rw [← Option.some_inj] rw [← getValue?_eq_some_getValue] rw [← getValue?_eq_some_getValue] - rw [getValue?_insertListConst_not_mem] - exact not_mem + rw [getValue?_insertListConst_of_contains_eq_false] + exact h -theorem getValue_insertListConst_mem [BEq α] [PartialEquivBEq α] +theorem getValue_insertListConst_of_mem [BEq α] [PartialEquivBEq α] {l : List ((_ : α) × β)} {toInsert : List (α × β)} - {k k' : α} {k_beq : k == k'} {v : β} - {distinct : toInsert.Pairwise (fun a b => (a.1 == b.1) = false)} - {mem : ⟨k,v⟩ ∈ toInsert} + {k k' : α} (k_beq : k == k') {v : β} + (distinct : toInsert.Pairwise (fun a b => (a.1 == b.1) = false)) + (mem : ⟨k, v⟩ ∈ toInsert) {h} : getValue k' (insertListConst l toInsert) h = v := by rw [← Option.some_inj] rw [← getValue?_eq_some_getValue] - rw [getValue?_insertListConst_mem (mem:=mem)] + rw [getValue?_insertListConst_of_mem (mem:=mem)] . exact k_beq . exact distinct /-- Internal implementation detail of the hash map -/ -def insertListIfNewUnit [BEq α] (l: List ((_ : α) × Unit))(toInsert: List α) : List ((_ : α) × Unit) := +def insertListIfNewUnit [BEq α] (l: List ((_ : α) × Unit))(toInsert: List α) : + List ((_ : α) × Unit) := match toInsert with | .nil => l | .cons hd tl => insertListIfNewUnit (insertEntryIfNew hd () l) tl -theorem insertListIfNewUnit_perm_of_perm_first [BEq α] [EquivBEq α] {l1 l2 : List ((_ : α) × Unit)} {toInsert : List α } (h : Perm l1 l2) - (distinct : DistinctKeys l1) : Perm (insertListIfNewUnit l1 toInsert) (insertListIfNewUnit l2 toInsert) := by +theorem insertListIfNewUnit_perm_of_perm_first [BEq α] [EquivBEq α] {l1 l2 : List ((_ : α) × Unit)} + {toInsert : List α } (h : Perm l1 l2) (distinct : DistinctKeys l1) : + Perm (insertListIfNewUnit l1 toInsert) (insertListIfNewUnit l2 toInsert) := by induction toInsert generalizing l1 l2 with | nil => simp [insertListIfNewUnit, h] | cons hd tl ih => @@ -2746,7 +2774,8 @@ theorem insertListIfNewUnit_perm_of_perm_first [BEq α] [EquivBEq α] {l1 l2 : L exact h · apply DistinctKeys.insertEntryIfNew distinct -theorem DistinctKeys.insertListIfNewUnit [BEq α] [PartialEquivBEq α] (l : List ((_ : α) × Unit))(toInsert : List α) (distinct: DistinctKeys l): +theorem DistinctKeys.insertListIfNewUnit [BEq α] [PartialEquivBEq α] (l : List ((_ : α) × Unit)) + (toInsert : List α) (distinct: DistinctKeys l): DistinctKeys (insertListIfNewUnit l toInsert) := by induction toInsert generalizing l with | nil => simp [List.insertListIfNewUnit, distinct] @@ -2756,7 +2785,8 @@ theorem DistinctKeys.insertListIfNewUnit [BEq α] [PartialEquivBEq α] (l : List apply insertEntryIfNew exact distinct -theorem containsKey_insertListIfNewUnit [BEq α] [PartialEquivBEq α] {l : List ((_ : α) × Unit)} {toInsert : List α} {k : α} : +theorem containsKey_insertListIfNewUnit [BEq α] [PartialEquivBEq α] {l : List ((_ : α) × Unit)} + {toInsert : List α} {k : α} : containsKey k (insertListIfNewUnit l toInsert) = (containsKey k l || toInsert.contains k) := by induction toInsert generalizing l with | nil => simp [insertListIfNewUnit] @@ -2765,8 +2795,20 @@ theorem containsKey_insertListIfNewUnit [BEq α] [PartialEquivBEq α] {l : List rw [ih, containsKey_insertEntryIfNew] rw [Bool.or_comm (hd == k), Bool.or_assoc, BEq.comm (a:=hd)] +theorem containsKey_insertListIfNewUnit_iff [BEq α] [PartialEquivBEq α] {l : List ((_ : α) × Unit)} + {toInsert : List α} {k : α} : + containsKey k (insertListIfNewUnit l toInsert) ↔ containsKey k l ∨ toInsert.contains k := by + induction toInsert generalizing l with + | nil => simp [insertListIfNewUnit] + | cons hd tl ih => + simp only [insertListIfNewUnit, List.contains_cons] + rw [ih, containsKey_insertEntryIfNew] + simp only [Bool.or_eq_true] + rw [or_comm (a:= (hd == k)= true), or_assoc, BEq.comm] + theorem containsKey_of_containsKey_insertListIfNewUnit [BEq α] [PartialEquivBEq α] {l : List ((_ : α) × Unit)} {toInsert : List α} {k : α} - (h₂ : toInsert.contains k = false) : containsKey k (insertListIfNewUnit l toInsert) → containsKey k l := by + (h₂ : toInsert.contains k = false) : containsKey k (insertListIfNewUnit l toInsert) → + containsKey k l := by intro h₁ rw [containsKey_insertListIfNewUnit] at h₁ simp only [Bool.or_eq_true] at h₁ @@ -2775,9 +2817,9 @@ theorem containsKey_of_containsKey_insertListIfNewUnit [BEq α] [PartialEquivBEq | inr h => simp only [h, Bool.true_eq_false] at h₂ -theorem getKey?_insertListIfNewUnit_not_mem [BEq α] [EquivBEq α] +theorem getKey?_insertListIfNewUnit_of_contains_eq_false [BEq α] [EquivBEq α] {l : List ((_ : α) × Unit)} {toInsert : List α} {k : α} - {not_mem : toInsert.contains k = false} : + (not_mem : toInsert.contains k = false) : getKey? k (insertListIfNewUnit l toInsert) = getKey? k l := by induction toInsert generalizing l with | nil => simp [insertListIfNewUnit] @@ -2797,11 +2839,11 @@ theorem getKey?_insertListIfNewUnit_not_mem [BEq α] [EquivBEq α] · simp only [List.contains_cons, Bool.or_eq_false_iff] at not_mem apply And.right not_mem -theorem getKey?_insertListIfNewUnit_mem [BEq α] [EquivBEq α] +theorem getKey?_insertListIfNewUnit_of_mem_of_contains_eq_false [BEq α] [EquivBEq α] {l : List ((_ : α) × Unit)} {toInsert : List α} - {k k' : α} {k_beq : k == k'} - {distinct : toInsert.Pairwise (fun a b => (a == b) = false)} - {mem : k ∈ toInsert} {mem' : containsKey k l = false}: + {k k' : α} (k_beq : k == k') + (distinct : toInsert.Pairwise (fun a b => (a == b) = false)) + (mem : k ∈ toInsert) (mem' : containsKey k l = false) : getKey? k' (insertListIfNewUnit l toInsert) = some k := by induction toInsert generalizing l with | nil => simp at mem @@ -2810,7 +2852,7 @@ theorem getKey?_insertListIfNewUnit_mem [BEq α] [EquivBEq α] simp only [List.mem_cons] at mem cases mem with | inl mem => - rw [getKey?_insertListIfNewUnit_not_mem, ← mem] + rw [getKey?_insertListIfNewUnit_of_contains_eq_false, ← mem] · simp only [insertEntryIfNew, mem', cond_eq_if, Bool.false_eq_true, ↓reduceIte, getKey?_cons, ite_eq_left_iff, Bool.not_eq_true] intro h @@ -2835,11 +2877,11 @@ theorem getKey?_insertListIfNewUnit_mem [BEq α] [EquivBEq α] apply (And.left distinct) k mem · exact mem' -theorem getKey?_insertListIfNewUnit_mem_both [BEq α] [EquivBEq α] +theorem getKey?_insertListIfNewUnit_of_mem_of_contains [BEq α] [EquivBEq α] {l : List ((_ : α) × Unit)} {toInsert : List α} {k : α} - {distinct : toInsert.Pairwise (fun a b => (a == b) = false)} - {mem : toInsert.contains k} {mem' : containsKey k l = true}: + (distinct : toInsert.Pairwise (fun a b => (a == b) = false)) + (mem : toInsert.contains k) (mem' : containsKey k l = true): getKey? k (insertListIfNewUnit l toInsert) = getKey? k l := by induction toInsert generalizing l with | nil => simp at mem @@ -2848,7 +2890,7 @@ theorem getKey?_insertListIfNewUnit_mem_both [BEq α] [EquivBEq α] cases mem with | inl mem => simp only [insertListIfNewUnit] - rw [getKey?_insertListIfNewUnit_not_mem, getKey?_insertEntryIfNew] + rw [getKey?_insertListIfNewUnit_of_contains_eq_false, getKey?_insertEntryIfNew] · simp only [BEq.symm mem, true_and, ite_eq_right_iff] have h':= containsKey_of_beq mem' mem intro h @@ -2875,7 +2917,7 @@ theorem getKey?_insertListIfNewUnit_mem_both [BEq α] [EquivBEq α] right apply mem' -theorem getKey_insertListIfNewUnit_not_mem [BEq α] [EquivBEq α] +theorem getKey_insertListIfNewUnit_of_contains_eq_false [BEq α] [EquivBEq α] {l : List ((_ : α) × Unit)} {toInsert : List α} {k : α} {not_mem : toInsert.contains k = false} {h} : @@ -2883,115 +2925,115 @@ theorem getKey_insertListIfNewUnit_not_mem [BEq α] [EquivBEq α] getKey k l (containsKey_of_containsKey_insertListIfNewUnit not_mem h) := by rw [← Option.some_inj] rw [← getKey?_eq_some_getKey] - rw [getKey?_insertListIfNewUnit_not_mem] + rw [getKey?_insertListIfNewUnit_of_contains_eq_false] . rw [getKey?_eq_some_getKey] . exact not_mem -theorem getKey_insertListIfNewUnit_mem [BEq α] [EquivBEq α] +theorem getKey_insertListIfNewUnit_of_mem_of_contains_eq_false [BEq α] [EquivBEq α] {l : List ((_ : α) × Unit)} {toInsert : List α} - {k k' : α} {k_beq : k == k'} - {distinct : toInsert.Pairwise (fun a b => (a == b) = false)} - {mem : k ∈ toInsert} + {k k' : α} (k_beq : k == k') + (distinct : toInsert.Pairwise (fun a b => (a == b) = false)) + (mem : k ∈ toInsert) {h} : containsKey k l = false → getKey k' (insertListIfNewUnit l toInsert) h = k := by intro mem' rw [← Option.some_inj] rw [← getKey?_eq_some_getKey] - rw [getKey?_insertListIfNewUnit_mem] + rw [getKey?_insertListIfNewUnit_of_mem_of_contains_eq_false] . exact k_beq . exact distinct . exact mem . exact mem' -theorem getKey_insertListIfNewUnit_mem_both [BEq α] [EquivBEq α] +theorem getKey_insertListIfNewUnit_of_mem_of_contains [BEq α] [EquivBEq α] {l : List ((_ : α) × Unit)} {toInsert : List α} {k : α} - {distinct : toInsert.Pairwise (fun a b => (a == b) = false)} - {mem : toInsert.contains k} {mem' : containsKey k l = true} {h}: + (distinct : toInsert.Pairwise (fun a b => (a == b) = false)) + (mem : toInsert.contains k) (mem' : containsKey k l = true) {h}: getKey k (insertListIfNewUnit l toInsert) h = getKey k l mem' := by rw [← Option.some_inj] rw [← getKey?_eq_some_getKey] rw [← getKey?_eq_some_getKey] - rw [getKey?_insertListIfNewUnit_mem_both] + rw [getKey?_insertListIfNewUnit_of_mem_of_contains] · exact distinct · exact mem · exact mem' -theorem getKey!_insertListIfNewUnit_not_mem [BEq α] [EquivBEq α] [Inhabited α] +theorem getKey!_insertListIfNewUnit_of_contains_eq_false [BEq α] [EquivBEq α] [Inhabited α] {l : List ((_ : α) × Unit)} {toInsert : List α} {k : α} - {not_mem : toInsert.contains k = false} : + (h : toInsert.contains k = false) : getKey! k (insertListIfNewUnit l toInsert) = getKey! k l := by rw [getKey!_eq_getKey?] - rw [getKey?_insertListIfNewUnit_not_mem] + rw [getKey?_insertListIfNewUnit_of_contains_eq_false] . rw [getKey!_eq_getKey?] - . exact not_mem + . exact h -theorem getKey!_insertListIfNewUnit_mem [BEq α] [EquivBEq α] [Inhabited α] +theorem getKey!_insertListIfNewUnit_of_mem_of_contains_eq_false [BEq α] [EquivBEq α] [Inhabited α] {l : List ((_ : α) × Unit)} {toInsert : List α} - {k k' : α} {k_beq : k == k'} - {distinct : toInsert.Pairwise (fun a b => (a == b) = false)} - {mem : k ∈ toInsert} {mem': containsKey k l = false} : + {k k' : α} (k_beq : k == k') + (distinct : toInsert.Pairwise (fun a b => (a == b) = false)) + (mem : k ∈ toInsert) (h': containsKey k l = false) : getKey! k' (insertListIfNewUnit l toInsert) = k := by rw [getKey!_eq_getKey?] - rw [getKey?_insertListIfNewUnit_mem] + rw [getKey?_insertListIfNewUnit_of_mem_of_contains_eq_false] . rw [Option.get!_some] . exact k_beq . exact distinct . exact mem - . exact mem' + . exact h' -theorem getKey!_insertListIfNewUnit_mem_both [BEq α] [EquivBEq α] [Inhabited α] +theorem getKey!_insertListIfNewUnit_of_mem_of_contains [BEq α] [EquivBEq α] [Inhabited α] {l : List ((_ : α) × Unit)} {toInsert : List α} {k : α} - {distinct : toInsert.Pairwise (fun a b => (a == b) = false)} - {mem : toInsert.contains k} {mem' : containsKey k l = true} : + (distinct : toInsert.Pairwise (fun a b => (a == b) = false)) + (mem : toInsert.contains k) (h' : containsKey k l = true) : getKey! k (insertListIfNewUnit l toInsert) = getKey! k l := by rw [getKey!_eq_getKey?] - rw [getKey?_insertListIfNewUnit_mem_both] + rw [getKey?_insertListIfNewUnit_of_mem_of_contains] . rw [getKey!_eq_getKey?] · exact distinct · exact mem - · exact mem' + · exact h' -theorem getKeyD_insertListIfNewUnit_not_mem [BEq α] [EquivBEq α] +theorem getKeyD_insertListIfNewUnit_of_contains_eq_false [BEq α] [EquivBEq α] {l : List ((_ : α) × Unit)} {toInsert : List α} {k fallback : α} - {not_mem : toInsert.contains k = false} : + (h : toInsert.contains k = false) : getKeyD k (insertListIfNewUnit l toInsert) fallback = getKeyD k l fallback := by rw [getKeyD_eq_getKey?] - rw [getKey?_insertListIfNewUnit_not_mem] + rw [getKey?_insertListIfNewUnit_of_contains_eq_false] . rw [getKeyD_eq_getKey?] - . exact not_mem + . exact h -theorem getKeyD_insertListIfNewUnit_mem [BEq α] [EquivBEq α] +theorem getKeyD_insertListIfNewUnit_of_mem_of_contains_eq_false [BEq α] [EquivBEq α] {l : List ((_ : α) × Unit)} {toInsert : List α} - {k k' fallback : α} {k_beq : k == k'} - {distinct : toInsert.Pairwise (fun a b => (a == b) = false)} - {mem : k ∈ toInsert } {mem': containsKey k l = false} : + {k k' fallback : α} (k_beq : k == k') + (distinct : toInsert.Pairwise (fun a b => (a == b) = false)) + {mem : k ∈ toInsert } {h': containsKey k l = false} : getKeyD k' (insertListIfNewUnit l toInsert) fallback = k := by rw [getKeyD_eq_getKey?] - rw [getKey?_insertListIfNewUnit_mem] + rw [getKey?_insertListIfNewUnit_of_mem_of_contains_eq_false] . rw [Option.getD_some] . exact k_beq . exact distinct . exact mem - . exact mem' + . exact h' -theorem getKeyD_insertListIfNewUnit_mem_both [BEq α] [EquivBEq α] +theorem getKeyD_insertListIfNewUnit_of_mem_of_contains [BEq α] [EquivBEq α] {l : List ((_ : α) × Unit)} {toInsert : List α} {k fallback : α} - {distinct : toInsert.Pairwise (fun a b => (a == b) = false)} - {mem : toInsert.contains k} {mem' : containsKey k l = true} : + (distinct : toInsert.Pairwise (fun a b => (a == b) = false)) + (mem : toInsert.contains k) (mem' : containsKey k l = true) : getKeyD k (insertListIfNewUnit l toInsert) fallback = getKeyD k l fallback := by rw [getKeyD_eq_getKey?] - rw [getKey?_insertListIfNewUnit_mem_both] + rw [getKey?_insertListIfNewUnit_of_mem_of_contains] . rw [getKeyD_eq_getKey?] · exact distinct · exact mem · exact mem' -theorem insertListIfNewUnit_length [BEq α] [EquivBEq α] +theorem length_insertListIfNewUnit [BEq α] [EquivBEq α] {l : List ((_ : α) × Unit)} {toInsert : List α} - {distinct_l : DistinctKeys l} + (distinct_l : DistinctKeys l) (distinct_toInsert : List.Pairwise (fun a b => (a == b) = false) toInsert) (distinct_both : ∀ (a : α), ¬ (containsKey a l ∧ toInsert.contains a)) : (insertListIfNewUnit l toInsert).length = l.length + toInsert.length := by @@ -3006,6 +3048,7 @@ theorem insertListIfNewUnit_length [BEq α] [EquivBEq α] Bool.not_eq_true] at distinct_both simp only [distinct_both, Bool.false_eq_true, ↓reduceIte] rw [Nat.add_assoc, Nat.add_comm 1 _] + · apply DistinctKeys.insertEntryIfNew distinct_l · simp only [pairwise_cons] at distinct_toInsert apply And.right distinct_toInsert · intro a @@ -3029,18 +3072,17 @@ theorem insertListIfNewUnit_length [BEq α] [EquivBEq α] | inr h => specialize distinct_both a h apply And.right distinct_both - · apply DistinctKeys.insertEntryIfNew distinct_l -theorem insertListIfNewUnit_not_isEmpty_if_start_not_isEmpty [BEq α] +theorem isEmpty_insertListIfNewUnit_eq_false_of_isEmpty_eq_false [BEq α] {l : List ((_ : α) × Unit)} {toInsert : List α} - {h : l.isEmpty = false} : + (h : l.isEmpty = false) : (List.insertListIfNewUnit l toInsert).isEmpty = false := by induction toInsert generalizing l with | nil => simp [insertListIfNewUnit, h] | cons hd tl ih => simp [insertListIfNewUnit, ih] -theorem insertListIfNewUnit_isEmpty [BEq α] +theorem isEmpty_insertListIfNewUnit [BEq α] {l : List ((_ : α) × Unit)} {toInsert : List α} : (List.insertListIfNewUnit l toInsert).isEmpty ↔ l.isEmpty ∧ toInsert.isEmpty := by induction toInsert with @@ -3049,7 +3091,7 @@ theorem insertListIfNewUnit_isEmpty [BEq α] simp only [insertListIfNewUnit, List.isEmpty_cons, Bool.false_eq_true, and_false, iff_false] apply ne_true_of_eq_false - apply insertListIfNewUnit_not_isEmpty_if_start_not_isEmpty + apply isEmpty_insertListIfNewUnit_eq_false_of_isEmpty_eq_false simp theorem getValue?_list_unit [BEq α] {l : List ((_ : α) × Unit)} {k : α}: @@ -3083,8 +3125,8 @@ theorem getValueD_list_unit [BEq α] {l : List ((_ : α) × Unit)} {k : α} {fal theorem getValue?_insertListIfNewUnit [BEq α] [PartialEquivBEq α] {l : List ((_ : α) × Unit)} {toInsert : List α} {k : α}: getValue? k (insertListIfNewUnit l toInsert) = - if (containsKey k l || toInsert.contains k) then some () else none := by - simp [← containsKey_insertListIfNewUnit, getValue?_list_unit] + if containsKey k l ∨ toInsert.contains k then some () else none := by + simp [← containsKey_insertListIfNewUnit_iff, getValue?_list_unit] theorem getValue_insertListIfNewUnit [BEq α] [PartialEquivBEq α] {l : List ((_ : α) × Unit)} {toInsert : List α} {k : α} {h}: diff --git a/src/Std/Data/DHashMap/Internal/Model.lean b/src/Std/Data/DHashMap/Internal/Model.lean index b221595ae2ff..8b7eb32018a0 100644 --- a/src/Std/Data/DHashMap/Internal/Model.lean +++ b/src/Std/Data/DHashMap/Internal/Model.lean @@ -529,7 +529,7 @@ theorem Const.insertMany_eq_insertListₘ [BEq α] [Hashable α] (m : Raw₀ α simp only [List.foldl_cons,insertListₘ] apply ih -theorem Const.insertManyIfNewUnit_eq_insertListIfNewUnit [BEq α] [Hashable α] (m : Raw₀ α (fun _ => Unit)) (l: List α): +theorem Const.insertManyIfNewUnit_eq_insertListIfNewUnitₘ [BEq α] [Hashable α] (m : Raw₀ α (fun _ => Unit)) (l: List α): (Const.insertManyIfNewUnit m l).1 = Const.insertListIfNewUnitₘ m l := by simp only [insertManyIfNewUnit, Id.run, Id.pure_eq, Id.bind_eq, List.forIn_yield_eq_foldl] suffices ∀ (t : { m' // ∀ (P : Raw₀ α (fun _ => Unit) → Prop), diff --git a/src/Std/Data/DHashMap/Internal/RawLemmas.lean b/src/Std/Data/DHashMap/Internal/RawLemmas.lean index fd574828cfec..cedc1b443c4a 100644 --- a/src/Std/Data/DHashMap/Internal/RawLemmas.lean +++ b/src/Std/Data/DHashMap/Internal/RawLemmas.lean @@ -70,9 +70,10 @@ variable [BEq α] [Hashable α] /-- Internal implementation detail of the hash map -/ scoped macro "wf_trivial" : tactic => `(tactic| repeat (first - | apply Raw₀.wfImp_insert | apply Raw₀.wfImp_insertIfNew | apply Raw₀.wfImp_insertMany | apply Raw₀.Const.wfImp_insertMany| apply Raw₀.Const.wfImp_insertManyIfNewUnit | apply Raw₀.wfImp_erase - | apply Raw.WF.out | assumption | apply Raw₀.wfImp_empty | apply Raw.WFImp.distinct - | apply Raw.WF.empty₀)) + | apply Raw₀.wfImp_insert | apply Raw₀.wfImp_insertIfNew | apply Raw₀.wfImp_insertMany + | apply Raw₀.Const.wfImp_insertMany| apply Raw₀.Const.wfImp_insertManyIfNewUnit + | apply Raw₀.wfImp_erase | apply Raw.WF.out | assumption | apply Raw₀.wfImp_empty + | apply Raw.WFImp.distinct | apply Raw.WF.empty₀)) /-- Internal implementation detail of the hash map -/ scoped macro "empty" : tactic => `(tactic| { intros; simp_all [List.isEmpty_iff] } ) @@ -90,8 +91,9 @@ private def queryNames : Array Name := ``Raw.pairwise_keys_iff_pairwise_keys] private def modifyNames : Array Name := - #[``toListModel_insert, ``toListModel_erase, ``toListModel_insertIfNew, ``toListModel_insertMany_list, - ``Const.toListModel_insertMany_list, ``Const.toListModel_insertManyIfNewUnit_list] + #[``toListModel_insert, ``toListModel_erase, ``toListModel_insertIfNew, + ``toListModel_insertMany_list, ``Const.toListModel_insertMany_list, + ``Const.toListModel_insertManyIfNewUnit_list] private def congrNames : MacroM (Array (TSyntax `term)) := do return #[← `(_root_.List.Perm.isEmpty_eq), ← `(containsKey_of_perm), @@ -841,381 +843,370 @@ theorem distinct_keys [EquivBEq α] [LawfulHashable α] (h : m.1.WF) : simp_to_model using (Raw.WF.out h).distinct.distinct @[simp] -theorem insertMany_list_nil : +theorem insertMany_nil : m.insertMany [] = m := by simp [insertMany, Id.run] @[simp] -theorem insertMany_list_singleton [EquivBEq α] [LawfulHashable α] {k : α} {v : β k} : - m.insertMany [⟨k,v⟩] = m.insert k v := by +theorem insertMany_list_singleton {k : α} {v : β k} : + m.insertMany [⟨k, v⟩] = m.insert k v := by simp [insertMany, Id.run] -theorem insertMany_list_cons {l : List ((a : α) × (β a))} {k : α} {v : β k} : - (m.insertMany (⟨k,v⟩ :: l)).1 = ((m.insert k v).insertMany l).1 := by +theorem insertMany_cons {l : List ((a : α) × β a)} {k : α} {v : β k} : + (m.insertMany (⟨k, v⟩ :: l)).1 = ((m.insert k v).insertMany l).1 := by simp only [insertMany_eq_insertListₘ] cases l with | nil => simp [insertListₘ] | cons hd tl => simp [insertListₘ] -theorem contains_insertMany_list [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l : List ((a : α) × (β a))} {k : α} : +theorem contains_insertMany_list [EquivBEq α] [LawfulHashable α] (h : m.1.WF) + {l : List ((a : α) × β a)} {k : α} : (m.insertMany l).1.contains k = (m.contains k || (l.map Sigma.fst).contains k) := by simp_to_model using List.containsKey_insertList -theorem contains_of_contains_insertMany_list [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l : List ((a : α) × (β a))} {k : α} : +theorem contains_of_contains_insertMany_list [EquivBEq α] [LawfulHashable α] (h : m.1.WF) + {l : List ((a : α) × β a)} {k : α} : (m.insertMany l).1.contains k → (l.map Sigma.fst).contains k = false → m.contains k := by simp_to_model using List.containsKey_of_containsKey_insertList -theorem get?_insertMany_list_not_mem [LawfulBEq α] - {l : List ((a : α) × (β a))} {k : α} - {not_mem : (l.map (Sigma.fst)).contains k = false} - (h : m.1.WF) : +theorem get?_insertMany_list_of_contains_eq_false [LawfulBEq α] (h : m.1.WF) + {l : List ((a : α) × β a)} {k : α} + (h' : (l.map (Sigma.fst)).contains k = false) : (m.insertMany l).1.get? k = m.get? k := by - simp_to_model using getValueCast?_insertList_not_mem - -theorem get?_insertMany_list_mem [LawfulBEq α] - {l : List ((a : α) × (β a))} {k k' : α} {k_eq : k == k'} {v : β k} - {distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)} - {mem : ⟨k,v⟩ ∈ l} - (h : m.1.WF) : - (m.insertMany l).1.get? k' = some (cast (by congr; apply LawfulBEq.eq_of_beq k_eq) v) := by - simp_to_model using getValueCast?_insertList_mem - -theorem get_insertMany_list_not_mem [LawfulBEq α] - {l : List ((a : α) × (β a))} {k : α} - {not_mem : (l.map (Sigma.fst)).contains k = false} - {h'} - (h : m.1.WF) : - (m.insertMany l).1.get k h' = m.get k (contains_of_contains_insertMany_list _ h h' not_mem) := by - simp_to_model using getValueCast_insertList_not_mem - -theorem get_insertMany_list_mem [LawfulBEq α] - {l : List ((a : α) × (β a))} {k k' : α} {k_eq : k == k'} {v : β k} - {distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)} - {mem : ⟨k,v⟩ ∈ l} - {h'} - (h : m.1.WF) : - (m.insertMany l).1.get k' h' = cast (by congr; apply LawfulBEq.eq_of_beq k_eq) v := by - simp_to_model using getValueCast_insertList_mem - -theorem get!_insertMany_list_not_mem [LawfulBEq α] - {l : List ((a : α) × (β a))} {k : α} [Inhabited (β k)] - {not_mem : (l.map (Sigma.fst)).contains k = false} - (h : m.1.WF) : + simp_to_model using getValueCast?_insertList_of_contains_eq_false + +theorem get?_insertMany_list_of_mem [LawfulBEq α] (h : m.1.WF) + {l : List ((a : α) × β a)} {k k' : α} (k_beq : k == k') {v : β k} + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) + (mem : ⟨k, v⟩ ∈ l) : + (m.insertMany l).1.get? k' = some (cast (by congr; apply LawfulBEq.eq_of_beq k_beq) v) := by + simp_to_model using getValueCast?_insertList_of_mem + +theorem get_insertMany_list_of_contains_eq_false [LawfulBEq α] (h : m.1.WF) + {l : List ((a : α) × β a)} {k : α} + (contains : (l.map (Sigma.fst)).contains k = false) + {h'} : + (m.insertMany l).1.get k h' = + m.get k (contains_of_contains_insertMany_list _ h h' contains) := by + simp_to_model using getValueCast_insertList_of_contains_eq_false + +theorem get_insertMany_list_of_mem [LawfulBEq α] (h : m.1.WF) + {l : List ((a : α) × β a)} {k k' : α} (k_beq : k == k') {v : β k} + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) + (mem : ⟨k, v⟩ ∈ l) + {h'} : + (m.insertMany l).1.get k' h' = cast (by congr; apply LawfulBEq.eq_of_beq k_beq) v := by + simp_to_model using getValueCast_insertList_of_mem + +theorem get!_insertMany_list_of_contains_eq_false [LawfulBEq α] (h : m.1.WF) + {l : List ((a : α) × β a)} {k : α} [Inhabited (β k)] + (h' : (l.map (Sigma.fst)).contains k = false) : (m.insertMany l).1.get! k = m.get! k := by - simp_to_model using getValueCast!_insertList_not_mem - -theorem get!_insertMany_list_mem [LawfulBEq α] - {l : List ((a : α) × (β a))} {k k' : α} {k_eq : k == k'} {v : β k} [Inhabited (β k')] - {distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)} - {mem : ⟨k,v⟩ ∈ l} - (h : m.1.WF) : - (m.insertMany l).1.get! k' = cast (by congr; apply LawfulBEq.eq_of_beq k_eq) v := by - simp_to_model using getValueCast!_insertList_mem - -theorem getD_insertMany_list_not_mem [LawfulBEq α] - {l : List ((a : α) × (β a))} {k : α} {fallback : β k} - {not_mem : (l.map (Sigma.fst)).contains k = false} - (h : m.1.WF) : + simp_to_model using getValueCast!_insertList_of_contains_eq_false + +theorem get!_insertMany_list_of_mem [LawfulBEq α] (h : m.1.WF) + {l : List ((a : α) × β a)} {k k' : α} (k_beq : k == k') {v : β k} [Inhabited (β k')] + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) + (mem : ⟨k, v⟩ ∈ l) : + (m.insertMany l).1.get! k' = cast (by congr; apply LawfulBEq.eq_of_beq k_beq) v := by + simp_to_model using getValueCast!_insertList_of_mem + +theorem getD_insertMany_list_of_contains_eq_false [LawfulBEq α] (h : m.1.WF) + {l : List ((a : α) × β a)} {k : α} {fallback : β k} + (not_mem : (l.map (Sigma.fst)).contains k = false) : (m.insertMany l).1.getD k fallback = m.getD k fallback := by - simp_to_model using getValueCastD_insertList_not_mem - -theorem getD_insertMany_list_mem [LawfulBEq α] - {l : List ((a : α) × (β a))} {k k' : α} {k_eq : k == k'} {v : β k} {fallback : β k'} - {distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)} - {mem : ⟨k,v⟩ ∈ l} - (h : m.1.WF) : - (m.insertMany l).1.getD k' fallback = cast (by congr; apply LawfulBEq.eq_of_beq k_eq) v := by - simp_to_model using getValueCastD_insertList_mem - -theorem getKey?_insertMany_list_not_mem [EquivBEq α] [LawfulHashable α] (h : m.1.WF) - {l : List ((a : α) × (β a))} {k : α} - {not_mem : (l.map Sigma.fst).contains k = false} : + simp_to_model using getValueCastD_insertList_of_contains_eq_false + +theorem getD_insertMany_list_of_mem [LawfulBEq α] (h : m.1.WF) + {l : List ((a : α) × β a)} {k k' : α} (k_beq : k == k') {v : β k} {fallback : β k'} + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) + (mem : ⟨k, v⟩ ∈ l) : + (m.insertMany l).1.getD k' fallback = cast (by congr; apply LawfulBEq.eq_of_beq k_beq) v := by + simp_to_model using getValueCastD_insertList_of_mem + +theorem getKey?_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] (h : m.1.WF) + {l : List ((a : α) × β a)} {k : α} + (h : (l.map Sigma.fst).contains k = false) : (m.insertMany l).1.getKey? k = m.getKey? k := by - simp_to_model using List.getKey?_insertList_not_mem + simp_to_model using getKey?_insertList_of_contains_eq_false -theorem getKey?_insertMany_list_mem [EquivBEq α] [LawfulHashable α] (h : m.1.WF) - {l : List ((a : α) × (β a))} - {k k' : α} {k_beq : k == k'} - {distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)} - {mem : k ∈ (l.map Sigma.fst)} : +theorem getKey?_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] (h : m.1.WF) + {l : List ((a : α) × β a)} + {k k' : α} (k_beq : k == k') + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) + (mem : k ∈ (l.map Sigma.fst)) : (m.insertMany l).1.getKey? k' = some k := by - simp_to_model using List.getKey?_insertList_mem + simp_to_model using getKey?_insertList_of_mem -theorem getKey_insertMany_list_not_mem [EquivBEq α] [LawfulHashable α] (h : m.1.WF) - {l : List ((a : α) × (β a))} {k : α} - {not_mem : (l.map Sigma.fst).contains k = false} +theorem getKey_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] (h : m.1.WF) + {l : List ((a : α) × β a)} {k : α} + (h₁ : (l.map Sigma.fst).contains k = false) {h'} : - (m.insertMany l).1.getKey k h' = m.getKey k (contains_of_contains_insertMany_list _ h h' not_mem) := by - simp_to_model using List.getKey_insertList_not_mem - -theorem getKey_insertMany_list_mem [EquivBEq α] [LawfulHashable α] (h : m.1.WF) - {l : List ((a : α) × (β a))} - {k k' : α} {k_beq : k == k'} - {distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)} - {mem : k ∈ (l.map Sigma.fst)} + (m.insertMany l).1.getKey k h' = + m.getKey k (contains_of_contains_insertMany_list _ h h' h₁) := by + simp_to_model using getKey_insertList_of_contains_eq_false + +theorem getKey_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] (h : m.1.WF) + {l : List ((a : α) × β a)} + {k k' : α} (k_beq : k == k') + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) + (mem : k ∈ (l.map Sigma.fst)) {h'} : (m.insertMany l).1.getKey k' h' = k := by - simp_to_model using List.getKey_insertList_mem + simp_to_model using getKey_insertList_of_mem -theorem getKey!_insertMany_list_not_mem [EquivBEq α] [LawfulHashable α] [Inhabited α] (h : m.1.WF) - {l : List ((a : α) × (β a))} {k : α} - {not_mem : (l.map Sigma.fst).contains k = false} : +theorem getKey!_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] [Inhabited α] + (h : m.1.WF) {l : List ((a : α) × β a)} {k : α} + (h' : (l.map Sigma.fst).contains k = false) : (m.insertMany l).1.getKey! k = m.getKey! k := by - simp_to_model using List.getKey!_insertList_not_mem + simp_to_model using getKey!_insertList_of_contains_eq_false -theorem getKey!_insertMany_list_mem [EquivBEq α] [LawfulHashable α] [Inhabited α] (h : m.1.WF) - {l : List ((a : α) × (β a))} - {k k' : α} {k_beq : k == k'} - {distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)} - {mem : k ∈ (l.map Sigma.fst)} : +theorem getKey!_insertMany_listof_mem [EquivBEq α] [LawfulHashable α] [Inhabited α] (h : m.1.WF) + {l : List ((a : α) × β a)} + {k k' : α} (k_beq : k == k') + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) + (mem : k ∈ (l.map Sigma.fst)) : (m.insertMany l).1.getKey! k' = k := by - simp_to_model using List.getKey!_insertList_mem + simp_to_model using getKey!_insertList_of_mem -theorem getKeyD_insertMany_list_not_mem [EquivBEq α] [LawfulHashable α] (h : m.1.WF) - {l : List ((a : α) × (β a))} {k fallback : α} - {not_mem : (l.map Sigma.fst).contains k = false} : +theorem getKeyD_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] (h : m.1.WF) + {l : List ((a : α) × β a)} {k fallback : α} + (h' : (l.map Sigma.fst).contains k = false) : (m.insertMany l).1.getKeyD k fallback = m.getKeyD k fallback := by - simp_to_model using List.getKeyD_insertList_not_mem + simp_to_model using getKeyD_insertList_of_contains_eq_false -theorem getKeyD_insertMany_list_mem [EquivBEq α] [LawfulHashable α] (h : m.1.WF) - {l : List ((a : α) × (β a))} - {k k' fallback : α} {k_beq : k == k'} - {distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)} - {mem : k ∈ (l.map Sigma.fst)} : +theorem getKeyD_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] (h : m.1.WF) + {l : List ((a : α) × β a)} + {k k' fallback : α} (k_beq : k == k') + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) + (mem : k ∈ (l.map Sigma.fst)) : (m.insertMany l).1.getKeyD k' fallback = k := by - simp_to_model using List.getKeyD_insertList_mem + simp_to_model using getKeyD_insertList_of_mem -theorem size_insertMany_list [EquivBEq α] [LawfulHashable α] - {l : List ((a : α) × (β a))} - {distinct : List.Pairwise (fun a b => (a.1 == b.1) = false) l} - (h : m.1.WF) : - (∀ (a : α), ¬ (m.contains a = true ∧ (l.map Sigma.fst).contains a = true)) → +theorem size_insertMany_list [EquivBEq α] [LawfulHashable α] (h : m.1.WF) + {l : List ((a : α) × β a)} {distinct : List.Pairwise (fun a b => (a.1 == b.1) = false) l}: + (∀ (a : α), ¬ (m.contains a ∧ (l.map Sigma.fst).contains a)) → (m.insertMany l).1.1.size = m.1.size + l.length := by - simp_to_model using List.insertList_length + simp_to_model using length_insertList -theorem insertMany_list_notEmpty_if_map_notEmpty [EquivBEq α] [LawfulHashable α] - {l : List ((a : α) × (β a))} (h : m.1.WF) : +theorem isEmpty_insertMany_list_eq_false_of_isEmpty_eq_false [EquivBEq α] [LawfulHashable α] + (h : m.1.WF) {l : List ((a : α) × β a)}: (m.1.isEmpty = false) → (m.insertMany l).1.1.isEmpty = false := by - simp_to_model using List.insertList_not_isEmpty_if_start_not_isEmpty + simp_to_model using isEmpty_insertList_eq_false_of_isEmpty_eq_false -theorem insertMany_list_isEmpty [EquivBEq α] [LawfulHashable α] - {l : List ((a : α) × (β a))} (h : m.1.WF) : +theorem isEmpty_insertMany_list [EquivBEq α] [LawfulHashable α] (h : m.1.WF) + {l : List ((a : α) × β a)} : (m.insertMany l).1.1.isEmpty ↔ m.1.isEmpty ∧ l.isEmpty := by - simp_to_model using List.insertList_isEmpty + simp_to_model using isEmpty_insertList section variable {β: Type v} (m: Raw₀ α (fun _ => β)) -theorem get?_insertMany_list_not_mem_const [EquivBEq α] [LawfulHashable α] - {l : List ((_ : α) × β)} {k : α} - {not_mem : (l.map Sigma.fst).contains k = false} - (h : m.1.WF) : +theorem get?_insertMany_list_of_contains_eq_false_const [EquivBEq α] [LawfulHashable α] (h : m.1.WF) + {l : List ((_ : α) × β)} {k : α} (h : (l.map Sigma.fst).contains k = false) : Const.get? (m.insertMany l).1 k = Const.get? m k := by - simp_to_model using getValue?_insertList_not_mem + simp_to_model using getValue?_insertList_of_contains_eq_false -theorem get?_insertMany_list_mem_const [EquivBEq α] [LawfulHashable α] - {l : List ((_ : α) × β)} {k k' : α} {k_eq: k == k'} {v : β} {mem : ⟨k,v⟩ ∈ l} - {distinct : List.Pairwise (fun a b => (a.1 == b.1) = false ) l} (h : m.1.WF): +theorem get?_insertMany_list_of_mem_const [EquivBEq α] [LawfulHashable α] (h : m.1.WF) + {l : List ((_ : α) × β)} {k k' : α} (k_beq : k == k') {v : β} (mem : ⟨k, v⟩ ∈ l) + (distinct : List.Pairwise (fun a b => (a.1 == b.1) = false ) l) : Const.get? (m.insertMany l).1 k' = v := by - simp_to_model using getValue?_insertList_mem + simp_to_model using getValue?_insertList_of_mem -theorem get_insertMany_list_not_mem_const [EquivBEq α] [LawfulHashable α] +theorem get_insertMany_list_of_contains_eq_false_const [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l : List ((_ : α) × β)} {k : α} - {not_mem : (l.map Sigma.fst).contains k = false} - (h : m.1.WF) {h'}: - Const.get (m.insertMany l).1 k h' = Const.get m k (contains_of_contains_insertMany_list _ h h' not_mem) := by - simp_to_model using getValue_insertList_not_mem - -theorem get_insertMany_list_mem_const [EquivBEq α] [LawfulHashable α] - {l : List ((_ : α) × β)} {k k' : α} {k_eq: k == k'} {v : β} {mem : ⟨k,v⟩ ∈ l} - {distinct : List.Pairwise (fun a b => (a.1 == b.1) = false ) l} (h : m.1.WF) {h'}: + (h₁ : (l.map Sigma.fst).contains k = false) {h'}: + Const.get (m.insertMany l).1 k h' = + Const.get m k (contains_of_contains_insertMany_list _ h h' h₁) := by + simp_to_model using getValue_insertList_of_contains_eq_false + +theorem get_insertMany_list_of_mem_const [EquivBEq α] [LawfulHashable α] (h : m.1.WF) + {l : List ((_ : α) × β)} {k k' : α} (k_beq : k == k') {v : β} (mem : ⟨k, v⟩ ∈ l) + (distinct : List.Pairwise (fun a b => (a.1 == b.1) = false ) l) {h'}: Const.get (m.insertMany l).1 k' h' = v := by - simp_to_model using getValue_insertList_mem + simp_to_model using getValue_insertList_of_mem -theorem get!_insertMany_list_not_mem_const [EquivBEq α] [LawfulHashable α] [Inhabited β] - {l : List ((_ : α) × β)} {k : α} - {not_mem : (l.map Sigma.fst).contains k = false} - (h : m.1.WF) : +theorem get!_insertMany_list_of_contains_eq_false_const [EquivBEq α] [LawfulHashable α] + [Inhabited β] (h : m.1.WF) {l : List ((_ : α) × β)} {k : α} + (h' : (l.map Sigma.fst).contains k = false) : Const.get! (m.insertMany l).1 k = Const.get! m k := by - simp_to_model using getValue!_insertList_not_mem + simp_to_model using getValue!_insertList_of_contains_eq_false -theorem get!_insertMany_list_mem_const [EquivBEq α] [LawfulHashable α] [Inhabited β] - {l : List ((_ : α) × β)} {k k' : α} {k_eq: k == k'} {v : β} {mem : ⟨k,v⟩ ∈ l} - {distinct : List.Pairwise (fun a b => (a.1 == b.1) = false ) l} (h : m.1.WF): +theorem get!_insertMany_list_of_mem_const [EquivBEq α] [LawfulHashable α] [Inhabited β] (h : m.1.WF) + {l : List ((_ : α) × β)} {k k' : α} (k_beq : k == k') {v : β} (mem: ⟨k, v⟩ ∈ l) + (distinct : List.Pairwise (fun a b => (a.1 == b.1) = false ) l) : Const.get! (m.insertMany l).1 k' = v := by - simp_to_model using getValue!_insertList_mem + simp_to_model using getValue!_insertList_of_mem -theorem getD_insertMany_list_not_mem_const [EquivBEq α] [LawfulHashable α] +theorem getD_insertMany_list_of_contains_eq_false_const [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l : List ((_ : α) × β)} {k : α} {fallback : β} - {not_mem : (l.map Sigma.fst).contains k = false} - (h : m.1.WF) : + (h' : (l.map Sigma.fst).contains k = false) : Const.getD (m.insertMany l).1 k fallback = Const.getD m k fallback:= by - simp_to_model using getValueD_insertList_not_mem + simp_to_model using getValueD_insertList_of_contains_eq_false -theorem getD_insertMany_list_mem_const [EquivBEq α] [LawfulHashable α] - {l : List ((_ : α) × β)} {k k' : α} {k_eq: k == k'} {v fallback: β} {mem : ⟨k,v⟩ ∈ l} - {distinct : List.Pairwise (fun a b => (a.1 == b.1) = false ) l} (h : m.1.WF): +theorem getD_insertMany_list_of_mem_const [EquivBEq α] [LawfulHashable α] (h : m.1.WF) + {l : List ((_ : α) × β)} {k k' : α} (k_beq : k == k') {v fallback: β} (mem: ⟨k, v⟩ ∈ l) + (distinct : List.Pairwise (fun a b => (a.1 == b.1) = false ) l) : Const.getD (m.insertMany l).1 k' fallback = v := by - simp_to_model using getValueD_insertList_mem + simp_to_model using getValueD_insertList_of_mem namespace Const @[simp] -theorem insertMany_list_nil : +theorem insertMany_nil : insertMany m [] = m := by simp [insertMany, Id.run] @[simp] -theorem insertMany_list_singleton [EquivBEq α] [LawfulHashable α] {k : α} {v : β} : - insertMany m [⟨k,v⟩] = m.insert k v := by +theorem insertMany_list_singleton {k : α} {v : β} : + insertMany m [⟨k, v⟩] = m.insert k v := by simp [insertMany, Id.run] -theorem insertMany_list_cons {l : List (α × β)} {k : α} {v : β} : - (insertMany m (⟨k,v⟩ :: l)).1 = (insertMany (m.insert k v) l).1 := by +theorem insertMany_cons {l : List (α × β)} {k : α} {v : β} : + (insertMany m (⟨k, v⟩ :: l)).1 = (insertMany (m.insert k v) l).1 := by simp only [insertMany_eq_insertListₘ] cases l with | nil => simp [insertListₘ] | cons hd tl => simp [insertListₘ] -theorem contains_insertMany_list [EquivBEq α] [LawfulHashable α] {l : List (α × β)} {k : α} (h : m.1.WF) : +theorem contains_insertMany_list [EquivBEq α] [LawfulHashable α] (h : m.1.WF) + {l : List (α × β)} {k : α} : (Const.insertMany m l).1.contains k = (m.contains k || (l.map Prod.fst).contains k) := by - simp_to_model using List.containsKey_insertListConst + simp_to_model using containsKey_insertListConst -theorem contains_of_contains_insertMany_list [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l : List ( α × β )} {k : α} : +theorem contains_of_contains_insertMany_list [EquivBEq α] [LawfulHashable α] (h : m.1.WF) + {l : List ( α × β )} {k : α} : (insertMany m l).1.contains k → (l.map Prod.fst).contains k = false → m.contains k := by - simp_to_model using List.containsKey_of_containsKey_insertListConst + simp_to_model using containsKey_of_containsKey_insertListConst -theorem getKey?_insertMany_list_not_mem [EquivBEq α] [LawfulHashable α] (h : m.1.WF) +theorem getKey?_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l : List (α × β)} {k : α} - {not_mem : (l.map Prod.fst).contains k = false} : + (h' : (l.map Prod.fst).contains k = false) : (insertMany m l).1.getKey? k = m.getKey? k := by - simp_to_model using List.getKey?_insertListConst_not_mem + simp_to_model using getKey?_insertListConst_of_contains_eq_false -theorem getKey?_insertMany_list_mem [EquivBEq α] [LawfulHashable α] (h : m.1.WF) +theorem getKey?_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l : List (α × β)} - {k k' : α} {k_beq : k == k'} - {distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)} - {mem : k ∈ (l.map Prod.fst)} : + {k k' : α} (k_beq : k == k') + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) + (mem : k ∈ (l.map Prod.fst)) : (insertMany m l).1.getKey? k' = some k := by - simp_to_model using List.getKey?_insertListConst_mem + simp_to_model using getKey?_insertListConst_of_mem -theorem getKey_insertMany_list_not_mem [EquivBEq α] [LawfulHashable α] (h : m.1.WF) +theorem getKey_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l : List (α × β)} {k : α} - {not_mem : (l.map Prod.fst).contains k = false} + {h₁ : (l.map Prod.fst).contains k = false} {h'} : - (insertMany m l).1.getKey k h' = m.getKey k (contains_of_contains_insertMany_list _ h h' not_mem) := by - simp_to_model using List.getKey_insertListConst_not_mem + (insertMany m l).1.getKey k h' = + m.getKey k (contains_of_contains_insertMany_list _ h h' h₁) := by + simp_to_model using getKey_insertListConst_of_contains_eq_false -theorem getKey_insertMany_list_mem [EquivBEq α] [LawfulHashable α] (h : m.1.WF) +theorem getKey_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l : List (α × β)} - {k k' : α} {k_beq : k == k'} - {distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)} - {mem : k ∈ (l.map Prod.fst)} + {k k' : α} (k_beq : k == k') + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) + (mem : k ∈ (l.map Prod.fst)) {h'} : (insertMany m l).1.getKey k' h' = k := by - simp_to_model using List.getKey_insertListConst_mem + simp_to_model using getKey_insertListConst_of_mem -theorem getKey!_insertMany_list_not_mem [EquivBEq α] [LawfulHashable α] [Inhabited α] (h : m.1.WF) - {l : List (α × β)} {k : α} - {not_mem : (l.map Prod.fst).contains k = false} : +theorem getKey!_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] [Inhabited α] + (h : m.1.WF) {l : List (α × β)} {k : α} + (h' : (l.map Prod.fst).contains k = false) : (insertMany m l).1.getKey! k = m.getKey! k := by - simp_to_model using List.getKey!_insertListConst_not_mem + simp_to_model using getKey!_insertListConst_of_contains_eq_false -theorem getKey!_insertMany_list_mem [EquivBEq α] [LawfulHashable α] [Inhabited α] (h : m.1.WF) +theorem getKey!_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] [Inhabited α] (h : m.1.WF) {l : List (α × β)} - {k k' : α} {k_beq : k == k'} - {distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)} - {mem : k ∈ (l.map Prod.fst)} : + {k k' : α} (k_beq : k == k') + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) + (mem : k ∈ (l.map Prod.fst)) : (insertMany m l).1.getKey! k' = k := by - simp_to_model using List.getKey!_insertListConst_mem + simp_to_model using getKey!_insertListConst_of_mem -theorem getKeyD_insertMany_list_not_mem [EquivBEq α] [LawfulHashable α] (h : m.1.WF) +theorem getKeyD_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l : List (α × β)} {k fallback : α} - {not_mem : (l.map Prod.fst).contains k = false} : + (h' : (l.map Prod.fst).contains k = false) : (insertMany m l).1.getKeyD k fallback = m.getKeyD k fallback := by - simp_to_model using List.getKeyD_insertListConst_not_mem + simp_to_model using getKeyD_insertListConst_of_contains_eq_false -theorem getKeyD_insertMany_list_mem [EquivBEq α] [LawfulHashable α] (h : m.1.WF) +theorem getKeyD_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l : List (α × β)} - {k k' fallback : α} {k_beq : k == k'} - {distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)} - {mem : k ∈ (l.map Prod.fst)} : + {k k' fallback : α} (k_beq : k == k') + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) + (mem : k ∈ (l.map Prod.fst)) : (insertMany m l).1.getKeyD k' fallback = k := by - simp_to_model using List.getKeyD_insertListConst_mem + simp_to_model using getKeyD_insertListConst_of_mem -theorem size_insertMany_list [EquivBEq α] [LawfulHashable α] +theorem size_insertMany_list [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l : List (α × β)} - {distinct : List.Pairwise (fun a b => (a.1 == b.1) = false) l} - (h : m.1.WF) : - (∀ (a : α), ¬ (m.contains a = true ∧ (l.map Prod.fst).contains a = true)) → + {distinct : List.Pairwise (fun a b => (a.1 == b.1) = false) l} : + (∀ (a : α), ¬ (m.contains a ∧ (l.map Prod.fst).contains a)) → (insertMany m l).1.1.size = m.1.size + l.length := by - simp_to_model using List.insertListConst_length + simp_to_model using length_insertListConst -theorem insertMany_list_notEmpty_if_map_notEmpty [EquivBEq α] [LawfulHashable α] - {l : List (α × β)} (h : m.1.WF) : +theorem isEmpty_insertMany_list_eq_false_of_isEmpty_eq_false [EquivBEq α] [LawfulHashable α] + (h : m.1.WF) {l : List (α × β)} : (m.1.isEmpty = false) → (insertMany m l).1.1.isEmpty = false := by - simp_to_model using List.insertListConst_not_isEmpty_if_start_not_isEmpty + simp_to_model using isEmpty_insertListConst_eq_false_of_isEmpty_eq_false -theorem insertMany_list_isEmpty [EquivBEq α] [LawfulHashable α] - {l : List (α × β)} (h : m.1.WF) : +theorem isEmpty_insertMany_list [EquivBEq α] [LawfulHashable α] (h : m.1.WF) + {l : List (α × β)} : (insertMany m l).1.1.isEmpty ↔ m.1.isEmpty ∧ l.isEmpty := by - simp_to_model using List.insertListConst_isEmpty + simp_to_model using isEmpty_insertListConst -theorem get?_insertMany_list_not_mem_const [EquivBEq α] [LawfulHashable α] +theorem get?_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l : List (α × β)} {k : α} - {not_mem : (l.map Prod.fst).contains k = false} - (h : m.1.WF) : + (h' : (l.map Prod.fst).contains k = false) : get? (insertMany m l).1 k = get? m k := by - simp_to_model using getValue?_insertListConst_not_mem + simp_to_model using getValue?_insertListConst_of_contains_eq_false -theorem get?_insertMany_list_mem_const [EquivBEq α] [LawfulHashable α] - {l : List (α × β)} {k k' : α} {k_eq: k == k'} {v : β} {mem : ⟨k,v⟩ ∈ l} - {distinct : List.Pairwise (fun a b => (a.1 == b.1) = false ) l} (h : m.1.WF): +theorem get?_insertMany_list_of_mem_const [EquivBEq α] [LawfulHashable α] (h : m.1.WF) + {l : List (α × β)} {k k' : α} (k_beq : k == k') {v : β} (mem: ⟨k, v⟩ ∈ l) + (distinct : List.Pairwise (fun a b => (a.1 == b.1) = false ) l) : get? (insertMany m l).1 k' = v := by - simp_to_model using getValue?_insertListConst_mem + simp_to_model using getValue?_insertListConst_of_mem -theorem get_insertMany_list_not_mem_const [EquivBEq α] [LawfulHashable α] +theorem get_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l : List (α × β)} {k : α} - {not_mem : (l.map Prod.fst).contains k = false} - (h : m.1.WF) {h'}: - get (insertMany m l).1 k h' = get m k (contains_of_contains_insertMany_list _ h h' not_mem) := by - simp_to_model using getValue_insertListConst_not_mem - -theorem get_insertMany_list_mem_const [EquivBEq α] [LawfulHashable α] - {l : List (α × β)} {k k' : α} {k_eq: k == k'} {v : β} {mem : ⟨k,v⟩ ∈ l} - {distinct : List.Pairwise (fun a b => (a.1 == b.1) = false ) l} (h : m.1.WF) {h'}: + (h₁ : (l.map Prod.fst).contains k = false) + {h'}: + get (insertMany m l).1 k h' = get m k (contains_of_contains_insertMany_list _ h h' h₁) := by + simp_to_model using getValue_insertListConst_of_contains_eq_false + +theorem get_insertMany_list_of_mem_const [EquivBEq α] [LawfulHashable α] (h : m.1.WF) + {l : List (α × β)} {k k' : α} (k_beq : k == k') {v : β} (mem: ⟨k, v⟩ ∈ l) + (distinct : List.Pairwise (fun a b => (a.1 == b.1) = false ) l) {h'}: get (insertMany m l).1 k' h' = v := by - simp_to_model using getValue_insertListConst_mem + simp_to_model using getValue_insertListConst_of_mem -theorem get!_insertMany_list_not_mem_const [EquivBEq α] [LawfulHashable α] [Inhabited β] - {l : List (α × β)} {k : α} - {not_mem : (l.map Prod.fst).contains k = false} - (h : m.1.WF) : +theorem get!_insertMany_list_of_contains_eq_false_const [EquivBEq α] [LawfulHashable α] + [Inhabited β] (h : m.1.WF) {l : List (α × β)} {k : α} + (h' : (l.map Prod.fst).contains k = false) : get! (insertMany m l).1 k = get! m k := by - simp_to_model using getValue!_insertListConst_not_mem + simp_to_model using getValue!_insertListConst_of_contains_eq_false -theorem get!_insertMany_list_mem_const [EquivBEq α] [LawfulHashable α] [Inhabited β] - {l : List (α × β)} {k k' : α} {k_eq: k == k'} {v : β} {mem : ⟨k,v⟩ ∈ l} - {distinct : List.Pairwise (fun a b => (a.1 == b.1) = false ) l} (h : m.1.WF): +theorem get!_insertMany_list_of_mem_const [EquivBEq α] [LawfulHashable α] [Inhabited β] (h : m.1.WF) + {l : List (α × β)} {k k' : α} (k_beq : k == k') {v : β} (mem: ⟨k, v⟩ ∈ l) + (distinct : List.Pairwise (fun a b => (a.1 == b.1) = false ) l) : get! (insertMany m l).1 k' = v := by - simp_to_model using getValue!_insertListConst_mem + simp_to_model using getValue!_insertListConst_of_mem -theorem getD_insertMany_list_not_mem_const [EquivBEq α] [LawfulHashable α] +theorem getD_insertMany_list_of_contains_eq_false_const [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l : List (α × β)} {k : α} {fallback : β} - {not_mem : (l.map Prod.fst).contains k = false} - (h : m.1.WF) : + (h' : (l.map Prod.fst).contains k = false) : getD (insertMany m l).1 k fallback = getD m k fallback:= by - simp_to_model using getValueD_insertListConst_not_mem + simp_to_model using getValueD_insertListConst_of_contains_eq_false -theorem getD_insertMany_list_mem_const [EquivBEq α] [LawfulHashable α] - {l : List (α × β)} {k k' : α} {k_eq: k == k'} {v fallback: β} {mem : ⟨k,v⟩ ∈ l} - {distinct : List.Pairwise (fun a b => (a.1 == b.1) = false ) l} (h : m.1.WF): +theorem getD_insertMany_list_of_mem_const [EquivBEq α] [LawfulHashable α] (h : m.1.WF) + {l : List (α × β)} {k k' : α} (k_beq : k == k') {v fallback: β} (mem: ⟨k, v⟩ ∈ l) + (distinct : List.Pairwise (fun a b => (a.1 == b.1) = false ) l) : getD (insertMany m l).1 k' fallback = v := by - simp_to_model using getValueD_insertListConst_mem + simp_to_model using getValueD_insertListConst_of_mem variable (m: Raw₀ α (fun _ => Unit)) @[simp] -theorem insertManyIfNewUnit_list_nil : +theorem insertManyIfNewUnit_nil : insertManyIfNewUnit m [] = m := by simp [insertManyIfNewUnit, Id.run] @@ -1224,119 +1215,124 @@ theorem insertManyIfNewUnit_list_singleton [EquivBEq α] [LawfulHashable α] {k insertManyIfNewUnit m [k] = m.insertIfNew k () := by simp [insertManyIfNewUnit, Id.run] -theorem contains_insertManyIfNewUnit_list [EquivBEq α] [LawfulHashable α] {l : List α} {k : α} (h : m.1.WF): +theorem insertManyIfNewUnit_cons {l : List α} {k : α} : + (insertManyIfNewUnit m (k :: l)).1 = (insertManyIfNewUnit (m.insertIfNew k ()) l).1 := by + simp only [insertManyIfNewUnit_eq_insertListIfNewUnitₘ] + cases l with + | nil => simp [insertListIfNewUnitₘ] + | cons hd tl => simp [insertListIfNewUnitₘ] + +theorem contains_insertManyIfNewUnit_list [EquivBEq α] [LawfulHashable α] (h : m.1.WF) + {l : List α} {k : α}: (insertManyIfNewUnit m l).1.contains k = (m.contains k || l.contains k) := by simp_to_model using containsKey_insertListIfNewUnit -theorem contains_of_contains_insertManyIfNewUnit_list [EquivBEq α] [LawfulHashable α] {l : List α} {k : α} (h : m.1.WF) - (h₂ : l.contains k = false) : (insertManyIfNewUnit m l).1.contains k = true → m.contains k := by +theorem contains_of_contains_insertManyIfNewUnit_list [EquivBEq α] [LawfulHashable α] (h : m.1.WF) + {l : List α} {k : α} (h₂ : l.contains k = false) : + (insertManyIfNewUnit m l).1.contains k → m.contains k := by simp_to_model using containsKey_of_containsKey_insertListIfNewUnit -theorem getKey?_insertManyIfNewUnit_list_not_mem [EquivBEq α] [LawfulHashable α] - {l : List α} {k : α} - {not_mem : l.contains k = false} (h : m.1.WF) : +theorem getKey?_insertManyIfNewUnit_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] + (h : m.1.WF) {l : List α} {k : α} (h' : l.contains k = false) : getKey? (insertManyIfNewUnit m l).1 k = getKey? m k := by - simp_to_model using getKey?_insertListIfNewUnit_not_mem - -theorem getKey?_insertManyIfNewUnit_list_mem [EquivBEq α] [LawfulHashable α] - {l : List α} {k k' : α} {k_beq : k == k'} - {distinct : l.Pairwise (fun a b => (a == b) = false)} - {mem : k ∈ l} (h : m.1.WF) : m.contains k = false → - getKey? (insertManyIfNewUnit m l).1 k' = some k := by - simp_to_model using getKey?_insertListIfNewUnit_mem - -theorem getKey?_insertManyIfNewUnit_list_mem_both [EquivBEq α] [LawfulHashable α] - {l : List α} {k : α} - {distinct : l.Pairwise (fun a b => (a == b) = false)} - {mem : l.contains k} (h : m.1.WF): - m.contains k = true → getKey? (insertManyIfNewUnit m l).1 k = getKey? m k := by - simp_to_model using getKey?_insertListIfNewUnit_mem_both - -theorem getKey_insertManyIfNewUnit_list_not_mem [EquivBEq α] [LawfulHashable α] - {l : List α} {k : α} - {not_mem : l.contains k = false} - {h'} (h: m.1.WF): + simp_to_model using getKey?_insertListIfNewUnit_of_contains_eq_false + +theorem getKey?_insertManyIfNewUnit_list_of_mem_of_contains_eq_false [EquivBEq α] [LawfulHashable α] + (h : m.1.WF) {l : List α} {k k' : α} (k_beq : k == k') + (distinct : l.Pairwise (fun a b => (a == b) = false)) (mem : k ∈ l) : + m.contains k = false → getKey? (insertManyIfNewUnit m l).1 k' = some k := by + simp_to_model using getKey?_insertListIfNewUnit_of_mem_of_contains_eq_false + +theorem getKey?_insertManyIfNewUnit_list_of_mem_of_contains [EquivBEq α] [LawfulHashable α] + (h : m.1.WF) {l : List α} {k : α} + (distinct : l.Pairwise (fun a b => (a == b) = false)) + (h': l.contains k) : + m.contains k → getKey? (insertManyIfNewUnit m l).1 k = getKey? m k := by + simp_to_model using getKey?_insertListIfNewUnit_of_mem_of_contains + +theorem getKey_insertManyIfNewUnit_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] + (h : m.1.WF) {l : List α} {k : α} (h₁ : l.contains k = false) {h'} : getKey (insertManyIfNewUnit m l).1 k h' = - getKey m k (contains_of_contains_insertManyIfNewUnit_list m h not_mem h') := by - simp_to_model using getKey_insertListIfNewUnit_not_mem - -theorem getKey_insertManyIfNewUnit_list_mem [EquivBEq α] [LawfulHashable α] - {l : List α} - {k k' : α} {k_beq : k == k'} - {distinct : l.Pairwise (fun a b => (a == b) = false)} - {mem : k ∈ l}{h'} (h: m.1.WF): + getKey m k (contains_of_contains_insertManyIfNewUnit_list m h h₁ h') := by + simp_to_model using getKey_insertListIfNewUnit_of_contains_eq_false + +theorem getKey_insertManyIfNewUnit_list_of_mem_of_contains_eq_false [EquivBEq α] [LawfulHashable α] + (h : m.1.WF) {l : List α} + {k k' : α} (k_beq : k == k') + (distinct : l.Pairwise (fun a b => (a == b) = false)) + (mem : k ∈ l) {h'} : m.contains k = false → getKey (insertManyIfNewUnit m l).1 k' h' = k := by - simp_to_model using getKey_insertListIfNewUnit_mem + simp_to_model using getKey_insertListIfNewUnit_of_mem_of_contains_eq_false -theorem getKey_insertManyIfNewUnit_list_mem_both [EquivBEq α] [LawfulHashable α] - {l : List α} {k : α} - {distinct : l.Pairwise (fun a b => (a == b) = false)} - {mem : l.contains k} {mem' : m.contains k = true} {h'} (h : m.1.WF): +theorem getKey_insertManyIfNewUnit_list_mem_of_mem_of_contains [EquivBEq α] [LawfulHashable α] + (h : m.1.WF) {l : List α} {k : α} + (distinct : l.Pairwise (fun a b => (a == b) = false)) + (h': l.contains k) {mem' : m.contains k} {h'} : getKey (insertManyIfNewUnit m l).1 k h' = getKey m k mem' := by - simp_to_model using getKey_insertListIfNewUnit_mem_both + simp_to_model using getKey_insertListIfNewUnit_of_mem_of_contains -theorem getKey!_insertManyIfNewUnit_list_not_mem [EquivBEq α] [LawfulHashable α] [Inhabited α] - {l : List α} {k : α} - {not_mem : l.contains k = false} (h : m.1.WF) : +theorem getKey!_insertManyIfNewUnit_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] + [Inhabited α] (h : m.1.WF) {l : List α} {k : α} + (h' : l.contains k = false) : getKey! (insertManyIfNewUnit m l).1 k = getKey! m k := by - simp_to_model using getKey!_insertListIfNewUnit_not_mem + simp_to_model using getKey!_insertListIfNewUnit_of_contains_eq_false -theorem getKey!_insertManyIfNewUnit_list_mem [EquivBEq α] [LawfulHashable α] [Inhabited α] - {l : List α} {k k' : α} {k_beq : k == k'} - {distinct : l.Pairwise (fun a b => (a == b) = false)} - {mem : k ∈ l} (h : m.1.WF) : contains m k = false → +theorem getKey!_insertManyIfNewUnit_list_of_mem_of_contains_eq_false [EquivBEq α] [LawfulHashable α] + [Inhabited α] (h : m.1.WF) {l : List α} {k k' : α} (k_beq : k == k') + (distinct : l.Pairwise (fun a b => (a == b) = false)) + (mem : k ∈ l) : contains m k = false → getKey! (insertManyIfNewUnit m l).1 k' = k := by - simp_to_model using getKey!_insertListIfNewUnit_mem + simp_to_model using getKey!_insertListIfNewUnit_of_mem_of_contains_eq_false -theorem getKey!_insertManyIfNewUnit_list_mem_both [EquivBEq α] [LawfulHashable α] [Inhabited α] - {l : List α} {k : α} - {distinct : l.Pairwise (fun a b => (a == b) = false)} - {mem : l.contains k} (h : m.1.WF) : m.contains k = true → +theorem getKey!_insertManyIfNewUnit_list_mem_of_mem_of_contains [EquivBEq α] [LawfulHashable α] + [Inhabited α] (h : m.1.WF) {l : List α} {k : α} + (distinct : l.Pairwise (fun a b => (a == b) = false)) + (h': l.contains k) : m.contains k → getKey! (insertManyIfNewUnit m l).1 k = getKey! m k := by - simp_to_model using getKey!_insertListIfNewUnit_mem_both + simp_to_model using getKey!_insertListIfNewUnit_of_mem_of_contains -theorem getKeyD_insertManyIfNewUnit_list_not_mem [EquivBEq α] [LawfulHashable α] - {l : List α} {k fallback : α} - {not_mem : l.contains k = false} (h : m.1.WF): +theorem getKeyD_insertManyIfNewUnit_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] + (h : m.1.WF) {l : List α} {k fallback : α} + (h' : l.contains k = false): getKeyD (insertManyIfNewUnit m l).1 k fallback = getKeyD m k fallback := by - simp_to_model using getKeyD_insertListIfNewUnit_not_mem + simp_to_model using getKeyD_insertListIfNewUnit_of_contains_eq_false -theorem getKeyD_insertManyIfNewUnit_list_mem [EquivBEq α] [LawfulHashable α] - {l : List α} {k k' fallback : α} {k_beq : k == k'} - {distinct : l.Pairwise (fun a b => (a == b) = false)} - {mem : k ∈ l } (h : m.1.WF) : m.contains k = false → +theorem getKeyD_insertManyIfNewUnit_list_of_mem_of_contains_eq_false [EquivBEq α] [LawfulHashable α] + (h : m.1.WF) {l : List α} {k k' fallback : α} (k_beq : k == k') + (distinct : l.Pairwise (fun a b => (a == b) = false)) + {mem : k ∈ l } : m.contains k = false → getKeyD (insertManyIfNewUnit m l).1 k' fallback = k := by - simp_to_model using getKeyD_insertListIfNewUnit_mem + simp_to_model using getKeyD_insertListIfNewUnit_of_mem_of_contains_eq_false -theorem getKeyD_insertManyIfNewUnit_list_mem_both [EquivBEq α] [LawfulHashable α] - {l : List α} {k fallback : α} - {distinct : l.Pairwise (fun a b => (a == b) = false)} - {mem : l.contains k} (h : m.1.WF) : m.contains k = true → +theorem getKeyD_insertManyIfNewUnit_list_of_mem_of_contains [EquivBEq α] [LawfulHashable α] + (h : m.1.WF) {l : List α} {k fallback : α} + (distinct : l.Pairwise (fun a b => (a == b) = false)) + (h': l.contains k) : m.contains k → getKeyD (insertManyIfNewUnit m l).1 k fallback = getKeyD m k fallback := by - simp_to_model using getKeyD_insertListIfNewUnit_mem_both + simp_to_model using getKeyD_insertListIfNewUnit_of_mem_of_contains -theorem insertManyIfNewUnit_list_length [EquivBEq α] [LawfulHashable α] +theorem size_insertManyIfNewUnit_list [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l : List α} - (distinct : List.Pairwise (fun a b => (a == b) = false) l) - (h : m.1.WF) : (∀ (a : α), ¬ (m.contains a = true ∧ l.contains a = true)) → + (distinct : List.Pairwise (fun a b => (a == b) = false) l): + (∀ (a : α), ¬ (m.contains a ∧ l.contains a)) → (insertManyIfNewUnit m l).1.1.size = m.1.size + l.length := by - simp_to_model using insertListIfNewUnit_length + simp_to_model using length_insertListIfNewUnit -theorem insertManyIfNewUnit_list_not_isEmpty_if_start_not_isEmpty [EquivBEq α] [LawfulHashable α] - {l : List α} (h : m.1.WF) : m.1.isEmpty = false → +theorem isEmpty_insertManyIfNewUnit_list_eq_false_of_isEmpty_eq_false [EquivBEq α] [LawfulHashable α] + (h : m.1.WF) {l : List α} : m.1.isEmpty = false → (insertManyIfNewUnit m l).1.1.isEmpty = false := by - simp_to_model using insertListIfNewUnit_not_isEmpty_if_start_not_isEmpty + simp_to_model using isEmpty_insertListIfNewUnit_eq_false_of_isEmpty_eq_false -theorem insertManyIfNewUnit_list_isEmpty [EquivBEq α] [LawfulHashable α] - {l : List α} (h : m.1.WF) : +theorem isEmpty_insertManyIfNewUnit_list [EquivBEq α] [LawfulHashable α] (h : m.1.WF) + {l : List α} : (insertManyIfNewUnit m l).1.1.isEmpty ↔ m.1.isEmpty ∧ l.isEmpty := by - simp_to_model using insertListIfNewUnit_isEmpty + simp_to_model using isEmpty_insertListIfNewUnit -theorem get?_insertManyIfNewUnit_list [EquivBEq α] [LawfulHashable α] - {l : List α} {k : α} (h : m.1.WF) : +theorem get?_insertManyIfNewUnit_list [EquivBEq α] [LawfulHashable α] (h : m.1.WF) + {l : List α} {k : α} : get? (insertManyIfNewUnit m l).1 k = - if ( m.contains k || l.contains k) then some () else none := by + if m.contains k ∨ l.contains k then some () else none := by simp_to_model using getValue?_insertListIfNewUnit theorem getValue_insertManyIfNewUnit_list diff --git a/src/Std/Data/DHashMap/Internal/WF.lean b/src/Std/Data/DHashMap/Internal/WF.lean index e70b20fda488..97b135282e0e 100644 --- a/src/Std/Data/DHashMap/Internal/WF.lean +++ b/src/Std/Data/DHashMap/Internal/WF.lean @@ -707,7 +707,9 @@ theorem wfImp_filter [BEq α] [Hashable α] [EquivBEq α] [LawfulHashable α] {m exact wfImp_filterₘ h /-! # `insertListₘ` -/ -theorem wfImp_insertListₘ [BEq α] [Hashable α] [EquivBEq α][LawfulHashable α] {m : Raw₀ α β} {l : List ((a : α) × β a)} (h : Raw.WFImp m.1) : Raw.WFImp (m.insertListₘ l).1 := by + +theorem wfImp_insertListₘ [BEq α] [Hashable α] [EquivBEq α][LawfulHashable α] {m : Raw₀ α β} + {l : List ((a : α) × β a)} (h : Raw.WFImp m.1) : Raw.WFImp (m.insertListₘ l).1 := by induction l generalizing m with | nil => simp only [insertListₘ] @@ -718,17 +720,18 @@ theorem wfImp_insertListₘ [BEq α] [Hashable α] [EquivBEq α][LawfulHashable apply wfImp_insert h theorem toListModel_insertListₘ [BEq α] [Hashable α] [EquivBEq α][LawfulHashable α] {m : Raw₀ α β} {l : List ((a : α) × β a)} (h : Raw.WFImp m.1) : - Perm (toListModel (insertListₘ m l).1.buckets) (List.insertList (toListModel m.1.buckets) l) := by -induction l generalizing m with -| nil => - simp[insertListₘ, List.insertList] -| cons hd tl ih => - simp only [insertListₘ, List.insertList] - apply Perm.trans - apply ih (wfImp_insert h) - apply List.insertList_perm_of_perm_first - apply toListModel_insert h - apply (wfImp_insert h).distinct + Perm (toListModel (insertListₘ m l).1.buckets) (List.insertList (toListModel m.1.buckets) l) := by + induction l generalizing m with + | nil => + simp[insertListₘ, List.insertList] + | cons hd tl ih => + simp only [insertListₘ, List.insertList] + apply Perm.trans + apply ih (wfImp_insert h) + apply List.insertList_perm_of_perm_first + apply toListModel_insert h + apply (wfImp_insert h).distinct + end Raw₀ namespace Raw @@ -765,7 +768,9 @@ theorem toListModel_insertMany_list [BEq α] [Hashable α] [EquivBEq α] [Lawful exact h /-! # `Const.insertListₘ` -/ -theorem Const.toListModel_insertListₘ {β : Type v} [BEq α] [Hashable α] [EquivBEq α] [LawfulHashable α] {m : Raw₀ α (fun _ => β)}{l : List (α × β)} (h : Raw.WFImp m.1): + +theorem Const.toListModel_insertListₘ {β : Type v} [BEq α] [Hashable α] [EquivBEq α] + [LawfulHashable α] {m : Raw₀ α (fun _ => β)}{l : List (α × β)} (h : Raw.WFImp m.1): Perm (toListModel (Const.insertListₘ m l).1.buckets) (insertListConst (toListModel m.1.buckets) l) := by induction l generalizing m with | nil => simp [Const.insertListₘ, insertListConst] @@ -778,7 +783,9 @@ theorem Const.toListModel_insertListₘ {β : Type v} [BEq α] [Hashable α] [Eq apply (wfImp_insert h).distinct /-! # `Const.insertMany` -/ -theorem Const.toListModel_insertMany_list {β : Type v} [BEq α] [Hashable α] [EquivBEq α] [LawfulHashable α] {m : Raw₀ α (fun _ => β)} {l : List (α × β)} (h : Raw.WFImp m.1): + +theorem Const.toListModel_insertMany_list {β : Type v} [BEq α] [Hashable α] [EquivBEq α] + [LawfulHashable α] {m : Raw₀ α (fun _ => β)} {l : List (α × β)} (h : Raw.WFImp m.1): Perm (toListModel (Const.insertMany m l).1.1.buckets) (insertListConst (toListModel m.1.buckets) l) := by rw [Const.insertMany_eq_insertListₘ] apply toListModel_insertListₘ h @@ -789,8 +796,11 @@ theorem Const.wfImp_insertMany {β : Type v} [BEq α] [Hashable α] [EquivBEq α Raw.WF.out ((Const.insertMany m l).2 _ Raw.WF.insert₀ (.wf m.2 h)) /-! # `Const.insertListIfNewUnitₘ` -/ -theorem Const.toListModel_insertListIfNewUnitₘ [BEq α] [Hashable α] [EquivBEq α] [LawfulHashable α] {m : Raw₀ α (fun _ => Unit)} {l : List α} (h : Raw.WFImp m.1): - Perm (toListModel (Const.insertListIfNewUnitₘ m l).1.buckets) (List.insertListIfNewUnit (toListModel m.1.buckets) l) := by + +theorem Const.toListModel_insertListIfNewUnitₘ [BEq α] [Hashable α] [EquivBEq α] [LawfulHashable α] + {m : Raw₀ α (fun _ => Unit)} {l : List α} (h : Raw.WFImp m.1): + Perm (toListModel (Const.insertListIfNewUnitₘ m l).1.buckets) + (List.insertListIfNewUnit (toListModel m.1.buckets) l) := by induction l generalizing m with | nil => simp [insertListIfNewUnitₘ, List.insertListIfNewUnit] | cons hd tl ih => @@ -802,13 +812,16 @@ theorem Const.toListModel_insertListIfNewUnitₘ [BEq α] [Hashable α] [EquivBE apply (wfImp_insertIfNew h).distinct /-! # `Const.insertManyIfNewUnit` -/ -theorem Const.toListModel_insertManyIfNewUnit_list [BEq α] [Hashable α] [EquivBEq α] [LawfulHashable α] {m : Raw₀ α (fun _ => Unit)} {l : List α} (h : Raw.WFImp m.1): - Perm (toListModel (Const.insertManyIfNewUnit m l).1.1.buckets) (List.insertListIfNewUnit (toListModel m.1.buckets) l) := by - rw [Const.insertManyIfNewUnit_eq_insertListIfNewUnit] + +theorem Const.toListModel_insertManyIfNewUnit_list [BEq α] [Hashable α] [EquivBEq α] + [LawfulHashable α] {m : Raw₀ α (fun _ => Unit)} {l : List α} (h : Raw.WFImp m.1): + Perm (toListModel (Const.insertManyIfNewUnit m l).1.1.buckets) + (List.insertListIfNewUnit (toListModel m.1.buckets) l) := by + rw [Const.insertManyIfNewUnit_eq_insertListIfNewUnitₘ] apply toListModel_insertListIfNewUnitₘ h -theorem Const.wfImp_insertManyIfNewUnit [BEq α] [Hashable α] [EquivBEq α] [LawfulHashable α] {ρ : Type w} - [ForIn Id ρ α] {m : Raw₀ α (fun _ => Unit)} {l : ρ} (h : Raw.WFImp m.1) : +theorem Const.wfImp_insertManyIfNewUnit [BEq α] [Hashable α] [EquivBEq α] [LawfulHashable α] + {ρ : Type w} [ForIn Id ρ α] {m : Raw₀ α (fun _ => Unit)} {l : ρ} (h : Raw.WFImp m.1) : Raw.WFImp (Const.insertManyIfNewUnit m l).1.1 := Raw.WF.out ((Const.insertManyIfNewUnit m l).2 _ Raw.WF.insertIfNew₀ (.wf m.2 h)) diff --git a/src/Std/Data/DHashMap/Raw.lean b/src/Std/Data/DHashMap/Raw.lean index 0f92a8c194ba..967dc4797d5b 100644 --- a/src/Std/Data/DHashMap/Raw.lean +++ b/src/Std/Data/DHashMap/Raw.lean @@ -389,7 +389,6 @@ appearance. (Raw₀.insertMany ⟨m, h⟩ l).1 else m -- will never happen for well-formed inputs - @[inline, inherit_doc Raw.insertMany] def Const.insertMany {β : Type v} [BEq α] [Hashable α] {ρ : Type w} [ForIn Id ρ (α × β)] (m : Raw α (fun _ => β)) (l : ρ) : Raw α (fun _ => β) := if h : 0 < m.buckets.size then From bddf6fbf450075bc128d350ff20960b6105b708f Mon Sep 17 00:00:00 2001 From: monsterkrampe Date: Tue, 3 Dec 2024 16:15:57 +0100 Subject: [PATCH 34/83] Adjust simp spacing --- src/Std/Data/DHashMap/Internal/Model.lean | 6 +++--- src/Std/Data/DHashMap/Internal/Raw.lean | 4 ++-- src/Std/Data/DHashMap/Internal/WF.lean | 2 +- 3 files changed, 6 insertions(+), 6 deletions(-) diff --git a/src/Std/Data/DHashMap/Internal/Model.lean b/src/Std/Data/DHashMap/Internal/Model.lean index 8b7eb32018a0..713e39efc098 100644 --- a/src/Std/Data/DHashMap/Internal/Model.lean +++ b/src/Std/Data/DHashMap/Internal/Model.lean @@ -480,7 +480,7 @@ theorem insertMany_eq_insertListₘ [BEq α] [Hashable α](m : Raw₀ α β) (l t.val.insertListₘ l from this _ intro t induction l generalizing m with - | nil => simp[insertListₘ] + | nil => simp [insertListₘ] | cons hd tl ih => simp only [List.foldl_cons,insertListₘ] apply ih @@ -524,7 +524,7 @@ theorem Const.insertMany_eq_insertListₘ [BEq α] [Hashable α] (m : Raw₀ α Const.insertListₘ t.val l from this _ intro t induction l generalizing m with - | nil => simp[insertListₘ] + | nil => simp [insertListₘ] | cons hd tl ih => simp only [List.foldl_cons,insertListₘ] apply ih @@ -538,7 +538,7 @@ theorem Const.insertManyIfNewUnit_eq_insertListIfNewUnitₘ [BEq α] [Hashable Const.insertListIfNewUnitₘ t.val l from this _ intro t induction l generalizing m with - | nil => simp[insertListIfNewUnitₘ] + | nil => simp [insertListIfNewUnitₘ] | cons hd tl ih => simp only [List.foldl_cons,insertListIfNewUnitₘ] apply ih diff --git a/src/Std/Data/DHashMap/Internal/Raw.lean b/src/Std/Data/DHashMap/Internal/Raw.lean index 4fcbb00d42a2..a735d30d7dd3 100644 --- a/src/Std/Data/DHashMap/Internal/Raw.lean +++ b/src/Std/Data/DHashMap/Internal/Raw.lean @@ -201,11 +201,11 @@ theorem filter_val [BEq α] [Hashable α] {m : Raw₀ α β} {f : (a : α) → theorem insertMany_eq [BEq α] [Hashable α] {m : Raw α β} (h : m.WF) {ρ : Type w} [ForIn Id ρ ((a : α) × β a)] {l : ρ} : m.insertMany l = Raw₀.insertMany ⟨m, h.size_buckets_pos⟩ l := by - simp[Raw.insertMany, h.size_buckets_pos] + simp [Raw.insertMany, h.size_buckets_pos] theorem insertMany_val [BEq α][Hashable α] {m : Raw₀ α β} {ρ : Type w} [ForIn Id ρ ((a : α) × β a)] {l : ρ} : m.val.insertMany l = m.insertMany l := by - simp[Raw.insertMany, m.2] + simp [Raw.insertMany, m.2] section diff --git a/src/Std/Data/DHashMap/Internal/WF.lean b/src/Std/Data/DHashMap/Internal/WF.lean index 97b135282e0e..7c199d9a3a09 100644 --- a/src/Std/Data/DHashMap/Internal/WF.lean +++ b/src/Std/Data/DHashMap/Internal/WF.lean @@ -723,7 +723,7 @@ theorem toListModel_insertListₘ [BEq α] [Hashable α] [EquivBEq α][LawfulHas Perm (toListModel (insertListₘ m l).1.buckets) (List.insertList (toListModel m.1.buckets) l) := by induction l generalizing m with | nil => - simp[insertListₘ, List.insertList] + simp [insertListₘ, List.insertList] | cons hd tl ih => simp only [insertListₘ, List.insertList] apply Perm.trans From ae25818a695f7b63cb9aca47342afe03aff3885f Mon Sep 17 00:00:00 2001 From: monsterkrampe Date: Tue, 3 Dec 2024 16:27:23 +0100 Subject: [PATCH 35/83] Incorporate smaller remarks from PR --- src/Std/Data/DHashMap/Internal/List/Associative.lean | 2 +- src/Std/Data/DHashMap/Internal/RawLemmas.lean | 6 +++--- src/Std/Data/DHashMap/Internal/WF.lean | 11 ----------- 3 files changed, 4 insertions(+), 15 deletions(-) diff --git a/src/Std/Data/DHashMap/Internal/List/Associative.lean b/src/Std/Data/DHashMap/Internal/List/Associative.lean index d52543f3d951..217cabfc61ab 100644 --- a/src/Std/Data/DHashMap/Internal/List/Associative.lean +++ b/src/Std/Data/DHashMap/Internal/List/Associative.lean @@ -2750,7 +2750,7 @@ theorem getValue_insertListConst_of_mem [BEq α] [PartialEquivBEq α] . exact distinct /-- Internal implementation detail of the hash map -/ -def insertListIfNewUnit [BEq α] (l: List ((_ : α) × Unit))(toInsert: List α) : +def insertListIfNewUnit [BEq α] (l: List ((_ : α) × Unit)) (toInsert: List α) : List ((_ : α) × Unit) := match toInsert with | .nil => l diff --git a/src/Std/Data/DHashMap/Internal/RawLemmas.lean b/src/Std/Data/DHashMap/Internal/RawLemmas.lean index cedc1b443c4a..d12cb9d954ad 100644 --- a/src/Std/Data/DHashMap/Internal/RawLemmas.lean +++ b/src/Std/Data/DHashMap/Internal/RawLemmas.lean @@ -984,14 +984,14 @@ theorem getKeyD_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] (h : m. simp_to_model using getKeyD_insertList_of_mem theorem size_insertMany_list [EquivBEq α] [LawfulHashable α] (h : m.1.WF) - {l : List ((a : α) × β a)} {distinct : List.Pairwise (fun a b => (a.1 == b.1) = false) l}: + {l : List ((a : α) × β a)} {distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)} : (∀ (a : α), ¬ (m.contains a ∧ (l.map Sigma.fst).contains a)) → (m.insertMany l).1.1.size = m.1.size + l.length := by simp_to_model using length_insertList theorem isEmpty_insertMany_list_eq_false_of_isEmpty_eq_false [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l : List ((a : α) × β a)}: - (m.1.isEmpty = false) → (m.insertMany l).1.1.isEmpty = false := by + (m.1.isEmpty = false) → (m.insertMany l).1.1.isEmpty = false := by simp_to_model using isEmpty_insertList_eq_false_of_isEmpty_eq_false theorem isEmpty_insertMany_list [EquivBEq α] [LawfulHashable α] (h : m.1.WF) @@ -1000,7 +1000,7 @@ theorem isEmpty_insertMany_list [EquivBEq α] [LawfulHashable α] (h : m.1.WF) simp_to_model using isEmpty_insertList section -variable {β: Type v} (m: Raw₀ α (fun _ => β)) +variable {β : Type v} (m : Raw₀ α (fun _ => β)) theorem get?_insertMany_list_of_contains_eq_false_const [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l : List ((_ : α) × β)} {k : α} (h : (l.map Sigma.fst).contains k = false) : diff --git a/src/Std/Data/DHashMap/Internal/WF.lean b/src/Std/Data/DHashMap/Internal/WF.lean index 7c199d9a3a09..95bde8df4309 100644 --- a/src/Std/Data/DHashMap/Internal/WF.lean +++ b/src/Std/Data/DHashMap/Internal/WF.lean @@ -708,17 +708,6 @@ theorem wfImp_filter [BEq α] [Hashable α] [EquivBEq α] [LawfulHashable α] {m /-! # `insertListₘ` -/ -theorem wfImp_insertListₘ [BEq α] [Hashable α] [EquivBEq α][LawfulHashable α] {m : Raw₀ α β} - {l : List ((a : α) × β a)} (h : Raw.WFImp m.1) : Raw.WFImp (m.insertListₘ l).1 := by - induction l generalizing m with - | nil => - simp only [insertListₘ] - exact h - | cons hd tl ih => - simp only [insertListₘ] - apply ih - apply wfImp_insert h - theorem toListModel_insertListₘ [BEq α] [Hashable α] [EquivBEq α][LawfulHashable α] {m : Raw₀ α β} {l : List ((a : α) × β a)} (h : Raw.WFImp m.1) : Perm (toListModel (insertListₘ m l).1.buckets) (List.insertList (toListModel m.1.buckets) l) := by induction l generalizing m with From 9bb3511e6cd0b30f565f5c02b093528c44135137 Mon Sep 17 00:00:00 2001 From: monsterkrampe Date: Tue, 3 Dec 2024 16:34:00 +0100 Subject: [PATCH 36/83] Shorten List.pairwise statements --- .../DHashMap/Internal/List/Associative.lean | 24 +++++++++---------- src/Std/Data/DHashMap/Internal/RawLemmas.lean | 24 +++++++++---------- 2 files changed, 24 insertions(+), 24 deletions(-) diff --git a/src/Std/Data/DHashMap/Internal/List/Associative.lean b/src/Std/Data/DHashMap/Internal/List/Associative.lean index 217cabfc61ab..9cf1beab8b9f 100644 --- a/src/Std/Data/DHashMap/Internal/List/Associative.lean +++ b/src/Std/Data/DHashMap/Internal/List/Associative.lean @@ -175,7 +175,7 @@ theorem getValueCast?_cons_self [BEq α] [LawfulBEq α] {l : List ((a : α) × theorem getValueCast?_of_mem [BEq α] [LawfulBEq α] {l : List ((a : α) × β a)} {k : α} (k' : α ) (v : β k) (k_beq : k == k') (mem : ⟨k, v⟩ ∈ l) - (distinct : List.Pairwise (fun a b => (a.1 == b.1) = false) l) : + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) : getValueCast? k' l = some (cast (by congr; apply LawfulBEq.eq_of_beq k_beq) v) := by induction l with | nil => simp at mem @@ -928,7 +928,7 @@ theorem DistinctKeys.nil [BEq α] : DistinctKeys ([] : List ((a : α) × β a)) ⟨by simp⟩ theorem DistinctKeys.def [BEq α] {l : List ((a : α) × β a)} : - DistinctKeys l ↔ List.Pairwise (fun a b => (a.1 == b.1) = false) l := + DistinctKeys l ↔ l.Pairwise (fun a b => (a.1 == b.1) = false) := ⟨fun h => by simpa [keys_eq_map, List.pairwise_map] using h.distinct, fun h => ⟨by simpa [keys_eq_map, List.pairwise_map] using h⟩⟩ @@ -2218,7 +2218,7 @@ theorem getKeyD_insertList_of_mem [BEq α] [EquivBEq α] theorem perm_insertList [BEq α] [ReflBEq α] [PartialEquivBEq α] {l toInsert : List ((a : α) × β a)} (distinct_l : DistinctKeys l) - (distinct_toInsert : List.Pairwise (fun a b => (a.1 == b.1) = false) toInsert) + (distinct_toInsert : toInsert.Pairwise (fun a b => (a.1 == b.1) = false)) (distinct_both : ∀ (a : α), ¬ (containsKey a l ∧ containsKey a toInsert)) : Perm (insertList l toInsert) (l ++ toInsert) := by rw [← DistinctKeys.def] at distinct_toInsert @@ -2263,7 +2263,7 @@ theorem perm_insertList [BEq α] [ReflBEq α] [PartialEquivBEq α] theorem length_insertList [BEq α] [EquivBEq α] {l toInsert : List ((a : α) × β a)} (distinct_l : DistinctKeys l) - (distinct_toInsert : List.Pairwise (fun a b => (a.1 == b.1) = false) toInsert) + (distinct_toInsert : toInsert.Pairwise (fun a b => (a.1 == b.1) = false)) (distinct_both : ∀ (a : α), ¬ (containsKey a l ∧ (toInsert.map Sigma.fst).contains a)) : (insertList l toInsert).length = l.length + toInsert.length := by rw [← DistinctKeys.def] at distinct_toInsert @@ -2350,7 +2350,7 @@ theorem getValue?_insertList_of_contains_eq_false [BEq α] [PartialEquivBEq α] theorem getValue?_insertList_of_mem [BEq α] [PartialEquivBEq α] {l toInsert : List ((_ : α) × β )} {k k' : α} (k_beq : k == k') {v : β} - (distinct : List.Pairwise (fun a b => (a.1 == b.1) = false) toInsert) + (distinct : toInsert.Pairwise (fun a b => (a.1 == b.1) = false)) (mem : ⟨k, v⟩ ∈ toInsert) : getValue? k' (insertList l toInsert) = v := by induction toInsert generalizing l with @@ -2394,7 +2394,7 @@ theorem getValue!_insertList_of_contains_eq_false [BEq α] [PartialEquivBEq α] theorem getValue!_insertList_of_mem [BEq α] [PartialEquivBEq α] [Inhabited β] {l toInsert : List ((_ : α) × β)} {k k' : α} {v: β} (k_beq : k == k') - (distinct : List.Pairwise (fun a b => (a.1 == b.1) = false) toInsert) (mem : ⟨k, v⟩ ∈ toInsert): + (distinct : toInsert.Pairwise (fun a b => (a.1 == b.1) = false)) (mem : ⟨k, v⟩ ∈ toInsert): getValue! k' (insertList l toInsert) = v := by simp only [getValue!_eq_getValue?] rw [getValue?_insertList_of_mem (mem:= mem)] @@ -2412,7 +2412,7 @@ theorem getValueD_insertList_of_contains_eq_false [BEq α] [PartialEquivBEq α] theorem getValueD_insertList_of_mem [BEq α] [PartialEquivBEq α] {l toInsert : List ((_ : α) × β)} {k k' : α} {v fallback: β} (k_beq : k == k') - (distinct : List.Pairwise (fun a b => (a.1 == b.1) = false) toInsert) (mem : ⟨k, v⟩ ∈ toInsert): + (distinct : toInsert.Pairwise (fun a b => (a.1 == b.1) = false)) (mem : ⟨k, v⟩ ∈ toInsert): getValueD k' (insertList l toInsert) fallback= v := by simp only [getValueD_eq_getValue?] rw [getValue?_insertList_of_mem (mem:= mem)] @@ -2627,7 +2627,7 @@ theorem getKeyD_insertListConst_of_mem [BEq α] [EquivBEq α] theorem length_insertListConst [BEq α] [EquivBEq α] {l : List ((_ : α) × β)} {toInsert : List (α × β)} (distinct_l : DistinctKeys l) - (distinct_toInsert : List.Pairwise (fun a b => (a.1 == b.1) = false) toInsert) + (distinct_toInsert : toInsert.Pairwise (fun a b => (a.1 == b.1) = false)) (distinct_both : ∀ (a : α), ¬ (containsKey a l ∧ (toInsert.map Prod.fst).contains a)) : (insertListConst l toInsert).length = l.length + toInsert.length := by rw [insertListConst_eq_insertList] @@ -2677,7 +2677,7 @@ theorem getValue?_insertListConst_of_contains_eq_false [BEq α] [PartialEquivBEq theorem getValue?_insertListConst_of_mem [BEq α] [PartialEquivBEq α] {l : List ((_ : α) × β)} {toInsert : List (α × β)} {k k' : α} (k_beq : k == k') {v : β} - (distinct : List.Pairwise (fun a b => (a.1 == b.1) = false) toInsert) + (distinct : toInsert.Pairwise (fun a b => (a.1 == b.1) = false)) (mem : ⟨k, v⟩ ∈ toInsert) : getValue? k' (insertListConst l toInsert) = v := by rw [insertListConst_eq_insertList] @@ -2698,7 +2698,7 @@ theorem getValue!_insertListConst_of_contains_eq_false [BEq α] [PartialEquivBEq theorem getValue!_insertListConst_of_mem [BEq α] [PartialEquivBEq α] [Inhabited β] {l : List ((_ : α) × β)} {toInsert : List (α × β)} {k k' : α} {v: β} (k_beq : k == k') - (distinct : List.Pairwise (fun a b => (a.1 == b.1) = false) toInsert) (mem : ⟨k, v⟩ ∈ toInsert): + (distinct : toInsert.Pairwise (fun a b => (a.1 == b.1) = false)) (mem : ⟨k, v⟩ ∈ toInsert): getValue! k' (insertListConst l toInsert) = v := by rw [insertListConst_eq_insertList, getValue!_insertList_of_mem (k_beq:=k_beq)] · apply pairwise_list_conversion_insertListConst distinct @@ -2716,7 +2716,7 @@ theorem getValueD_insertListConst_of_contains_eq_false [BEq α] [PartialEquivBEq theorem getValueD_insertListConst_of_mem [BEq α] [PartialEquivBEq α] {l : List ((_ : α) × β)} {toInsert : List (α × β)} {k k' : α} {v fallback: β} (k_beq : k == k') - (distinct : List.Pairwise (fun a b => (a.1 == b.1) = false) toInsert) (mem : ⟨k, v⟩ ∈ toInsert): + (distinct : toInsert.Pairwise (fun a b => (a.1 == b.1) = false)) (mem : ⟨k, v⟩ ∈ toInsert): getValueD k' (insertListConst l toInsert) fallback= v := by rw [insertListConst_eq_insertList, getValueD_insertList_of_mem (k_beq:= k_beq)] · apply pairwise_list_conversion_insertListConst distinct @@ -3034,7 +3034,7 @@ theorem getKeyD_insertListIfNewUnit_of_mem_of_contains [BEq α] [EquivBEq α] theorem length_insertListIfNewUnit [BEq α] [EquivBEq α] {l : List ((_ : α) × Unit)} {toInsert : List α} (distinct_l : DistinctKeys l) - (distinct_toInsert : List.Pairwise (fun a b => (a == b) = false) toInsert) + (distinct_toInsert : toInsert.Pairwise (fun a b => (a == b) = false)) (distinct_both : ∀ (a : α), ¬ (containsKey a l ∧ toInsert.contains a)) : (insertListIfNewUnit l toInsert).length = l.length + toInsert.length := by induction toInsert generalizing l with diff --git a/src/Std/Data/DHashMap/Internal/RawLemmas.lean b/src/Std/Data/DHashMap/Internal/RawLemmas.lean index d12cb9d954ad..474cf4c8aa96 100644 --- a/src/Std/Data/DHashMap/Internal/RawLemmas.lean +++ b/src/Std/Data/DHashMap/Internal/RawLemmas.lean @@ -1009,7 +1009,7 @@ theorem get?_insertMany_list_of_contains_eq_false_const [EquivBEq α] [LawfulHas theorem get?_insertMany_list_of_mem_const [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l : List ((_ : α) × β)} {k k' : α} (k_beq : k == k') {v : β} (mem : ⟨k, v⟩ ∈ l) - (distinct : List.Pairwise (fun a b => (a.1 == b.1) = false ) l) : + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false )) : Const.get? (m.insertMany l).1 k' = v := by simp_to_model using getValue?_insertList_of_mem @@ -1022,7 +1022,7 @@ theorem get_insertMany_list_of_contains_eq_false_const [EquivBEq α] [LawfulHash theorem get_insertMany_list_of_mem_const [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l : List ((_ : α) × β)} {k k' : α} (k_beq : k == k') {v : β} (mem : ⟨k, v⟩ ∈ l) - (distinct : List.Pairwise (fun a b => (a.1 == b.1) = false ) l) {h'}: + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false )) {h'}: Const.get (m.insertMany l).1 k' h' = v := by simp_to_model using getValue_insertList_of_mem @@ -1034,7 +1034,7 @@ theorem get!_insertMany_list_of_contains_eq_false_const [EquivBEq α] [LawfulHas theorem get!_insertMany_list_of_mem_const [EquivBEq α] [LawfulHashable α] [Inhabited β] (h : m.1.WF) {l : List ((_ : α) × β)} {k k' : α} (k_beq : k == k') {v : β} (mem: ⟨k, v⟩ ∈ l) - (distinct : List.Pairwise (fun a b => (a.1 == b.1) = false ) l) : + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false )) : Const.get! (m.insertMany l).1 k' = v := by simp_to_model using getValue!_insertList_of_mem @@ -1046,7 +1046,7 @@ theorem getD_insertMany_list_of_contains_eq_false_const [EquivBEq α] [LawfulHas theorem getD_insertMany_list_of_mem_const [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l : List ((_ : α) × β)} {k k' : α} (k_beq : k == k') {v fallback: β} (mem: ⟨k, v⟩ ∈ l) - (distinct : List.Pairwise (fun a b => (a.1 == b.1) = false ) l) : + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false )) : Const.getD (m.insertMany l).1 k' fallback = v := by simp_to_model using getValueD_insertList_of_mem @@ -1139,14 +1139,14 @@ theorem getKeyD_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] (h : m. theorem size_insertMany_list [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l : List (α × β)} - {distinct : List.Pairwise (fun a b => (a.1 == b.1) = false) l} : + {distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)} : (∀ (a : α), ¬ (m.contains a ∧ (l.map Prod.fst).contains a)) → (insertMany m l).1.1.size = m.1.size + l.length := by simp_to_model using length_insertListConst theorem isEmpty_insertMany_list_eq_false_of_isEmpty_eq_false [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l : List (α × β)} : - (m.1.isEmpty = false) → (insertMany m l).1.1.isEmpty = false := by + (m.1.isEmpty = false) → (insertMany m l).1.1.isEmpty = false := by simp_to_model using isEmpty_insertListConst_eq_false_of_isEmpty_eq_false theorem isEmpty_insertMany_list [EquivBEq α] [LawfulHashable α] (h : m.1.WF) @@ -1162,7 +1162,7 @@ theorem get?_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable theorem get?_insertMany_list_of_mem_const [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l : List (α × β)} {k k' : α} (k_beq : k == k') {v : β} (mem: ⟨k, v⟩ ∈ l) - (distinct : List.Pairwise (fun a b => (a.1 == b.1) = false ) l) : + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false )) : get? (insertMany m l).1 k' = v := by simp_to_model using getValue?_insertListConst_of_mem @@ -1175,7 +1175,7 @@ theorem get_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable theorem get_insertMany_list_of_mem_const [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l : List (α × β)} {k k' : α} (k_beq : k == k') {v : β} (mem: ⟨k, v⟩ ∈ l) - (distinct : List.Pairwise (fun a b => (a.1 == b.1) = false ) l) {h'}: + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false )) {h'}: get (insertMany m l).1 k' h' = v := by simp_to_model using getValue_insertListConst_of_mem @@ -1187,7 +1187,7 @@ theorem get!_insertMany_list_of_contains_eq_false_const [EquivBEq α] [LawfulHas theorem get!_insertMany_list_of_mem_const [EquivBEq α] [LawfulHashable α] [Inhabited β] (h : m.1.WF) {l : List (α × β)} {k k' : α} (k_beq : k == k') {v : β} (mem: ⟨k, v⟩ ∈ l) - (distinct : List.Pairwise (fun a b => (a.1 == b.1) = false ) l) : + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false )) : get! (insertMany m l).1 k' = v := by simp_to_model using getValue!_insertListConst_of_mem @@ -1199,7 +1199,7 @@ theorem getD_insertMany_list_of_contains_eq_false_const [EquivBEq α] [LawfulHas theorem getD_insertMany_list_of_mem_const [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l : List (α × β)} {k k' : α} (k_beq : k == k') {v fallback: β} (mem: ⟨k, v⟩ ∈ l) - (distinct : List.Pairwise (fun a b => (a.1 == b.1) = false ) l) : + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false )) : getD (insertMany m l).1 k' fallback = v := by simp_to_model using getValueD_insertListConst_of_mem @@ -1261,7 +1261,7 @@ theorem getKey_insertManyIfNewUnit_list_of_mem_of_contains_eq_false [EquivBEq α {k k' : α} (k_beq : k == k') (distinct : l.Pairwise (fun a b => (a == b) = false)) (mem : k ∈ l) {h'} : - m.contains k = false → + m.contains k = false → getKey (insertManyIfNewUnit m l).1 k' h' = k := by simp_to_model using getKey_insertListIfNewUnit_of_mem_of_contains_eq_false @@ -1314,7 +1314,7 @@ theorem getKeyD_insertManyIfNewUnit_list_of_mem_of_contains [EquivBEq α] [Lawfu theorem size_insertManyIfNewUnit_list [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l : List α} - (distinct : List.Pairwise (fun a b => (a == b) = false) l): + (distinct : l.Pairwise (fun a b => (a == b) = false)): (∀ (a : α), ¬ (m.contains a ∧ l.contains a)) → (insertManyIfNewUnit m l).1.1.size = m.1.size + l.length := by simp_to_model using length_insertListIfNewUnit From 3266d6c5a14c3498341f9e761902c448dfbb1f58 Mon Sep 17 00:00:00 2001 From: monsterkrampe Date: Tue, 3 Dec 2024 16:50:23 +0100 Subject: [PATCH 37/83] Fix two more small spacing issues --- src/Std/Data/DHashMap/Internal/RawLemmas.lean | 2 +- src/Std/Data/DHashMap/Internal/WF.lean | 1 + 2 files changed, 2 insertions(+), 1 deletion(-) diff --git a/src/Std/Data/DHashMap/Internal/RawLemmas.lean b/src/Std/Data/DHashMap/Internal/RawLemmas.lean index 474cf4c8aa96..91adba6ee39f 100644 --- a/src/Std/Data/DHashMap/Internal/RawLemmas.lean +++ b/src/Std/Data/DHashMap/Internal/RawLemmas.lean @@ -1247,7 +1247,7 @@ theorem getKey?_insertManyIfNewUnit_list_of_mem_of_contains [EquivBEq α] [Lawfu (h : m.1.WF) {l : List α} {k : α} (distinct : l.Pairwise (fun a b => (a == b) = false)) (h': l.contains k) : - m.contains k → getKey? (insertManyIfNewUnit m l).1 k = getKey? m k := by + m.contains k → getKey? (insertManyIfNewUnit m l).1 k = getKey? m k := by simp_to_model using getKey?_insertListIfNewUnit_of_mem_of_contains theorem getKey_insertManyIfNewUnit_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] diff --git a/src/Std/Data/DHashMap/Internal/WF.lean b/src/Std/Data/DHashMap/Internal/WF.lean index 95bde8df4309..abb0c2d1924d 100644 --- a/src/Std/Data/DHashMap/Internal/WF.lean +++ b/src/Std/Data/DHashMap/Internal/WF.lean @@ -742,6 +742,7 @@ theorem WF.out [BEq α] [Hashable α] [i₁ : EquivBEq α] [i₂ : LawfulHashabl end Raw namespace Raw₀ + /-! # `insertMany` -/ theorem wfImp_insertMany [BEq α] [Hashable α] [EquivBEq α] [LawfulHashable α] {ρ : Type w} From 422ed96c77dddfa23c997fd5769a39ab013e6912 Mon Sep 17 00:00:00 2001 From: monsterkrampe Date: Wed, 4 Dec 2024 16:42:09 +0100 Subject: [PATCH 38/83] Unify getValue? proofs with getValueCast? --- .../DHashMap/Internal/List/Associative.lean | 213 +++++++++--------- src/Std/Data/DHashMap/Internal/Model.lean | 2 +- 2 files changed, 112 insertions(+), 103 deletions(-) diff --git a/src/Std/Data/DHashMap/Internal/List/Associative.lean b/src/Std/Data/DHashMap/Internal/List/Associative.lean index 9cf1beab8b9f..8289a222d421 100644 --- a/src/Std/Data/DHashMap/Internal/List/Associative.lean +++ b/src/Std/Data/DHashMap/Internal/List/Associative.lean @@ -2329,60 +2329,76 @@ theorem isEmpty_insertList [BEq α] simp section -variable {β: Type v} +variable {β : Type v} theorem getValue?_insertList_of_contains_eq_false [BEq α] [PartialEquivBEq α] {l toInsert : List ((_ : α) × β)} {k : α} (h : (toInsert.map Sigma.fst).contains k = false) : getValue? k (insertList l toInsert) = getValue? k l := by + rw [← containsKey_eq_contains_map_fst] at h induction toInsert generalizing l with | nil => simp [insertList] | cons hd tl ih => + rw [containsKey_cons, Bool.or_eq_false_iff] at h simp only [insertList] - rw [ih] - · apply getValue?_insertEntry_of_false - simp only [List.map_cons, List.contains_cons, Bool.or_eq_false_iff] at h - apply BEq.symm_false - apply (And.left h) - · simp only [List.map_cons, List.contains_cons, Bool.or_eq_false_iff] at h - apply And.right h - -theorem getValue?_insertList_of_mem [BEq α] [PartialEquivBEq α] + rw [ih (And.right h)] + rw [getValue?_insertEntry] + simp [And.left h] + +theorem getValue?_insertList_of_mem [BEq α] [EquivBEq α] {l toInsert : List ((_ : α) × β )} {k k' : α} (k_beq : k == k') {v : β} (distinct : toInsert.Pairwise (fun a b => (a.1 == b.1) = false)) + (distinct2 : DistinctKeys l) (mem : ⟨k, v⟩ ∈ toInsert) : getValue? k' (insertList l toInsert) = v := by + rw [← DistinctKeys.def] at distinct induction toInsert generalizing l with | nil => simp at mem | cons hd tl ih => - simp only [insertList] - simp only [List.mem_cons] at mem + rw [getValue?_of_perm distinct2.insertList (insertList_cons_perm distinct2 distinct)] + rw [getValue?_insertEntry] + rw [List.mem_cons] at mem cases mem with - | inl mem => - rw [getValue?_insertList_of_contains_eq_false] - simp only [← mem] - · rw [getValue?_insertEntry_of_beq] - exact k_beq - · simp only [pairwise_cons] at distinct - false_or_by_contra - rename_i h - simp only [Bool.not_eq_false, List.contains_iff_exists_mem_beq, List.mem_map] at h - rcases h with ⟨a', h, k_a'⟩ - rcases h with ⟨a, a_mem, a_fst⟩ - rcases distinct with ⟨distinct, _⟩ - specialize distinct a a_mem - rw [← mem, a_fst] at distinct - simp only at distinct - have k_a': k == a' := by - apply PartialEquivBEq.trans k_beq k_a' - rw [k_a'] at distinct - simp at distinct + | inl mem => rw [← mem]; simp [k_beq] | inr mem => - apply ih - · simp only [pairwise_cons] at distinct - apply And.right distinct - · exact mem + rw [if_neg, ih distinct.tail] + . exact distinct2 + . exact mem + . have := distinct.containsKey_eq_false + rw [containsKey_eq_false_iff] at this + rw [Bool.not_eq_true] + apply BEq.neq_of_neq_of_beq + . apply this; exact mem + . simp [k_beq] + +theorem getValue_insertList_of_contains_eq_false [BEq α] [PartialEquivBEq α] + {l toInsert : List ((_ : α) × β)} {k : α} + (h' : (toInsert.map Sigma.fst).contains k = false) + {h} : + getValue k (insertList l toInsert) h = + getValue k l (containsKey_of_containsKey_insertList h h') := by + rw [← Option.some_inj] + rw [← getValue?_eq_some_getValue] + rw [← getValue?_eq_some_getValue] + rw [getValue?_insertList_of_contains_eq_false] + exact h' + +theorem getValue_insertList_of_mem [BEq α] [EquivBEq α] + {l toInsert : List ((_ : α) × β)} + {k k' : α} (k_beq : k == k') {v : β} + (distinct : toInsert.Pairwise (fun a b => (a.1 == b.1) = false)) + (distinct2 : DistinctKeys l) + (mem : ⟨k, v⟩ ∈ toInsert) + {h} : + getValue k' (insertList l toInsert) h = v := by + rw [← Option.some_inj] + rw [← getValue?_eq_some_getValue] + rw [getValue?_insertList_of_mem] + . exact k_beq + . exact distinct + . exact distinct2 + . exact mem theorem getValue!_insertList_of_contains_eq_false [BEq α] [PartialEquivBEq α] [Inhabited β] {l toInsert : List ((_ : α) × β)} {k : α} @@ -2392,15 +2408,19 @@ theorem getValue!_insertList_of_contains_eq_false [BEq α] [PartialEquivBEq α] rw [getValue?_insertList_of_contains_eq_false] exact h -theorem getValue!_insertList_of_mem [BEq α] [PartialEquivBEq α] [Inhabited β] +theorem getValue!_insertList_of_mem [BEq α] [EquivBEq α] [Inhabited β] {l toInsert : List ((_ : α) × β)} {k k' : α} {v: β} (k_beq : k == k') - (distinct : toInsert.Pairwise (fun a b => (a.1 == b.1) = false)) (mem : ⟨k, v⟩ ∈ toInsert): + (distinct : toInsert.Pairwise (fun a b => (a.1 == b.1) = false)) + (distinct2 : DistinctKeys l) + (mem : ⟨k, v⟩ ∈ toInsert): getValue! k' (insertList l toInsert) = v := by simp only [getValue!_eq_getValue?] - rw [getValue?_insertList_of_mem (mem:= mem)] - · simp + rw [getValue?_insertList_of_mem] + · rw [Option.get!_some] · exact k_beq · exact distinct + . exact distinct2 + . exact mem theorem getValueD_insertList_of_contains_eq_false [BEq α] [PartialEquivBEq α] {l toInsert : List ((_ : α) × β)} {k : α} {fallback : β} @@ -2410,43 +2430,22 @@ theorem getValueD_insertList_of_contains_eq_false [BEq α] [PartialEquivBEq α] rw [getValue?_insertList_of_contains_eq_false] exact h -theorem getValueD_insertList_of_mem [BEq α] [PartialEquivBEq α] +theorem getValueD_insertList_of_mem [BEq α] [EquivBEq α] {l toInsert : List ((_ : α) × β)} {k k' : α} {v fallback: β} (k_beq : k == k') - (distinct : toInsert.Pairwise (fun a b => (a.1 == b.1) = false)) (mem : ⟨k, v⟩ ∈ toInsert): + (distinct : toInsert.Pairwise (fun a b => (a.1 == b.1) = false)) + (distinct2 : DistinctKeys l) + (mem : ⟨k, v⟩ ∈ toInsert): getValueD k' (insertList l toInsert) fallback= v := by simp only [getValueD_eq_getValue?] - rw [getValue?_insertList_of_mem (mem:= mem)] - · simp + rw [getValue?_insertList_of_mem] + · rw [Option.getD_some] · exact k_beq · exact distinct - -theorem getValue_insertList_of_contains_eq_false [BEq α] [PartialEquivBEq α] - {l toInsert : List ((_ : α) × β)} {k : α} - (h' : (toInsert.map Sigma.fst).contains k = false) - {h} : - getValue k (insertList l toInsert) h = - getValue k l (containsKey_of_containsKey_insertList h h') := by - rw [← Option.some_inj] - rw [← getValue?_eq_some_getValue] - rw [← getValue?_eq_some_getValue] - rw [getValue?_insertList_of_contains_eq_false] - exact h' - -theorem getValue_insertList_of_mem [BEq α] [PartialEquivBEq α] - {l toInsert : List ((_ : α) × β)} - {k k' : α} (k_beq : k == k') {v : β} - (distinct : toInsert.Pairwise (fun a b => (a.1 == b.1) = false)) - (mem : ⟨k, v⟩ ∈ toInsert) - {h} : - getValue k' (insertList l toInsert) h = v := by - rw [← Option.some_inj] - rw [← getValue?_eq_some_getValue] - rw [getValue?_insertList_of_mem (mem:=mem)] - . exact k_beq - . exact distinct + . exact distinct2 + . exact mem /-- Internal implementation detail of the hash map -/ -def insertListConst [BEq α] (l: List ((_ : α) × β))(toInsert: List (α × β)) : List ((_ : α) × β) := +def insertListConst [BEq α] (l: List ((_ : α) × β)) (toInsert: List (α × β)) : List ((_ : α) × β) := match toInsert with | .nil => l | .cons hd tl => insertListConst (insertEntry hd.1 hd.2 l) tl @@ -2674,19 +2673,48 @@ theorem getValue?_insertListConst_of_contains_eq_false [BEq α] [PartialEquivBEq rw [list_conversion_insertListConst] exact h -theorem getValue?_insertListConst_of_mem [BEq α] [PartialEquivBEq α] +theorem getValue?_insertListConst_of_mem [BEq α] [EquivBEq α] {l : List ((_ : α) × β)} {toInsert : List (α × β)} {k k' : α} (k_beq : k == k') {v : β} (distinct : toInsert.Pairwise (fun a b => (a.1 == b.1) = false)) + (distinct2 : DistinctKeys l) (mem : ⟨k, v⟩ ∈ toInsert) : getValue? k' (insertListConst l toInsert) = v := by rw [insertListConst_eq_insertList] apply getValue?_insertList_of_mem (k_beq:=k_beq) · apply pairwise_list_conversion_insertListConst distinct - · simp only [List.mem_map, Prod.exists] + · exact distinct2 + . simp only [List.mem_map, Prod.exists] exists k exists v +theorem getValue_insertListConst_of_contains_eq_false [BEq α] [PartialEquivBEq α] + {l : List ((_ : α) × β)} {toInsert : List (α × β)} {k : α} + {h : (toInsert.map Prod.fst).contains k = false} + {h'} : + getValue k (insertListConst l toInsert) h' = + getValue k l (containsKey_of_containsKey_insertListConst h' h) := by + rw [← Option.some_inj] + rw [← getValue?_eq_some_getValue] + rw [← getValue?_eq_some_getValue] + rw [getValue?_insertListConst_of_contains_eq_false] + exact h + +theorem getValue_insertListConst_of_mem [BEq α] [EquivBEq α] + {l : List ((_ : α) × β)} {toInsert : List (α × β)} + {k k' : α} (k_beq : k == k') {v : β} + (distinct : toInsert.Pairwise (fun a b => (a.1 == b.1) = false)) + (distinct2 : DistinctKeys l) + (mem : ⟨k, v⟩ ∈ toInsert) + {h} : + getValue k' (insertListConst l toInsert) h = v := by + rw [← Option.some_inj] + rw [← getValue?_eq_some_getValue] + rw [getValue?_insertListConst_of_mem (mem:=mem)] + . exact k_beq + . exact distinct + . exact distinct2 + theorem getValue!_insertListConst_of_contains_eq_false [BEq α] [PartialEquivBEq α] [Inhabited β] {l : List ((_ : α) × β)} {toInsert : List (α × β)} {k : α} (h : (toInsert.map Prod.fst).contains k = false) : @@ -2696,12 +2724,15 @@ theorem getValue!_insertListConst_of_contains_eq_false [BEq α] [PartialEquivBEq rw [list_conversion_insertListConst] exact h -theorem getValue!_insertListConst_of_mem [BEq α] [PartialEquivBEq α] [Inhabited β] +theorem getValue!_insertListConst_of_mem [BEq α] [EquivBEq α] [Inhabited β] {l : List ((_ : α) × β)} {toInsert : List (α × β)} {k k' : α} {v: β} (k_beq : k == k') - (distinct : toInsert.Pairwise (fun a b => (a.1 == b.1) = false)) (mem : ⟨k, v⟩ ∈ toInsert): + (distinct : toInsert.Pairwise (fun a b => (a.1 == b.1) = false)) + (distinct2 : DistinctKeys l) + (mem : ⟨k, v⟩ ∈ toInsert): getValue! k' (insertListConst l toInsert) = v := by rw [insertListConst_eq_insertList, getValue!_insertList_of_mem (k_beq:=k_beq)] · apply pairwise_list_conversion_insertListConst distinct + . exact distinct2 · simp only [List.mem_map, Prod.exists] exists k exists v @@ -2714,41 +2745,19 @@ theorem getValueD_insertListConst_of_contains_eq_false [BEq α] [PartialEquivBEq rw [list_conversion_insertListConst] exact not_mem -theorem getValueD_insertListConst_of_mem [BEq α] [PartialEquivBEq α] +theorem getValueD_insertListConst_of_mem [BEq α] [EquivBEq α] {l : List ((_ : α) × β)} {toInsert : List (α × β)} {k k' : α} {v fallback: β} (k_beq : k == k') - (distinct : toInsert.Pairwise (fun a b => (a.1 == b.1) = false)) (mem : ⟨k, v⟩ ∈ toInsert): + (distinct : toInsert.Pairwise (fun a b => (a.1 == b.1) = false)) + (distinct2 : DistinctKeys l) + (mem : ⟨k, v⟩ ∈ toInsert): getValueD k' (insertListConst l toInsert) fallback= v := by rw [insertListConst_eq_insertList, getValueD_insertList_of_mem (k_beq:= k_beq)] · apply pairwise_list_conversion_insertListConst distinct + . exact distinct2 · simp exists k exists v -theorem getValue_insertListConst_of_contains_eq_false [BEq α] [PartialEquivBEq α] - {l : List ((_ : α) × β)} {toInsert : List (α × β)} {k : α} - {h : (toInsert.map Prod.fst).contains k = false} - {h'} : - getValue k (insertListConst l toInsert) h' = - getValue k l (containsKey_of_containsKey_insertListConst h' h) := by - rw [← Option.some_inj] - rw [← getValue?_eq_some_getValue] - rw [← getValue?_eq_some_getValue] - rw [getValue?_insertListConst_of_contains_eq_false] - exact h - -theorem getValue_insertListConst_of_mem [BEq α] [PartialEquivBEq α] - {l : List ((_ : α) × β)} {toInsert : List (α × β)} - {k k' : α} (k_beq : k == k') {v : β} - (distinct : toInsert.Pairwise (fun a b => (a.1 == b.1) = false)) - (mem : ⟨k, v⟩ ∈ toInsert) - {h} : - getValue k' (insertListConst l toInsert) h = v := by - rw [← Option.some_inj] - rw [← getValue?_eq_some_getValue] - rw [getValue?_insertListConst_of_mem (mem:=mem)] - . exact k_beq - . exact distinct - /-- Internal implementation detail of the hash map -/ def insertListIfNewUnit [BEq α] (l: List ((_ : α) × Unit)) (toInsert: List α) : List ((_ : α) × Unit) := @@ -3130,12 +3139,12 @@ theorem getValue?_insertListIfNewUnit [BEq α] [PartialEquivBEq α] theorem getValue_insertListIfNewUnit [BEq α] [PartialEquivBEq α] {l : List ((_ : α) × Unit)} {toInsert : List α} {k : α} {h}: - getValue k (insertListIfNewUnit l toInsert) h = () := by + getValue k (insertListIfNewUnit l toInsert) h = () := by rw [getValue_list_unit] theorem getValue!_insertListIfNewUnit [BEq α] [PartialEquivBEq α] {l : List ((_ : α) × Unit)} {toInsert : List α} {k : α} : - getValue! k (insertListIfNewUnit l toInsert) = () := by + getValue! k (insertListIfNewUnit l toInsert) = () := by rw [getValue!_list_unit] theorem getValueD_insertListIfNewUnit [BEq α] [PartialEquivBEq α] diff --git a/src/Std/Data/DHashMap/Internal/Model.lean b/src/Std/Data/DHashMap/Internal/Model.lean index 713e39efc098..30db8950f7f6 100644 --- a/src/Std/Data/DHashMap/Internal/Model.lean +++ b/src/Std/Data/DHashMap/Internal/Model.lean @@ -323,7 +323,7 @@ def filterₘ (m : Raw₀ α β) (f : (a : α) → β a → Bool) : Raw₀ α β ⟨withComputedSize (updateAllBuckets m.1.buckets fun l => l.filter f), by simpa using m.2⟩ /-- Internal implementation detail of the hash map -/ -def insertListₘ [BEq α] [Hashable α](m : Raw₀ α β) (l : List ((a : α) × β a)) : Raw₀ α β := +def insertListₘ [BEq α] [Hashable α] (m : Raw₀ α β) (l : List ((a : α) × β a)) : Raw₀ α β := match l with | .nil => m | .cons hd tl => insertListₘ (m.insert hd.1 hd.2) tl From 5ce3a536a7e66dde2b807954cf05f23e2a820c42 Mon Sep 17 00:00:00 2001 From: monsterkrampe Date: Wed, 4 Dec 2024 17:59:06 +0100 Subject: [PATCH 39/83] Simplify insertListConst --- .../DHashMap/Internal/List/Associative.lean | 209 +++++++----------- src/Std/Data/DHashMap/Internal/WF.lean | 7 +- 2 files changed, 88 insertions(+), 128 deletions(-) diff --git a/src/Std/Data/DHashMap/Internal/List/Associative.lean b/src/Std/Data/DHashMap/Internal/List/Associative.lean index 8289a222d421..bbc0c82752d3 100644 --- a/src/Std/Data/DHashMap/Internal/List/Associative.lean +++ b/src/Std/Data/DHashMap/Internal/List/Associative.lean @@ -2444,91 +2444,46 @@ theorem getValueD_insertList_of_mem [BEq α] [EquivBEq α] . exact distinct2 . exact mem -/-- Internal implementation detail of the hash map -/ -def insertListConst [BEq α] (l: List ((_ : α) × β)) (toInsert: List (α × β)) : List ((_ : α) × β) := - match toInsert with - | .nil => l - | .cons hd tl => insertListConst (insertEntry hd.1 hd.2 l) tl - -theorem insertListConst_eq_insertList [BEq α] {l : List ((_ : α) × β)} {toInsert : List (α × β)}: - insertListConst l toInsert = insertList l (List.map (fun a => ⟨a.1, a.2⟩) toInsert) := by - induction toInsert generalizing l with - | nil => simp [insertList, insertListConst] - | cons hd tl ih => - simp [insertList, insertListConst] - apply ih -theorem DistinctKeys.insertListConst [BEq α] [PartialEquivBEq α] {l₁ : List ((_ : α) × β)} - {l₂ : List (α × β)} (h : DistinctKeys l₁) : - DistinctKeys (insertListConst l₁ l₂) := by - rw [insertListConst_eq_insertList] - apply DistinctKeys.insertList h - -theorem insertListConst_perm_of_perm_first [BEq α] [EquivBEq α] {l1 l2: List ((_ : α) × β)} - {toInsert : List (α × β)} (h : Perm l1 l2) (distinct : DistinctKeys l1) : - Perm (insertListConst l1 toInsert) (insertListConst l2 toInsert) := by - simp only [insertListConst_eq_insertList] - apply insertList_perm_of_perm_first h distinct - -theorem list_conversion_insertListConst {toInsert : List (α × β)}: - (List.map Sigma.fst (List.map (fun a => (⟨a.fst, a.snd⟩ : (_ : α ) × β )) toInsert)) = - (List.map Prod.fst toInsert) := by - apply List.ext_get - · simp - · intro n h1 h2 - simp +-- FIXME: This already exists in mathlib... +def Prod.toSigma (p : α × β) : ((_ : α) × β) := ⟨p.fst, p.snd⟩ -theorem list_conversion_insertListConst_containsKey_true_iff [BEq α] [PartialEquivBEq α] - {toInsert : List (α × β)} {a : α}: - containsKey a (List.map (fun a => (⟨a.fst, a.snd⟩ : ((_ : α ) × β) )) toInsert) = true ↔ - (toInsert.map Prod.fst).contains a = true := by - rw [List.contains_iff_exists_mem_beq] - rw [containsKey_iff_exists] - simp only [keys_def, list_conversion_insertListConst] - -theorem pairwise_list_conversion_insertListConst [BEq α] {toInsert : List (α × β)} - (distinct : toInsert.Pairwise (fun a b => (a.1 == b.1) = false) ): - List.Pairwise (fun a b => (a.1 == b.1) = false) - (List.map (fun a => (⟨a.1, a.2⟩ : ((_ : α ) × β)) ) toInsert) := by - induction toInsert with - | nil => simp - | cons hd tl ih => - simp only [List.map_cons, pairwise_cons] - simp only [List.map_cons, pairwise_cons] at distinct - rcases distinct with ⟨left,right⟩ - constructor - · intro a' a_mem - simp only [List.mem_map, Prod.exists] at a_mem - rcases a_mem with ⟨a,b,ab_tl, ab_a⟩ - rw [← ab_a] - simp only - specialize left (a,b) ab_tl - apply left - · apply ih right +def List.toSigma (l : List (α × β)) : List ((_ : α) × β) := l.map Prod.toSigma + +@[simp] +theorem sigma_fst_after_prod_toSigma_eq_prod_fst : + Sigma.fst ∘ Prod.toSigma = @Prod.fst α β := by + apply funext + simp [Prod.toSigma] + +-- TODO: maybe insertList_prod would be a better name? +/-- Internal implementation detail of the hash map -/ +def insertListConst [BEq α] (l: List ((_ : α) × β)) (toInsert: List (α × β)) : List ((_ : α) × β) := + insertList l toInsert.toSigma theorem containsKey_insertListConst [BEq α] [PartialEquivBEq α] {l : List ((_ : α) × β)} {toInsert : List (α × β)} {k : α} : containsKey k (insertListConst l toInsert) = (containsKey k l || (toInsert.map Prod.fst).contains k) := by - rw [insertListConst_eq_insertList, containsKey_insertList,list_conversion_insertListConst] + unfold insertListConst + rw [containsKey_insertList] + simp [List.toSigma] theorem containsKey_of_containsKey_insertListConst [BEq α] [PartialEquivBEq α] {l : List ((_ : α) × β)} {toInsert : List (α × β)} {k : α} (h₁ : containsKey k (insertListConst l toInsert)) (h₂ : (toInsert.map Prod.fst).contains k = false) : containsKey k l := by - rw [insertListConst_eq_insertList] at h₁ + unfold insertListConst at h₁ apply containsKey_of_containsKey_insertList h₁ - rw [list_conversion_insertListConst] - exact h₂ + simpa [List.toSigma] theorem getKey?_insertListConst_of_contains_eq_false [BEq α] [EquivBEq α] {l : List ((_ : α) × β)} {toInsert : List (α × β)} {k : α} (h : (toInsert.map Prod.fst).contains k = false) : getKey? k (insertListConst l toInsert) = getKey? k l := by - rw [insertListConst_eq_insertList] + unfold insertListConst apply getKey?_insertList_of_contains_eq_false - rw [list_conversion_insertListConst] - exact h + simpa [List.toSigma] theorem getKey?_insertListConst_of_mem [BEq α] [EquivBEq α] {l : List ((_ : α) × β)} {toInsert : List (α × β)} @@ -2537,13 +2492,13 @@ theorem getKey?_insertListConst_of_mem [BEq α] [EquivBEq α] (distinct2 : DistinctKeys l) (mem : k ∈ (toInsert.map Prod.fst)) : getKey? k' (insertListConst l toInsert) = some k := by - rw [insertListConst_eq_insertList] + unfold insertListConst apply getKey?_insertList_of_mem · exact k_beq - · apply pairwise_list_conversion_insertListConst distinct + · simpa [List.toSigma, List.pairwise_map, Prod.toSigma] · exact distinct2 - · rw [list_conversion_insertListConst] - assumption + · simp only [List.toSigma, List.map_map, sigma_fst_after_prod_toSigma_eq_prod_fst] + exact mem theorem getKey_insertListConst_of_contains_eq_false [BEq α] [EquivBEq α] {l : List ((_ : α) × β)} {toInsert : List (α × β)} {k : α} @@ -2577,10 +2532,10 @@ theorem getKey!_insertListConst_of_contains_eq_false [BEq α] [EquivBEq α] [Inh {l : List ((_ : α) × β)} {toInsert : List (α × β)} {k : α} (h : (toInsert.map Prod.fst).contains k = false) : getKey! k (insertListConst l toInsert) = getKey! k l := by - rw [insertListConst_eq_insertList] - apply getKey!_insertList_of_contains_eq_false - rw [list_conversion_insertListConst] - exact h + rw [getKey!_eq_getKey?] + rw [getKey?_insertListConst_of_contains_eq_false] + . rw [getKey!_eq_getKey?] + . exact h theorem getKey!_insertListConst_of_mem [BEq α] [EquivBEq α] [Inhabited α] {l : List ((_ : α) × β)} {toInsert : List (α × β)} @@ -2589,23 +2544,22 @@ theorem getKey!_insertListConst_of_mem [BEq α] [EquivBEq α] [Inhabited α] (distinct2 : DistinctKeys l) (mem : k ∈ (toInsert.map Prod.fst)) : getKey! k' (insertListConst l toInsert) = k := by - rw [insertListConst_eq_insertList] - apply getKey!_insertList_of_mem - · exact k_beq - · apply pairwise_list_conversion_insertListConst - exact distinct - · exact distinct2 - · rw [list_conversion_insertListConst] - exact mem + rw [getKey!_eq_getKey?] + rw [getKey?_insertListConst_of_mem] + . rw [Option.get!_some] + . exact k_beq + . exact distinct + . exact distinct2 + . exact mem theorem getKeyD_insertListConst_of_contains_eq_false [BEq α] [EquivBEq α] {l : List ((_ : α) × β)} {toInsert : List (α × β)} {k fallback : α} (h : (toInsert.map Prod.fst).contains k = false) : getKeyD k (insertListConst l toInsert) fallback = getKeyD k l fallback := by - rw [insertListConst_eq_insertList] - apply getKeyD_insertList_of_contains_eq_false - rw [list_conversion_insertListConst] - exact h + rw [getKeyD_eq_getKey?] + rw [getKey?_insertListConst_of_contains_eq_false] + . rw [getKeyD_eq_getKey?] + . exact h theorem getKeyD_insertListConst_of_mem [BEq α] [EquivBEq α] {l : List ((_ : α) × β)} {toInsert : List (α × β)} @@ -2614,14 +2568,13 @@ theorem getKeyD_insertListConst_of_mem [BEq α] [EquivBEq α] (distinct2 : DistinctKeys l) (mem : k ∈ (toInsert.map Prod.fst)) : getKeyD k' (insertListConst l toInsert) fallback = k := by - rw [insertListConst_eq_insertList] - apply getKeyD_insertList_of_mem - · exact k_beq - · apply pairwise_list_conversion_insertListConst - exact distinct - · exact distinct2 - · rw [list_conversion_insertListConst] - exact mem + rw [getKeyD_eq_getKey?] + rw [getKey?_insertListConst_of_mem] + . rw [Option.getD_some] + . exact k_beq + . exact distinct + . exact distinct2 + . exact mem theorem length_insertListConst [BEq α] [EquivBEq α] {l : List ((_ : α) × β)} {toInsert : List (α × β)} @@ -2629,11 +2582,11 @@ theorem length_insertListConst [BEq α] [EquivBEq α] (distinct_toInsert : toInsert.Pairwise (fun a b => (a.1 == b.1) = false)) (distinct_both : ∀ (a : α), ¬ (containsKey a l ∧ (toInsert.map Prod.fst).contains a)) : (insertListConst l toInsert).length = l.length + toInsert.length := by - rw [insertListConst_eq_insertList] + unfold insertListConst rw [length_insertList] - · simp + · simp [List.toSigma] · exact distinct_l - · apply pairwise_list_conversion_insertListConst distinct_toInsert + · simpa [List.toSigma, List.pairwise_map, Prod.toSigma] · intro a simp only [not_and, Bool.not_eq_true] intro h @@ -2646,32 +2599,38 @@ theorem length_insertListConst [BEq α] [EquivBEq α] rw [List.contains_eq_any_beq] at distinct_both simp only [List.any_map, List.any_eq_false, Function.comp_apply, Bool.not_eq_true, Prod.forall] at distinct_both - apply distinct_both + intro x hx + apply distinct_both x.fst x.snd + simp only [List.toSigma, List.mem_map] at hx + rcases hx with ⟨_, _, eq⟩ + rw [← eq] + simpa [Prod.toSigma] theorem isEmpty_insertListConst_eq_false_of_isEmpty_eq_false [BEq α] {l : List ((_ : α) × β)} {toInsert : List (α × β)} (h : l.isEmpty = false) : (List.insertListConst l toInsert).isEmpty = false := by induction toInsert generalizing l with - | nil => simp [insertListConst, h] + | nil => simp [insertListConst, insertList, List.toSigma, h] | cons hd tl ih => - simp [insertListConst, ih] + unfold insertListConst at ih + unfold List.toSigma at ih + simp [insertListConst, insertList, ih] theorem isEmpty_insertListConst [BEq α] {l : List ((_ : α) × β)} {toInsert : List (α × β)} : (List.insertListConst l toInsert).isEmpty ↔ l.isEmpty ∧ toInsert.isEmpty := by - rw [insertListConst_eq_insertList] + unfold insertListConst rw [isEmpty_insertList] - simp only [List.isEmpty_eq_true, List.map_eq_nil_iff] + simp only [List.isEmpty_eq_true, List.map_eq_nil_iff, List.toSigma] theorem getValue?_insertListConst_of_contains_eq_false [BEq α] [PartialEquivBEq α] {l : List ((_ : α) × β)} {toInsert : List (α × β)} {k : α} (h : (toInsert.map Prod.fst).contains k = false) : getValue? k (insertListConst l toInsert) = getValue? k l := by - rw [insertListConst_eq_insertList] + unfold insertListConst rw [getValue?_insertList_of_contains_eq_false] - rw [list_conversion_insertListConst] - exact h + simpa [List.toSigma] theorem getValue?_insertListConst_of_mem [BEq α] [EquivBEq α] {l : List ((_ : α) × β)} {toInsert : List (α × β)} @@ -2680,13 +2639,12 @@ theorem getValue?_insertListConst_of_mem [BEq α] [EquivBEq α] (distinct2 : DistinctKeys l) (mem : ⟨k, v⟩ ∈ toInsert) : getValue? k' (insertListConst l toInsert) = v := by - rw [insertListConst_eq_insertList] + unfold insertListConst apply getValue?_insertList_of_mem (k_beq:=k_beq) - · apply pairwise_list_conversion_insertListConst distinct + · simpa [List.toSigma, List.pairwise_map, Prod.toSigma] · exact distinct2 - . simp only [List.mem_map, Prod.exists] - exists k - exists v + . simp only [List.toSigma, List.mem_map] + exists (k, v) theorem getValue_insertListConst_of_contains_eq_false [BEq α] [PartialEquivBEq α] {l : List ((_ : α) × β)} {toInsert : List (α × β)} {k : α} @@ -2719,9 +2677,8 @@ theorem getValue!_insertListConst_of_contains_eq_false [BEq α] [PartialEquivBEq {l : List ((_ : α) × β)} {toInsert : List (α × β)} {k : α} (h : (toInsert.map Prod.fst).contains k = false) : getValue! k (insertListConst l toInsert) = getValue! k l := by - rw [insertListConst_eq_insertList] - rw [getValue!_insertList_of_contains_eq_false] - rw [list_conversion_insertListConst] + simp only [getValue!_eq_getValue?] + rw [getValue?_insertListConst_of_contains_eq_false] exact h theorem getValue!_insertListConst_of_mem [BEq α] [EquivBEq α] [Inhabited β] @@ -2730,19 +2687,20 @@ theorem getValue!_insertListConst_of_mem [BEq α] [EquivBEq α] [Inhabited β] (distinct2 : DistinctKeys l) (mem : ⟨k, v⟩ ∈ toInsert): getValue! k' (insertListConst l toInsert) = v := by - rw [insertListConst_eq_insertList, getValue!_insertList_of_mem (k_beq:=k_beq)] - · apply pairwise_list_conversion_insertListConst distinct + simp only [getValue!_eq_getValue?] + rw [getValue?_insertListConst_of_mem] + · rw [Option.get!_some] + · exact k_beq + · exact distinct . exact distinct2 - · simp only [List.mem_map, Prod.exists] - exists k - exists v + . exact mem theorem getValueD_insertListConst_of_contains_eq_false [BEq α] [PartialEquivBEq α] {l : List ((_ : α) × β)} {toInsert : List (α × β)} {k : α} {fallback : β} (not_mem : (toInsert.map Prod.fst).contains k = false) : getValueD k (insertListConst l toInsert) fallback = getValueD k l fallback := by - rw [insertListConst_eq_insertList, getValueD_insertList_of_contains_eq_false] - rw [list_conversion_insertListConst] + simp only [getValueD_eq_getValue?] + rw [getValue?_insertListConst_of_contains_eq_false] exact not_mem theorem getValueD_insertListConst_of_mem [BEq α] [EquivBEq α] @@ -2751,12 +2709,13 @@ theorem getValueD_insertListConst_of_mem [BEq α] [EquivBEq α] (distinct2 : DistinctKeys l) (mem : ⟨k, v⟩ ∈ toInsert): getValueD k' (insertListConst l toInsert) fallback= v := by - rw [insertListConst_eq_insertList, getValueD_insertList_of_mem (k_beq:= k_beq)] - · apply pairwise_list_conversion_insertListConst distinct + simp only [getValueD_eq_getValue?] + rw [getValue?_insertListConst_of_mem] + · rw [Option.getD_some] + · exact k_beq + · exact distinct . exact distinct2 - · simp - exists k - exists v + . exact mem /-- Internal implementation detail of the hash map -/ def insertListIfNewUnit [BEq α] (l: List ((_ : α) × Unit)) (toInsert: List α) : diff --git a/src/Std/Data/DHashMap/Internal/WF.lean b/src/Std/Data/DHashMap/Internal/WF.lean index abb0c2d1924d..9c9a2520dca0 100644 --- a/src/Std/Data/DHashMap/Internal/WF.lean +++ b/src/Std/Data/DHashMap/Internal/WF.lean @@ -763,12 +763,13 @@ theorem Const.toListModel_insertListₘ {β : Type v} [BEq α] [Hashable α] [Eq [LawfulHashable α] {m : Raw₀ α (fun _ => β)}{l : List (α × β)} (h : Raw.WFImp m.1): Perm (toListModel (Const.insertListₘ m l).1.buckets) (insertListConst (toListModel m.1.buckets) l) := by induction l generalizing m with - | nil => simp [Const.insertListₘ, insertListConst] + | nil => simp [Const.insertListₘ, insertListConst, insertList, List.toSigma] | cons hd tl ih => - simp [Const.insertListₘ, insertListConst] + simp only [Const.insertListₘ, insertListConst] apply Perm.trans apply ih (wfImp_insert h) - apply List.insertListConst_perm_of_perm_first + unfold insertListConst + apply List.insertList_perm_of_perm_first apply toListModel_insert h apply (wfImp_insert h).distinct From be3fc596927d45280c88a54b4ef6e48b0ce1375e Mon Sep 17 00:00:00 2001 From: monsterkrampe Date: Thu, 5 Dec 2024 16:05:00 +0100 Subject: [PATCH 40/83] Remove and simplify according to some PR comments --- .../DHashMap/Internal/List/Associative.lean | 206 +++--------------- src/Std/Data/DHashMap/Internal/RawLemmas.lean | 6 +- src/Std/Data/DHashMap/Internal/WF.lean | 2 +- 3 files changed, 39 insertions(+), 175 deletions(-) diff --git a/src/Std/Data/DHashMap/Internal/List/Associative.lean b/src/Std/Data/DHashMap/Internal/List/Associative.lean index bbc0c82752d3..15fab05f9f79 100644 --- a/src/Std/Data/DHashMap/Internal/List/Associative.lean +++ b/src/Std/Data/DHashMap/Internal/List/Associative.lean @@ -38,14 +38,6 @@ theorem assoc_induction {motive : List ((a : α) × β a) → Prop} (nil : motiv | [] => nil | ⟨_, _⟩ :: _ => cons _ _ _ (assoc_induction nil cons _) -theorem keys_def {l: List ((a : α ) × β a)}: - keys l = l.map Sigma.fst := by - induction l with - | nil => simp [keys] - | cons hd tl ih => - simp only [keys, List.map_cons, List.cons.injEq, true_and] - apply ih - /-- Internal implementation detail of the hash map -/ def getEntry? [BEq α] (a : α) : List ((a : α) × β a) → Option ((a : α) × β a) | [] => none @@ -172,35 +164,6 @@ theorem getValueCast?_cons_self [BEq α] [LawfulBEq α] {l : List ((a : α) × getValueCast? k (⟨k, v⟩ :: l) = some v := by rw [getValueCast?_cons_of_true BEq.refl, cast_eq] -theorem getValueCast?_of_mem [BEq α] [LawfulBEq α] - {l : List ((a : α) × β a)} - {k : α} (k' : α ) (v : β k) (k_beq : k == k') (mem : ⟨k, v⟩ ∈ l) - (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) : - getValueCast? k' l = some (cast (by congr; apply LawfulBEq.eq_of_beq k_beq) v) := by - induction l with - | nil => simp at mem - | cons hd tl ih => - simp only [getValueCast?, beq_iff_eq] - simp only [List.mem_cons] at mem - cases mem with - | inl mem => - rw [← mem] - simp [k_beq] - | inr mem => - have hd_k' : (hd.fst == k') = false := by - simp only [beq_eq_false_iff_ne, ne_eq, pairwise_cons] at distinct - simp only [beq_eq_false_iff_ne, ne_eq] - rcases distinct with ⟨distinct, _⟩ - specialize distinct ⟨k, v⟩ mem - simp only at distinct - simp only [beq_iff_eq] at k_beq - rw [k_beq] at distinct - apply distinct - simp only [hd_k', Bool.false_eq_true, ↓reduceDIte] - apply ih - · exact mem - · exact List.Pairwise.of_cons distinct - theorem getValue?_eq_getValueCast? [BEq α] [LawfulBEq α] {β : Type v} {l : List ((_ : α) × β)} {a : α} : getValue? a l = getValueCast? a l := by induction l using assoc_induction <;> simp_all [getValueCast?_cons, getValue?_cons] @@ -1038,13 +1001,6 @@ theorem insertEntry_of_containsKey_eq_false [BEq α] {l : List ((a : α) × β a (h : containsKey k l = false) : insertEntry k v l = ⟨k, v⟩ :: l := by simp [insertEntry, h] -theorem perm_insertEntry_of_containsKey_eq_false [BEq α] - {l : List ((a : α) × β a)} {k : α} {v : β k} - (h : containsKey k l = false) : Perm (insertEntry k v l) (⟨k, v⟩ :: l) := by - cases l with - | nil => simp - | cons hd tl => simp [insertEntry, h] - theorem DistinctKeys.insertEntry [BEq α] [PartialEquivBEq α] {l : List ((a : α) × β a)} {k : α} {v : β k} (h : DistinctKeys l) : DistinctKeys (insertEntry k v l) := by cases h' : containsKey k l @@ -2215,50 +2171,25 @@ theorem getKeyD_insertList_of_mem [BEq α] [EquivBEq α] . exact distinct2 . exact mem -theorem perm_insertList [BEq α] [ReflBEq α] [PartialEquivBEq α] +theorem perm_insertList [BEq α] [EquivBEq α] {l toInsert : List ((a : α) × β a)} (distinct_l : DistinctKeys l) (distinct_toInsert : toInsert.Pairwise (fun a b => (a.1 == b.1) = false)) - (distinct_both : ∀ (a : α), ¬ (containsKey a l ∧ containsKey a toInsert)) : + (distinct_both : ∀ (a : α), ¬ (containsKey a l ∧ (toInsert.map Sigma.fst).contains a)) : Perm (insertList l toInsert) (l ++ toInsert) := by rw [← DistinctKeys.def] at distinct_toInsert + simp only [not_and, Bool.not_eq_true] at distinct_both induction toInsert generalizing l with | nil => simp only [insertList, List.append_nil, Perm.refl] | cons hd tl ih => - simp only [insertList] - apply Perm.trans - · apply ih - · simp only [insertEntry] - specialize distinct_both hd.1 - simp only [containsKey, BEq.refl, Bool.true_or, and_true, Bool.not_eq_true] at distinct_both - simp only [distinct_both, cond_false, distinctKeys_cons_iff, and_true] - exact distinct_l - · apply DistinctKeys.tail distinct_toInsert - · intro a - specialize distinct_both a - simp only [containsKey, Bool.or_eq_true, not_and, not_or, Bool.not_eq_true] at distinct_both - simp only [containsKey_insertEntry, Bool.or_eq_true, not_and, Bool.not_eq_true] - by_cases hd_a : hd.fst == a - · rw [distinctKeys_cons_iff] at distinct_toInsert - have contains_a_tl : containsKey a tl = false := by - false_or_by_contra - rename_i h - simp only [Bool.not_eq_false] at h - rw [containsKey_congr (PartialEquivBEq.symm hd_a)] at h - simp only [h, Bool.true_eq_false, and_false] at distinct_toInsert - simp [contains_a_tl] - · simp only [hd_a, Bool.false_eq_true, false_or] - simp only [Bool.not_eq_true] at hd_a - simp only [hd_a, true_and] at distinct_both - apply distinct_both - · simp only [insertEntry] - specialize distinct_both hd.1 - simp only [containsKey, BEq.refl, Bool.true_or, and_true, Bool.not_eq_true] at distinct_both - simp only [distinct_both, cond_false, List.cons_append] - have h_hd : hd = ⟨hd.fst, hd.snd⟩ := rfl - rw [← h_hd] - apply Perm.symm - apply List.perm_middle + simp only [List.map_cons, List.contains_cons, Bool.or_eq_false_iff] at distinct_both + refine (insertList_cons_perm distinct_l distinct_toInsert).trans ?_ + rw [insertEntry_of_containsKey_eq_false] + · refine (Perm.cons _ (ih distinct_l (distinct_toInsert).tail ?_)).trans List.perm_middle.symm + exact fun a ha => (distinct_both a ha).2 + · simp only [containsKey_insertList, Bool.or_eq_false_iff, ← containsKey_eq_contains_map_fst] + refine ⟨Bool.not_eq_true _ ▸ fun h => ?_, distinct_toInsert.containsKey_eq_false⟩ + simpa using (distinct_both _ h).1 theorem length_insertList [BEq α] [EquivBEq α] {l toInsert : List ((a : α) × β a)} @@ -2266,46 +2197,7 @@ theorem length_insertList [BEq α] [EquivBEq α] (distinct_toInsert : toInsert.Pairwise (fun a b => (a.1 == b.1) = false)) (distinct_both : ∀ (a : α), ¬ (containsKey a l ∧ (toInsert.map Sigma.fst).contains a)) : (insertList l toInsert).length = l.length + toInsert.length := by - rw [← DistinctKeys.def] at distinct_toInsert - induction toInsert generalizing l with - | nil => simp [insertList] - | cons hd tl ih => - simp only [insertList, List.length_cons] - rw [ih] - · rw [length_insertEntry] - split - · rename_i h - specialize distinct_both hd.1 - simp only [not_and, Bool.not_eq_true] at distinct_both - specialize distinct_both h - simp [containsKey] at distinct_both - · rw [Nat.add_assoc, Nat.add_comm 1 _] - · apply DistinctKeys.insertEntry distinct_l - · apply DistinctKeys.tail distinct_toInsert - · intro a - specialize distinct_both a - simp only [List.map_cons, List.contains_cons, Bool.or_eq_true, not_and, not_or, - Bool.not_eq_true] at distinct_both - simp only [containsKey_insertEntry, Bool.or_eq_true, not_and, Bool.not_eq_true] - intro h - cases h with - | inl h => - rw [DistinctKeys.def] at distinct_toInsert - simp at distinct_toInsert - rcases distinct_toInsert with ⟨left,_⟩ - simp only [List.contains_eq_any_beq, List.any_map, List.any_eq_false, Function.comp_apply, - Bool.not_eq_true] - intro x x_mem - specialize left x x_mem - false_or_by_contra - rename_i h' - simp only [Bool.not_eq_false] at h' - have hd_x: hd.fst == x.fst := BEq.trans h h' - rw [hd_x] at left - simp only [Bool.true_eq_false] at left - | inr h => - apply And.right (distinct_both h) - + simpa using (perm_insertList distinct_l distinct_toInsert distinct_both).length_eq theorem isEmpty_insertList_eq_false_of_isEmpty_eq_false [BEq α] {l toInsert : List ((a : α) × β a)} @@ -2445,21 +2337,20 @@ theorem getValueD_insertList_of_mem [BEq α] [EquivBEq α] . exact mem --- FIXME: This already exists in mathlib... -def Prod.toSigma (p : α × β) : ((_ : α) × β) := ⟨p.fst, p.snd⟩ - -def List.toSigma (l : List (α × β)) : List ((_ : α) × β) := l.map Prod.toSigma +-- TODO: this should be unified with Mathlib... +/-- Internal implementation detail of the hash map -/ +def Prod.toSigma_HashMapInternal (p : α × β) : ((_ : α) × β) := ⟨p.fst, p.snd⟩ @[simp] -theorem sigma_fst_after_prod_toSigma_eq_prod_fst : - Sigma.fst ∘ Prod.toSigma = @Prod.fst α β := by +theorem Prod.fst_comp_toSigma_HashMapInternal : + Sigma.fst ∘ Prod.toSigma_HashMapInternal = @Prod.fst α β := by apply funext - simp [Prod.toSigma] + simp [Prod.toSigma_HashMapInternal] -- TODO: maybe insertList_prod would be a better name? /-- Internal implementation detail of the hash map -/ def insertListConst [BEq α] (l: List ((_ : α) × β)) (toInsert: List (α × β)) : List ((_ : α) × β) := - insertList l toInsert.toSigma + insertList l (toInsert.map Prod.toSigma_HashMapInternal) theorem containsKey_insertListConst [BEq α] [PartialEquivBEq α] {l : List ((_ : α) × β)} {toInsert : List (α × β)} {k : α} : @@ -2467,7 +2358,7 @@ theorem containsKey_insertListConst [BEq α] [PartialEquivBEq α] (containsKey k l || (toInsert.map Prod.fst).contains k) := by unfold insertListConst rw [containsKey_insertList] - simp [List.toSigma] + simp theorem containsKey_of_containsKey_insertListConst [BEq α] [PartialEquivBEq α] {l : List ((_ : α) × β)} {toInsert : List (α × β)} {k : α} @@ -2475,7 +2366,7 @@ theorem containsKey_of_containsKey_insertListConst [BEq α] [PartialEquivBEq α] (h₂ : (toInsert.map Prod.fst).contains k = false) : containsKey k l := by unfold insertListConst at h₁ apply containsKey_of_containsKey_insertList h₁ - simpa [List.toSigma] + simpa theorem getKey?_insertListConst_of_contains_eq_false [BEq α] [EquivBEq α] {l : List ((_ : α) × β)} {toInsert : List (α × β)} {k : α} @@ -2483,7 +2374,7 @@ theorem getKey?_insertListConst_of_contains_eq_false [BEq α] [EquivBEq α] getKey? k (insertListConst l toInsert) = getKey? k l := by unfold insertListConst apply getKey?_insertList_of_contains_eq_false - simpa [List.toSigma] + simpa theorem getKey?_insertListConst_of_mem [BEq α] [EquivBEq α] {l : List ((_ : α) × β)} {toInsert : List (α × β)} @@ -2495,9 +2386,9 @@ theorem getKey?_insertListConst_of_mem [BEq α] [EquivBEq α] unfold insertListConst apply getKey?_insertList_of_mem · exact k_beq - · simpa [List.toSigma, List.pairwise_map, Prod.toSigma] + · simpa [List.pairwise_map] · exact distinct2 - · simp only [List.toSigma, List.map_map, sigma_fst_after_prod_toSigma_eq_prod_fst] + · simp only [List.map_map, Prod.fst_comp_toSigma_HashMapInternal] exact mem theorem getKey_insertListConst_of_contains_eq_false [BEq α] [EquivBEq α] @@ -2584,37 +2475,19 @@ theorem length_insertListConst [BEq α] [EquivBEq α] (insertListConst l toInsert).length = l.length + toInsert.length := by unfold insertListConst rw [length_insertList] - · simp [List.toSigma] + · simp · exact distinct_l - · simpa [List.toSigma, List.pairwise_map, Prod.toSigma] - · intro a - simp only [not_and, Bool.not_eq_true] - intro h - specialize distinct_both a - simp only [not_and, Bool.not_eq_true] at distinct_both - specialize distinct_both h - rw [List.contains_eq_any_beq] - simp only [List.map_map, List.any_map, List.any_eq_false, Function.comp_apply, Bool.not_eq_true, - Prod.forall] - rw [List.contains_eq_any_beq] at distinct_both - simp only [List.any_map, List.any_eq_false, Function.comp_apply, Bool.not_eq_true, - Prod.forall] at distinct_both - intro x hx - apply distinct_both x.fst x.snd - simp only [List.toSigma, List.mem_map] at hx - rcases hx with ⟨_, _, eq⟩ - rw [← eq] - simpa [Prod.toSigma] + · simpa [List.pairwise_map] + · simpa using distinct_both theorem isEmpty_insertListConst_eq_false_of_isEmpty_eq_false [BEq α] {l : List ((_ : α) × β)} {toInsert : List (α × β)} (h : l.isEmpty = false) : (List.insertListConst l toInsert).isEmpty = false := by induction toInsert generalizing l with - | nil => simp [insertListConst, insertList, List.toSigma, h] + | nil => simp [insertListConst, insertList, h] | cons hd tl ih => unfold insertListConst at ih - unfold List.toSigma at ih simp [insertListConst, insertList, ih] theorem isEmpty_insertListConst [BEq α] @@ -2622,7 +2495,7 @@ theorem isEmpty_insertListConst [BEq α] (List.insertListConst l toInsert).isEmpty ↔ l.isEmpty ∧ toInsert.isEmpty := by unfold insertListConst rw [isEmpty_insertList] - simp only [List.isEmpty_eq_true, List.map_eq_nil_iff, List.toSigma] + simp only [List.isEmpty_eq_true, List.map_eq_nil_iff] theorem getValue?_insertListConst_of_contains_eq_false [BEq α] [PartialEquivBEq α] {l : List ((_ : α) × β)} {toInsert : List (α × β)} {k : α} @@ -2630,7 +2503,7 @@ theorem getValue?_insertListConst_of_contains_eq_false [BEq α] [PartialEquivBEq getValue? k (insertListConst l toInsert) = getValue? k l := by unfold insertListConst rw [getValue?_insertList_of_contains_eq_false] - simpa [List.toSigma] + simpa theorem getValue?_insertListConst_of_mem [BEq α] [EquivBEq α] {l : List ((_ : α) × β)} {toInsert : List (α × β)} @@ -2641,9 +2514,9 @@ theorem getValue?_insertListConst_of_mem [BEq α] [EquivBEq α] getValue? k' (insertListConst l toInsert) = v := by unfold insertListConst apply getValue?_insertList_of_mem (k_beq:=k_beq) - · simpa [List.toSigma, List.pairwise_map, Prod.toSigma] + · simpa [List.pairwise_map] · exact distinct2 - . simp only [List.toSigma, List.mem_map] + . simp only [List.mem_map] exists (k, v) theorem getValue_insertListConst_of_contains_eq_false [BEq α] [PartialEquivBEq α] @@ -3073,22 +2946,13 @@ theorem getValue?_list_unit [BEq α] {l : List ((_ : α) × Unit)} {k : α}: · simp [hd_k, ih] theorem getValue_list_unit [BEq α] {l : List ((_ : α) × Unit)} {k : α} {h}: - getValue k l h = () := by - induction l with - | nil => simp - | cons hd tl ih => simp + getValue k l h = () := by simp theorem getValue!_list_unit [BEq α] {l : List ((_ : α) × Unit)} {k : α} : - getValue! k l = () := by - induction l with - | nil => simp - | cons hd tl ih => simp + getValue! k l = () := by simp theorem getValueD_list_unit [BEq α] {l : List ((_ : α) × Unit)} {k : α} {fallback : Unit} : - getValueD k l fallback = () := by - induction l with - | nil => simp - | cons hd tl ih => simp + getValueD k l fallback = () := by simp theorem getValue?_insertListIfNewUnit [BEq α] [PartialEquivBEq α] {l : List ((_ : α) × Unit)} {toInsert : List α} {k : α}: diff --git a/src/Std/Data/DHashMap/Internal/RawLemmas.lean b/src/Std/Data/DHashMap/Internal/RawLemmas.lean index 91adba6ee39f..5fd710f8d0cd 100644 --- a/src/Std/Data/DHashMap/Internal/RawLemmas.lean +++ b/src/Std/Data/DHashMap/Internal/RawLemmas.lean @@ -1046,7 +1046,7 @@ theorem getD_insertMany_list_of_contains_eq_false_const [EquivBEq α] [LawfulHas theorem getD_insertMany_list_of_mem_const [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l : List ((_ : α) × β)} {k k' : α} (k_beq : k == k') {v fallback: β} (mem: ⟨k, v⟩ ∈ l) - (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false )) : + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) : Const.getD (m.insertMany l).1 k' fallback = v := by simp_to_model using getValueD_insertList_of_mem @@ -1088,7 +1088,7 @@ theorem getKey?_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] (h : m. {l : List (α × β)} {k k' : α} (k_beq : k == k') (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) - (mem : k ∈ (l.map Prod.fst)) : + (mem : k ∈ l.map Prod.fst) : (insertMany m l).1.getKey? k' = some k := by simp_to_model using getKey?_insertListConst_of_mem @@ -1229,7 +1229,7 @@ theorem contains_insertManyIfNewUnit_list [EquivBEq α] [LawfulHashable α] (h : theorem contains_of_contains_insertManyIfNewUnit_list [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l : List α} {k : α} (h₂ : l.contains k = false) : - (insertManyIfNewUnit m l).1.contains k → m.contains k := by + (insertManyIfNewUnit m l).1.contains k → m.contains k := by simp_to_model using containsKey_of_containsKey_insertListIfNewUnit theorem getKey?_insertManyIfNewUnit_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] diff --git a/src/Std/Data/DHashMap/Internal/WF.lean b/src/Std/Data/DHashMap/Internal/WF.lean index 9c9a2520dca0..37ceaad6485e 100644 --- a/src/Std/Data/DHashMap/Internal/WF.lean +++ b/src/Std/Data/DHashMap/Internal/WF.lean @@ -763,7 +763,7 @@ theorem Const.toListModel_insertListₘ {β : Type v} [BEq α] [Hashable α] [Eq [LawfulHashable α] {m : Raw₀ α (fun _ => β)}{l : List (α × β)} (h : Raw.WFImp m.1): Perm (toListModel (Const.insertListₘ m l).1.buckets) (insertListConst (toListModel m.1.buckets) l) := by induction l generalizing m with - | nil => simp [Const.insertListₘ, insertListConst, insertList, List.toSigma] + | nil => simp [Const.insertListₘ, insertListConst, insertList] | cons hd tl ih => simp only [Const.insertListₘ, insertListConst] apply Perm.trans From 65728131594853242594b997aa7415ff81c0e654 Mon Sep 17 00:00:00 2001 From: monsterkrampe Date: Thu, 5 Dec 2024 16:11:24 +0100 Subject: [PATCH 41/83] Break some long lines --- src/Std/Data/DHashMap/Internal/List/Associative.lean | 3 ++- src/Std/Data/DHashMap/Internal/WF.lean | 3 ++- 2 files changed, 4 insertions(+), 2 deletions(-) diff --git a/src/Std/Data/DHashMap/Internal/List/Associative.lean b/src/Std/Data/DHashMap/Internal/List/Associative.lean index 15fab05f9f79..1c6d9c43130e 100644 --- a/src/Std/Data/DHashMap/Internal/List/Associative.lean +++ b/src/Std/Data/DHashMap/Internal/List/Associative.lean @@ -2647,7 +2647,8 @@ theorem containsKey_insertListIfNewUnit_iff [BEq α] [PartialEquivBEq α] {l : L simp only [Bool.or_eq_true] rw [or_comm (a:= (hd == k)= true), or_assoc, BEq.comm] -theorem containsKey_of_containsKey_insertListIfNewUnit [BEq α] [PartialEquivBEq α] {l : List ((_ : α) × Unit)} {toInsert : List α} {k : α} +theorem containsKey_of_containsKey_insertListIfNewUnit [BEq α] [PartialEquivBEq α] + {l : List ((_ : α) × Unit)} {toInsert : List α} {k : α} (h₂ : toInsert.contains k = false) : containsKey k (insertListIfNewUnit l toInsert) → containsKey k l := by intro h₁ diff --git a/src/Std/Data/DHashMap/Internal/WF.lean b/src/Std/Data/DHashMap/Internal/WF.lean index 37ceaad6485e..ef6ebed33fb3 100644 --- a/src/Std/Data/DHashMap/Internal/WF.lean +++ b/src/Std/Data/DHashMap/Internal/WF.lean @@ -708,7 +708,8 @@ theorem wfImp_filter [BEq α] [Hashable α] [EquivBEq α] [LawfulHashable α] {m /-! # `insertListₘ` -/ -theorem toListModel_insertListₘ [BEq α] [Hashable α] [EquivBEq α][LawfulHashable α] {m : Raw₀ α β} {l : List ((a : α) × β a)} (h : Raw.WFImp m.1) : +theorem toListModel_insertListₘ [BEq α] [Hashable α] [EquivBEq α][LawfulHashable α] + {m : Raw₀ α β} {l : List ((a : α) × β a)} (h : Raw.WFImp m.1) : Perm (toListModel (insertListₘ m l).1.buckets) (List.insertList (toListModel m.1.buckets) l) := by induction l generalizing m with | nil => From be75a4c72254730ebaa8129648c94490b88e43ab Mon Sep 17 00:00:00 2001 From: monsterkrampe Date: Thu, 5 Dec 2024 16:43:26 +0100 Subject: [PATCH 42/83] Rename theorems to resolve Const namespace issue --- .../DHashMap/Internal/List/Associative.lean | 14 ++-- src/Std/Data/DHashMap/Internal/RawLemmas.lean | 81 ++++++++++--------- 2 files changed, 48 insertions(+), 47 deletions(-) diff --git a/src/Std/Data/DHashMap/Internal/List/Associative.lean b/src/Std/Data/DHashMap/Internal/List/Associative.lean index 1c6d9c43130e..49e329891881 100644 --- a/src/Std/Data/DHashMap/Internal/List/Associative.lean +++ b/src/Std/Data/DHashMap/Internal/List/Associative.lean @@ -2719,7 +2719,7 @@ theorem getKey?_insertListIfNewUnit_of_mem_of_contains_eq_false [BEq α] [EquivB apply (And.left distinct) k mem · exact mem' -theorem getKey?_insertListIfNewUnit_of_mem_of_contains [BEq α] [EquivBEq α] +theorem getKey?_insertListIfNewUnit_of_contains_of_contains [BEq α] [EquivBEq α] {l : List ((_ : α) × Unit)} {toInsert : List α} {k : α} (distinct : toInsert.Pairwise (fun a b => (a == b) = false)) @@ -2787,7 +2787,7 @@ theorem getKey_insertListIfNewUnit_of_mem_of_contains_eq_false [BEq α] [EquivBE . exact mem . exact mem' -theorem getKey_insertListIfNewUnit_of_mem_of_contains [BEq α] [EquivBEq α] +theorem getKey_insertListIfNewUnit_of_contains_of_contains [BEq α] [EquivBEq α] {l : List ((_ : α) × Unit)} {toInsert : List α} {k : α} (distinct : toInsert.Pairwise (fun a b => (a == b) = false)) @@ -2796,7 +2796,7 @@ theorem getKey_insertListIfNewUnit_of_mem_of_contains [BEq α] [EquivBEq α] rw [← Option.some_inj] rw [← getKey?_eq_some_getKey] rw [← getKey?_eq_some_getKey] - rw [getKey?_insertListIfNewUnit_of_mem_of_contains] + rw [getKey?_insertListIfNewUnit_of_contains_of_contains] · exact distinct · exact mem · exact mem' @@ -2824,14 +2824,14 @@ theorem getKey!_insertListIfNewUnit_of_mem_of_contains_eq_false [BEq α] [EquivB . exact mem . exact h' -theorem getKey!_insertListIfNewUnit_of_mem_of_contains [BEq α] [EquivBEq α] [Inhabited α] +theorem getKey!_insertListIfNewUnit_of_contains_of_contains [BEq α] [EquivBEq α] [Inhabited α] {l : List ((_ : α) × Unit)} {toInsert : List α} {k : α} (distinct : toInsert.Pairwise (fun a b => (a == b) = false)) (mem : toInsert.contains k) (h' : containsKey k l = true) : getKey! k (insertListIfNewUnit l toInsert) = getKey! k l := by rw [getKey!_eq_getKey?] - rw [getKey?_insertListIfNewUnit_of_mem_of_contains] + rw [getKey?_insertListIfNewUnit_of_contains_of_contains] . rw [getKey!_eq_getKey?] · exact distinct · exact mem @@ -2860,14 +2860,14 @@ theorem getKeyD_insertListIfNewUnit_of_mem_of_contains_eq_false [BEq α] [EquivB . exact mem . exact h' -theorem getKeyD_insertListIfNewUnit_of_mem_of_contains [BEq α] [EquivBEq α] +theorem getKeyD_insertListIfNewUnit_of_contains_of_contains [BEq α] [EquivBEq α] {l : List ((_ : α) × Unit)} {toInsert : List α} {k fallback : α} (distinct : toInsert.Pairwise (fun a b => (a == b) = false)) (mem : toInsert.contains k) (mem' : containsKey k l = true) : getKeyD k (insertListIfNewUnit l toInsert) fallback = getKeyD k l fallback := by rw [getKeyD_eq_getKey?] - rw [getKey?_insertListIfNewUnit_of_mem_of_contains] + rw [getKey?_insertListIfNewUnit_of_contains_of_contains] . rw [getKeyD_eq_getKey?] · exact distinct · exact mem diff --git a/src/Std/Data/DHashMap/Internal/RawLemmas.lean b/src/Std/Data/DHashMap/Internal/RawLemmas.lean index 5fd710f8d0cd..32ea6957488e 100644 --- a/src/Std/Data/DHashMap/Internal/RawLemmas.lean +++ b/src/Std/Data/DHashMap/Internal/RawLemmas.lean @@ -1002,55 +1002,56 @@ theorem isEmpty_insertMany_list [EquivBEq α] [LawfulHashable α] (h : m.1.WF) section variable {β : Type v} (m : Raw₀ α (fun _ => β)) -theorem get?_insertMany_list_of_contains_eq_false_const [EquivBEq α] [LawfulHashable α] (h : m.1.WF) +namespace Const + +theorem get?_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l : List ((_ : α) × β)} {k : α} (h : (l.map Sigma.fst).contains k = false) : Const.get? (m.insertMany l).1 k = Const.get? m k := by simp_to_model using getValue?_insertList_of_contains_eq_false -theorem get?_insertMany_list_of_mem_const [EquivBEq α] [LawfulHashable α] (h : m.1.WF) +theorem get?_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l : List ((_ : α) × β)} {k k' : α} (k_beq : k == k') {v : β} (mem : ⟨k, v⟩ ∈ l) (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false )) : Const.get? (m.insertMany l).1 k' = v := by simp_to_model using getValue?_insertList_of_mem -theorem get_insertMany_list_of_contains_eq_false_const [EquivBEq α] [LawfulHashable α] (h : m.1.WF) +theorem get_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l : List ((_ : α) × β)} {k : α} (h₁ : (l.map Sigma.fst).contains k = false) {h'}: Const.get (m.insertMany l).1 k h' = Const.get m k (contains_of_contains_insertMany_list _ h h' h₁) := by simp_to_model using getValue_insertList_of_contains_eq_false -theorem get_insertMany_list_of_mem_const [EquivBEq α] [LawfulHashable α] (h : m.1.WF) +theorem get_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l : List ((_ : α) × β)} {k k' : α} (k_beq : k == k') {v : β} (mem : ⟨k, v⟩ ∈ l) (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false )) {h'}: Const.get (m.insertMany l).1 k' h' = v := by simp_to_model using getValue_insertList_of_mem -theorem get!_insertMany_list_of_contains_eq_false_const [EquivBEq α] [LawfulHashable α] +theorem get!_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] [Inhabited β] (h : m.1.WF) {l : List ((_ : α) × β)} {k : α} (h' : (l.map Sigma.fst).contains k = false) : Const.get! (m.insertMany l).1 k = Const.get! m k := by simp_to_model using getValue!_insertList_of_contains_eq_false -theorem get!_insertMany_list_of_mem_const [EquivBEq α] [LawfulHashable α] [Inhabited β] (h : m.1.WF) +theorem get!_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] [Inhabited β] (h : m.1.WF) {l : List ((_ : α) × β)} {k k' : α} (k_beq : k == k') {v : β} (mem: ⟨k, v⟩ ∈ l) (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false )) : Const.get! (m.insertMany l).1 k' = v := by simp_to_model using getValue!_insertList_of_mem -theorem getD_insertMany_list_of_contains_eq_false_const [EquivBEq α] [LawfulHashable α] (h : m.1.WF) +theorem getD_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l : List ((_ : α) × β)} {k : α} {fallback : β} (h' : (l.map Sigma.fst).contains k = false) : Const.getD (m.insertMany l).1 k fallback = Const.getD m k fallback:= by simp_to_model using getValueD_insertList_of_contains_eq_false -theorem getD_insertMany_list_of_mem_const [EquivBEq α] [LawfulHashable α] (h : m.1.WF) +theorem getD_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l : List ((_ : α) × β)} {k k' : α} (k_beq : k == k') {v fallback: β} (mem: ⟨k, v⟩ ∈ l) (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) : Const.getD (m.insertMany l).1 k' fallback = v := by simp_to_model using getValueD_insertList_of_mem -namespace Const @[simp] theorem insertMany_nil : insertMany m [] = m := by @@ -1068,23 +1069,23 @@ theorem insertMany_cons {l : List (α × β)} {k : α} {v : β} : | nil => simp [insertListₘ] | cons hd tl => simp [insertListₘ] -theorem contains_insertMany_list [EquivBEq α] [LawfulHashable α] (h : m.1.WF) +theorem contains_constInsertMany_list [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l : List (α × β)} {k : α} : (Const.insertMany m l).1.contains k = (m.contains k || (l.map Prod.fst).contains k) := by simp_to_model using containsKey_insertListConst -theorem contains_of_contains_insertMany_list [EquivBEq α] [LawfulHashable α] (h : m.1.WF) +theorem contains_of_contains_constInsertMany_list [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l : List ( α × β )} {k : α} : (insertMany m l).1.contains k → (l.map Prod.fst).contains k = false → m.contains k := by simp_to_model using containsKey_of_containsKey_insertListConst -theorem getKey?_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] (h : m.1.WF) +theorem getKey?_constInsertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l : List (α × β)} {k : α} (h' : (l.map Prod.fst).contains k = false) : (insertMany m l).1.getKey? k = m.getKey? k := by simp_to_model using getKey?_insertListConst_of_contains_eq_false -theorem getKey?_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] (h : m.1.WF) +theorem getKey?_constInsertMany_list_of_mem [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l : List (α × β)} {k k' : α} (k_beq : k == k') (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) @@ -1092,15 +1093,15 @@ theorem getKey?_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] (h : m. (insertMany m l).1.getKey? k' = some k := by simp_to_model using getKey?_insertListConst_of_mem -theorem getKey_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] (h : m.1.WF) +theorem getKey_constInsertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l : List (α × β)} {k : α} {h₁ : (l.map Prod.fst).contains k = false} {h'} : (insertMany m l).1.getKey k h' = - m.getKey k (contains_of_contains_insertMany_list _ h h' h₁) := by + m.getKey k (contains_of_contains_constInsertMany_list _ h h' h₁) := by simp_to_model using getKey_insertListConst_of_contains_eq_false -theorem getKey_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] (h : m.1.WF) +theorem getKey_constInsertMany_list_of_mem [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l : List (α × β)} {k k' : α} (k_beq : k == k') (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) @@ -1109,13 +1110,13 @@ theorem getKey_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] (h : m.1 (insertMany m l).1.getKey k' h' = k := by simp_to_model using getKey_insertListConst_of_mem -theorem getKey!_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] [Inhabited α] +theorem getKey!_constInsertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] [Inhabited α] (h : m.1.WF) {l : List (α × β)} {k : α} (h' : (l.map Prod.fst).contains k = false) : (insertMany m l).1.getKey! k = m.getKey! k := by simp_to_model using getKey!_insertListConst_of_contains_eq_false -theorem getKey!_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] [Inhabited α] (h : m.1.WF) +theorem getKey!_constInsertMany_list_of_mem [EquivBEq α] [LawfulHashable α] [Inhabited α] (h : m.1.WF) {l : List (α × β)} {k k' : α} (k_beq : k == k') (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) @@ -1123,13 +1124,13 @@ theorem getKey!_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] [Inhabi (insertMany m l).1.getKey! k' = k := by simp_to_model using getKey!_insertListConst_of_mem -theorem getKeyD_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] (h : m.1.WF) +theorem getKeyD_constInsertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l : List (α × β)} {k fallback : α} (h' : (l.map Prod.fst).contains k = false) : (insertMany m l).1.getKeyD k fallback = m.getKeyD k fallback := by simp_to_model using getKeyD_insertListConst_of_contains_eq_false -theorem getKeyD_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] (h : m.1.WF) +theorem getKeyD_constInsertMany_list_of_mem [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l : List (α × β)} {k k' fallback : α} (k_beq : k == k') (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) @@ -1137,67 +1138,67 @@ theorem getKeyD_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] (h : m. (insertMany m l).1.getKeyD k' fallback = k := by simp_to_model using getKeyD_insertListConst_of_mem -theorem size_insertMany_list [EquivBEq α] [LawfulHashable α] (h : m.1.WF) +theorem size_constInsertMany_list [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l : List (α × β)} {distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)} : (∀ (a : α), ¬ (m.contains a ∧ (l.map Prod.fst).contains a)) → (insertMany m l).1.1.size = m.1.size + l.length := by simp_to_model using length_insertListConst -theorem isEmpty_insertMany_list_eq_false_of_isEmpty_eq_false [EquivBEq α] [LawfulHashable α] +theorem isEmpty_constInsertMany_list_eq_false_of_isEmpty_eq_false [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l : List (α × β)} : (m.1.isEmpty = false) → (insertMany m l).1.1.isEmpty = false := by simp_to_model using isEmpty_insertListConst_eq_false_of_isEmpty_eq_false -theorem isEmpty_insertMany_list [EquivBEq α] [LawfulHashable α] (h : m.1.WF) +theorem isEmpty_constInsertMany_list [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l : List (α × β)} : (insertMany m l).1.1.isEmpty ↔ m.1.isEmpty ∧ l.isEmpty := by simp_to_model using isEmpty_insertListConst -theorem get?_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] (h : m.1.WF) +theorem get?_constInsertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l : List (α × β)} {k : α} (h' : (l.map Prod.fst).contains k = false) : get? (insertMany m l).1 k = get? m k := by simp_to_model using getValue?_insertListConst_of_contains_eq_false -theorem get?_insertMany_list_of_mem_const [EquivBEq α] [LawfulHashable α] (h : m.1.WF) +theorem get?_constInsertMany_list_of_mem [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l : List (α × β)} {k k' : α} (k_beq : k == k') {v : β} (mem: ⟨k, v⟩ ∈ l) (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false )) : get? (insertMany m l).1 k' = v := by simp_to_model using getValue?_insertListConst_of_mem -theorem get_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] (h : m.1.WF) +theorem get_constInsertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l : List (α × β)} {k : α} (h₁ : (l.map Prod.fst).contains k = false) {h'}: - get (insertMany m l).1 k h' = get m k (contains_of_contains_insertMany_list _ h h' h₁) := by + get (insertMany m l).1 k h' = get m k (contains_of_contains_constInsertMany_list _ h h' h₁) := by simp_to_model using getValue_insertListConst_of_contains_eq_false -theorem get_insertMany_list_of_mem_const [EquivBEq α] [LawfulHashable α] (h : m.1.WF) +theorem get_constInsertMany_list_of_mem [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l : List (α × β)} {k k' : α} (k_beq : k == k') {v : β} (mem: ⟨k, v⟩ ∈ l) (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false )) {h'}: get (insertMany m l).1 k' h' = v := by simp_to_model using getValue_insertListConst_of_mem -theorem get!_insertMany_list_of_contains_eq_false_const [EquivBEq α] [LawfulHashable α] +theorem get!_constInsertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] [Inhabited β] (h : m.1.WF) {l : List (α × β)} {k : α} (h' : (l.map Prod.fst).contains k = false) : get! (insertMany m l).1 k = get! m k := by simp_to_model using getValue!_insertListConst_of_contains_eq_false -theorem get!_insertMany_list_of_mem_const [EquivBEq α] [LawfulHashable α] [Inhabited β] (h : m.1.WF) +theorem get!_constInsertMany_list_of_mem [EquivBEq α] [LawfulHashable α] [Inhabited β] (h : m.1.WF) {l : List (α × β)} {k k' : α} (k_beq : k == k') {v : β} (mem: ⟨k, v⟩ ∈ l) (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false )) : get! (insertMany m l).1 k' = v := by simp_to_model using getValue!_insertListConst_of_mem -theorem getD_insertMany_list_of_contains_eq_false_const [EquivBEq α] [LawfulHashable α] (h : m.1.WF) +theorem getD_constInsertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l : List (α × β)} {k : α} {fallback : β} (h' : (l.map Prod.fst).contains k = false) : getD (insertMany m l).1 k fallback = getD m k fallback:= by simp_to_model using getValueD_insertListConst_of_contains_eq_false -theorem getD_insertMany_list_of_mem_const [EquivBEq α] [LawfulHashable α] (h : m.1.WF) +theorem getD_constInsertMany_list_of_mem [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l : List (α × β)} {k k' : α} (k_beq : k == k') {v fallback: β} (mem: ⟨k, v⟩ ∈ l) (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false )) : getD (insertMany m l).1 k' fallback = v := by @@ -1243,12 +1244,12 @@ theorem getKey?_insertManyIfNewUnit_list_of_mem_of_contains_eq_false [EquivBEq m.contains k = false → getKey? (insertManyIfNewUnit m l).1 k' = some k := by simp_to_model using getKey?_insertListIfNewUnit_of_mem_of_contains_eq_false -theorem getKey?_insertManyIfNewUnit_list_of_mem_of_contains [EquivBEq α] [LawfulHashable α] +theorem getKey?_insertManyIfNewUnit_list_of_contains_of_contains [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l : List α} {k : α} (distinct : l.Pairwise (fun a b => (a == b) = false)) (h': l.contains k) : m.contains k → getKey? (insertManyIfNewUnit m l).1 k = getKey? m k := by - simp_to_model using getKey?_insertListIfNewUnit_of_mem_of_contains + simp_to_model using getKey?_insertListIfNewUnit_of_contains_of_contains theorem getKey_insertManyIfNewUnit_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l : List α} {k : α} (h₁ : l.contains k = false) {h'} : @@ -1265,12 +1266,12 @@ theorem getKey_insertManyIfNewUnit_list_of_mem_of_contains_eq_false [EquivBEq α getKey (insertManyIfNewUnit m l).1 k' h' = k := by simp_to_model using getKey_insertListIfNewUnit_of_mem_of_contains_eq_false -theorem getKey_insertManyIfNewUnit_list_mem_of_mem_of_contains [EquivBEq α] [LawfulHashable α] +theorem getKey_insertManyIfNewUnit_list_mem_of_contains_of_contains [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l : List α} {k : α} (distinct : l.Pairwise (fun a b => (a == b) = false)) (h': l.contains k) {mem' : m.contains k} {h'} : getKey (insertManyIfNewUnit m l).1 k h' = getKey m k mem' := by - simp_to_model using getKey_insertListIfNewUnit_of_mem_of_contains + simp_to_model using getKey_insertListIfNewUnit_of_contains_of_contains theorem getKey!_insertManyIfNewUnit_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] [Inhabited α] (h : m.1.WF) {l : List α} {k : α} @@ -1285,12 +1286,12 @@ theorem getKey!_insertManyIfNewUnit_list_of_mem_of_contains_eq_false [EquivBEq getKey! (insertManyIfNewUnit m l).1 k' = k := by simp_to_model using getKey!_insertListIfNewUnit_of_mem_of_contains_eq_false -theorem getKey!_insertManyIfNewUnit_list_mem_of_mem_of_contains [EquivBEq α] [LawfulHashable α] +theorem getKey!_insertManyIfNewUnit_list_mem_of_contains_of_contains [EquivBEq α] [LawfulHashable α] [Inhabited α] (h : m.1.WF) {l : List α} {k : α} (distinct : l.Pairwise (fun a b => (a == b) = false)) (h': l.contains k) : m.contains k → getKey! (insertManyIfNewUnit m l).1 k = getKey! m k := by - simp_to_model using getKey!_insertListIfNewUnit_of_mem_of_contains + simp_to_model using getKey!_insertListIfNewUnit_of_contains_of_contains theorem getKeyD_insertManyIfNewUnit_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l : List α} {k fallback : α} @@ -1305,12 +1306,12 @@ theorem getKeyD_insertManyIfNewUnit_list_of_mem_of_contains_eq_false [EquivBEq getKeyD (insertManyIfNewUnit m l).1 k' fallback = k := by simp_to_model using getKeyD_insertListIfNewUnit_of_mem_of_contains_eq_false -theorem getKeyD_insertManyIfNewUnit_list_of_mem_of_contains [EquivBEq α] [LawfulHashable α] +theorem getKeyD_insertManyIfNewUnit_list_of_contains_of_contains [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l : List α} {k fallback : α} (distinct : l.Pairwise (fun a b => (a == b) = false)) (h': l.contains k) : m.contains k → getKeyD (insertManyIfNewUnit m l).1 k fallback = getKeyD m k fallback := by - simp_to_model using getKeyD_insertListIfNewUnit_of_mem_of_contains + simp_to_model using getKeyD_insertListIfNewUnit_of_contains_of_contains theorem size_insertManyIfNewUnit_list [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l : List α} From 3689bd63eeb52cf6455b85febca640bc195f8714 Mon Sep 17 00:00:00 2001 From: monsterkrampe Date: Thu, 5 Dec 2024 16:50:49 +0100 Subject: [PATCH 43/83] Convert negated ands into simp normal form --- .../DHashMap/Internal/List/Associative.lean | 18 ++++++++++-------- src/Std/Data/DHashMap/Internal/RawLemmas.lean | 6 +++--- 2 files changed, 13 insertions(+), 11 deletions(-) diff --git a/src/Std/Data/DHashMap/Internal/List/Associative.lean b/src/Std/Data/DHashMap/Internal/List/Associative.lean index 49e329891881..8fca1f97c03f 100644 --- a/src/Std/Data/DHashMap/Internal/List/Associative.lean +++ b/src/Std/Data/DHashMap/Internal/List/Associative.lean @@ -2175,10 +2175,9 @@ theorem perm_insertList [BEq α] [EquivBEq α] {l toInsert : List ((a : α) × β a)} (distinct_l : DistinctKeys l) (distinct_toInsert : toInsert.Pairwise (fun a b => (a.1 == b.1) = false)) - (distinct_both : ∀ (a : α), ¬ (containsKey a l ∧ (toInsert.map Sigma.fst).contains a)) : + (distinct_both : ∀ (a : α), containsKey a l -> (toInsert.map Sigma.fst).contains a = false) : Perm (insertList l toInsert) (l ++ toInsert) := by rw [← DistinctKeys.def] at distinct_toInsert - simp only [not_and, Bool.not_eq_true] at distinct_both induction toInsert generalizing l with | nil => simp only [insertList, List.append_nil, Perm.refl] | cons hd tl ih => @@ -2195,7 +2194,7 @@ theorem length_insertList [BEq α] [EquivBEq α] {l toInsert : List ((a : α) × β a)} (distinct_l : DistinctKeys l) (distinct_toInsert : toInsert.Pairwise (fun a b => (a.1 == b.1) = false)) - (distinct_both : ∀ (a : α), ¬ (containsKey a l ∧ (toInsert.map Sigma.fst).contains a)) : + (distinct_both : ∀ (a : α), containsKey a l -> (toInsert.map Sigma.fst).contains a = false) : (insertList l toInsert).length = l.length + toInsert.length := by simpa using (perm_insertList distinct_l distinct_toInsert distinct_both).length_eq @@ -2471,7 +2470,7 @@ theorem length_insertListConst [BEq α] [EquivBEq α] {l : List ((_ : α) × β)} {toInsert : List (α × β)} (distinct_l : DistinctKeys l) (distinct_toInsert : toInsert.Pairwise (fun a b => (a.1 == b.1) = false)) - (distinct_both : ∀ (a : α), ¬ (containsKey a l ∧ (toInsert.map Prod.fst).contains a)) : + (distinct_both : ∀ (a : α), containsKey a l -> (toInsert.map Prod.fst).contains a = false) : (insertListConst l toInsert).length = l.length + toInsert.length := by unfold insertListConst rw [length_insertList] @@ -2877,7 +2876,7 @@ theorem length_insertListIfNewUnit [BEq α] [EquivBEq α] {l : List ((_ : α) × Unit)} {toInsert : List α} (distinct_l : DistinctKeys l) (distinct_toInsert : toInsert.Pairwise (fun a b => (a == b) = false)) - (distinct_both : ∀ (a : α), ¬ (containsKey a l ∧ toInsert.contains a)) : + (distinct_both : ∀ (a : α), containsKey a l -> toInsert.contains a = false) : (insertListIfNewUnit l toInsert).length = l.length + toInsert.length := by induction toInsert generalizing l with | nil => simp [insertListIfNewUnit] @@ -2888,13 +2887,15 @@ theorem length_insertListIfNewUnit [BEq α] [EquivBEq α] specialize distinct_both hd simp only [List.contains_cons, BEq.refl, Bool.true_or, and_true, Bool.not_eq_true] at distinct_both - simp only [distinct_both, Bool.false_eq_true, ↓reduceIte] - rw [Nat.add_assoc, Nat.add_comm 1 _] + cases eq : containsKey hd l with + | true => simp [eq] at distinct_both + | false => + simp only [Bool.false_eq_true, ↓reduceIte] + rw [Nat.add_assoc, Nat.add_comm 1 _] · apply DistinctKeys.insertEntryIfNew distinct_l · simp only [pairwise_cons] at distinct_toInsert apply And.right distinct_toInsert · intro a - simp only [not_and, Bool.not_eq_true] simp only [List.contains_cons, Bool.or_eq_true, not_and, not_or, Bool.not_eq_true] at distinct_both rw [containsKey_insertEntryIfNew] @@ -2913,6 +2914,7 @@ theorem length_insertListIfNewUnit [BEq α] [EquivBEq α] exact left | inr h => specialize distinct_both a h + rw [Bool.or_eq_false_iff] at distinct_both apply And.right distinct_both theorem isEmpty_insertListIfNewUnit_eq_false_of_isEmpty_eq_false [BEq α] diff --git a/src/Std/Data/DHashMap/Internal/RawLemmas.lean b/src/Std/Data/DHashMap/Internal/RawLemmas.lean index 32ea6957488e..32fa01ec5c41 100644 --- a/src/Std/Data/DHashMap/Internal/RawLemmas.lean +++ b/src/Std/Data/DHashMap/Internal/RawLemmas.lean @@ -985,7 +985,7 @@ theorem getKeyD_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] (h : m. theorem size_insertMany_list [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l : List ((a : α) × β a)} {distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)} : - (∀ (a : α), ¬ (m.contains a ∧ (l.map Sigma.fst).contains a)) → + (∀ (a : α), m.contains a -> (l.map Sigma.fst).contains a = false) → (m.insertMany l).1.1.size = m.1.size + l.length := by simp_to_model using length_insertList @@ -1141,7 +1141,7 @@ theorem getKeyD_constInsertMany_list_of_mem [EquivBEq α] [LawfulHashable α] (h theorem size_constInsertMany_list [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l : List (α × β)} {distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)} : - (∀ (a : α), ¬ (m.contains a ∧ (l.map Prod.fst).contains a)) → + (∀ (a : α), m.contains a -> (l.map Prod.fst).contains a = false) → (insertMany m l).1.1.size = m.1.size + l.length := by simp_to_model using length_insertListConst @@ -1316,7 +1316,7 @@ theorem getKeyD_insertManyIfNewUnit_list_of_contains_of_contains [EquivBEq α] [ theorem size_insertManyIfNewUnit_list [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l : List α} (distinct : l.Pairwise (fun a b => (a == b) = false)): - (∀ (a : α), ¬ (m.contains a ∧ l.contains a)) → + (∀ (a : α), m.contains a -> l.contains a = false) → (insertManyIfNewUnit m l).1.1.size = m.1.size + l.length := by simp_to_model using length_insertListIfNewUnit From b8d325d24ac059180a980dbcde45c6595e2addb7 Mon Sep 17 00:00:00 2001 From: monsterkrampe Date: Tue, 10 Dec 2024 19:05:02 +0100 Subject: [PATCH 44/83] Try getEntry?_insertList approach --- .../DHashMap/Internal/List/Associative.lean | 110 +++++++++++------- 1 file changed, 69 insertions(+), 41 deletions(-) diff --git a/src/Std/Data/DHashMap/Internal/List/Associative.lean b/src/Std/Data/DHashMap/Internal/List/Associative.lean index 8fca1f97c03f..2fcf6fd35cd4 100644 --- a/src/Std/Data/DHashMap/Internal/List/Associative.lean +++ b/src/Std/Data/DHashMap/Internal/List/Associative.lean @@ -1907,6 +1907,45 @@ theorem insertList_cons_perm [BEq α] [EquivBEq α] {l₁ l₂ : List ((a : α) have := List.rel_of_pairwise_cons hl₂ (List.mem_cons_self _ _) simp [BEq.trans hh (BEq.symm hp)] at this +theorem getEntry?_insertList [BEq α] [EquivBEq α] + {l toInsert : List ((a : α) × β a)} + (distinct_l : DistinctKeys l) + (distinct_toInsert : toInsert.Pairwise (fun a b => (a.1 == b.1) = false)) (k : α) : + getEntry? k (insertList l toInsert) = (getEntry? k toInsert).or (getEntry? k l) := by + induction toInsert generalizing l with + | nil => simp [insertList] + | cons h t ih => + rw [getEntry?_of_perm distinct_l.insertList + (insertList_cons_perm distinct_l (DistinctKeys.def.2 distinct_toInsert)), + getEntry?_insertEntry] + cases hk : h.1 == k + · simp only [Bool.false_eq_true, ↓reduceIte] + rw [ih distinct_l distinct_toInsert.tail, getEntry?_cons_of_false hk] + · simp only [↓reduceIte] + rw [getEntry?_cons_of_true hk, Option.some_or] + +theorem getEntry?_insertList_of_contains_eq_false [BEq α] [PartialEquivBEq α] + {l toInsert : List ((a : α) × β a)} {k : α} + (not_contains : containsKey k toInsert = false) : + getEntry? k (insertList l toInsert) = getEntry? k l := by + induction toInsert generalizing l with + | nil => simp [insertList] + | cons h t ih => + unfold insertList + rw [containsKey_cons_eq_false] at not_contains + rw [ih not_contains.right, getEntry?_insertEntry] + simp [not_contains] + +theorem getEntry?_insertList_of_contains_eq_true [BEq α] [EquivBEq α] + {l toInsert : List ((a : α) × β a)} {k : α} + (distinct_l : DistinctKeys l) + (distinct_toInsert : toInsert.Pairwise (fun a b => (a.1 == b.1) = false)) + (contains : containsKey k toInsert = true) : + getEntry? k (insertList l toInsert) = getEntry? k toInsert := by + rw [getEntry?_insertList distinct_l distinct_toInsert] + rw [containsKey_eq_isSome_getEntry?] at contains + exact Option.or_of_isSome contains + theorem containsKey_eq_contains_map_fst [BEq α] [PartialEquivBEq α] {l : List ((a : α) × β a)} {k : α} : containsKey k l = (l.map Sigma.fst).contains k := by induction l with @@ -1940,42 +1979,31 @@ theorem getValueCast?_insertList_of_contains_eq_false [BEq α] [LawfulBEq α] (not_contains : (toInsert.map Sigma.fst).contains k = false) : getValueCast? k (insertList l toInsert) = getValueCast? k l := by rw [← containsKey_eq_contains_map_fst] at not_contains - induction toInsert generalizing l with - | nil => simp [insertList] - | cons hd tl ih => - rw [containsKey_cons, Bool.or_eq_false_iff] at not_contains - simp only [insertList] - rw [ih (And.right not_contains)] - rw [getValueCast?_insertEntry] - simp [And.left not_contains] + rw [getValueCast?_eq_getEntry?, getValueCast?_eq_getEntry?] + apply Option.dmap_congr + rw [getEntry?_insertList_of_contains_eq_false not_contains] theorem getValueCast?_insertList_of_mem [BEq α] [LawfulBEq α] {l toInsert : List ((a : α) × β a)} {k k' : α} (k_beq : k == k') (v : β k) - (distinct : toInsert.Pairwise (fun a b => (a.1 == b.1) = false)) - (distinct2 : DistinctKeys l) + (distinct_l : DistinctKeys l) + (distinct_toInsert : toInsert.Pairwise (fun a b => (a.1 == b.1) = false)) (mem : ⟨k, v⟩ ∈ toInsert) : getValueCast? k' (insertList l toInsert) = some (cast (by congr; exact LawfulBEq.eq_of_beq k_beq) v) := by - rw [← DistinctKeys.def] at distinct - induction toInsert generalizing l with - | nil => simp at mem - | cons hd tl ih => - rw [getValueCast?_of_perm distinct2.insertList (insertList_cons_perm distinct2 distinct)] - rw [getValueCast?_insertEntry] - rw [List.mem_cons] at mem - cases mem with - | inl mem => rw [← mem]; simp [k_beq] - | inr mem => - rw [dif_neg, ih distinct.tail] - . exact distinct2 - . exact mem - . have := distinct.containsKey_eq_false - rw [containsKey_eq_false_iff] at this - rw [Bool.not_eq_true] - apply BEq.neq_of_neq_of_beq - . apply this; exact mem - . simp [k_beq] + rw [getValueCast?_eq_getEntry?] + have : getEntry? k' (insertList l toInsert) = getEntry? k' toInsert := by + apply getEntry?_insertList_of_contains_eq_true distinct_l distinct_toInsert + apply containsKey_of_beq _ k_beq + exact containsKey_of_mem mem + rw [Option.dmap_congr this] + rw [← getValueCast?_eq_getEntry?] + rw [List.mem_iff_append] at mem + rcases mem with ⟨before, after, eq⟩ + rw [← DistinctKeys.def, eq] at distinct_toInsert + have perm : (before ++ ⟨k, v⟩ :: after).Perm (⟨k, v⟩ :: (before ++ after)) := List.perm_middle + rw [eq, getValueCast?_of_perm distinct_toInsert perm] + rw [getValueCast?_cons_of_true k_beq] theorem getValueCast_insertList_of_contains_eq_false [BEq α] [LawfulBEq α] {l toInsert : List ((a : α) × β a)} {k : α} @@ -1992,8 +2020,8 @@ theorem getValueCast_insertList_of_contains_eq_false [BEq α] [LawfulBEq α] theorem getValueCast_insertList_of_mem [BEq α] [LawfulBEq α] {l toInsert : List ((a : α) × β a)} {k k' : α} (k_beq : k == k') (v : β k) - (distinct : toInsert.Pairwise (fun a b => (a.1 == b.1) = false)) - (distinct2 : DistinctKeys l) + (distinct_l : DistinctKeys l) + (distinct_toInsert : toInsert.Pairwise (fun a b => (a.1 == b.1) = false)) (mem : ⟨k, v⟩ ∈ toInsert) {h} : getValueCast k' (insertList l toInsert) h = @@ -2002,8 +2030,8 @@ theorem getValueCast_insertList_of_mem [BEq α] [LawfulBEq α] rw [← getValueCast?_eq_some_getValueCast] rw [getValueCast?_insertList_of_mem] . exact k_beq - . exact distinct - . exact distinct2 + . exact distinct_l + . exact distinct_toInsert . exact mem theorem getValueCast!_insertList_of_contains_eq_false [BEq α] [LawfulBEq α] @@ -2017,15 +2045,15 @@ theorem getValueCast!_insertList_of_contains_eq_false [BEq α] [LawfulBEq α] theorem getValueCast!_insertList_of_mem [BEq α] [LawfulBEq α] {l toInsert : List ((a : α) × β a)} {k k' : α} (k_beq : k == k') (v : β k) [Inhabited (β k')] - (distinct : toInsert.Pairwise (fun a b => (a.1 == b.1) = false)) - (distinct2 : DistinctKeys l) + (distinct_l : DistinctKeys l) + (distinct_toInsert : toInsert.Pairwise (fun a b => (a.1 == b.1) = false)) (mem : ⟨k, v⟩ ∈ toInsert) : getValueCast! k' (insertList l toInsert) = cast (by congr; exact LawfulBEq.eq_of_beq k_beq) v := by rw [getValueCast!_eq_getValueCast?, getValueCast?_insertList_of_mem] . rw [Option.get!_some]; exact k_beq - . exact distinct - . exact distinct2 + . exact distinct_l + . exact distinct_toInsert . exact mem theorem getValueCastD_insertList_of_contains_eq_false [BEq α] [LawfulBEq α] @@ -2039,15 +2067,15 @@ theorem getValueCastD_insertList_of_contains_eq_false [BEq α] [LawfulBEq α] theorem getValueCastD_insertList_of_mem [BEq α] [LawfulBEq α] {l toInsert : List ((a : α) × β a)} {k k' : α} (k_beq : k == k') {v : β k} {fallback : β k'} - (distinct : toInsert.Pairwise (fun a b => (a.1 == b.1) = false)) - (distinct2 : DistinctKeys l) + (distinct_l : DistinctKeys l) + (distinct_toInsert : toInsert.Pairwise (fun a b => (a.1 == b.1) = false)) (mem : ⟨k, v⟩ ∈ toInsert) : getValueCastD k' (insertList l toInsert) fallback = cast (by congr; exact LawfulBEq.eq_of_beq k_beq) v := by rw [getValueCastD_eq_getValueCast?, getValueCast?_insertList_of_mem] . rw [Option.getD_some]; exact k_beq - . exact distinct - . exact distinct2 + . exact distinct_l + . exact distinct_toInsert . exact mem theorem getKey?_insertList_of_contains_eq_false [BEq α] [EquivBEq α] From 14a1da00103c7bb4bb1ac292ffe6f7c841819345 Mon Sep 17 00:00:00 2001 From: monsterkrampe Date: Wed, 11 Dec 2024 11:11:29 +0100 Subject: [PATCH 45/83] Use getEntry?_insertList approach where applicable --- .../DHashMap/Internal/List/Associative.lean | 227 ++++++++---------- 1 file changed, 101 insertions(+), 126 deletions(-) diff --git a/src/Std/Data/DHashMap/Internal/List/Associative.lean b/src/Std/Data/DHashMap/Internal/List/Associative.lean index 2fcf6fd35cd4..82952ee54074 100644 --- a/src/Std/Data/DHashMap/Internal/List/Associative.lean +++ b/src/Std/Data/DHashMap/Internal/List/Associative.lean @@ -1946,6 +1946,18 @@ theorem getEntry?_insertList_of_contains_eq_true [BEq α] [EquivBEq α] rw [containsKey_eq_isSome_getEntry?] at contains exact Option.or_of_isSome contains +theorem getEntry?_of_mem [BEq α] [PartialEquivBEq α] + {l : List ((a : α) × β a)} (hl : DistinctKeys l) + {k k' : α} (hk : k == k') {v : β k} (hkv : ⟨k, v⟩ ∈ l) : + getEntry? k' l = some ⟨k, v⟩ := by + induction l using assoc_induction with + | nil => simp at hkv + | cons k₁ v₁ t ih => + obtain ⟨⟨⟩⟩|hkv := List.mem_cons.1 hkv + · rw [getEntry?_cons_of_true hk] + · rw [getEntry?_cons_of_false, ih hl.tail hkv] + exact BEq.neq_of_neq_of_beq (containsKey_eq_false_iff.1 hl.containsKey_eq_false _ hkv) hk + theorem containsKey_eq_contains_map_fst [BEq α] [PartialEquivBEq α] {l : List ((a : α) × β a)} {k : α} : containsKey k l = (l.map Sigma.fst).contains k := by induction l with @@ -1996,14 +2008,10 @@ theorem getValueCast?_insertList_of_mem [BEq α] [LawfulBEq α] apply getEntry?_insertList_of_contains_eq_true distinct_l distinct_toInsert apply containsKey_of_beq _ k_beq exact containsKey_of_mem mem + rw [← DistinctKeys.def] at distinct_toInsert + rw [getEntry?_of_mem distinct_toInsert k_beq mem] at this rw [Option.dmap_congr this] - rw [← getValueCast?_eq_getEntry?] - rw [List.mem_iff_append] at mem - rcases mem with ⟨before, after, eq⟩ - rw [← DistinctKeys.def, eq] at distinct_toInsert - have perm : (before ++ ⟨k, v⟩ :: after).Perm (⟨k, v⟩ :: (before ++ after)) := List.perm_middle - rw [eq, getValueCast?_of_perm distinct_toInsert perm] - rw [getValueCast?_cons_of_true k_beq] + simp theorem getValueCast_insertList_of_contains_eq_false [BEq α] [LawfulBEq α] {l toInsert : List ((a : α) × β a)} {k : α} @@ -2083,45 +2091,29 @@ theorem getKey?_insertList_of_contains_eq_false [BEq α] [EquivBEq α] (not_contains : (toInsert.map Sigma.fst).contains k = false) : getKey? k (insertList l toInsert) = getKey? k l := by rw [← containsKey_eq_contains_map_fst] at not_contains - induction toInsert generalizing l with - | nil => simp [insertList] - | cons hd tl ih => - simp only [insertList] - rw [ih] - · rw [getKey?_insertEntry] - simp only [containsKey, Bool.or_eq_false_iff] at not_contains - simp [And.left not_contains] - · simp only [containsKey, Bool.or_eq_false_iff] at not_contains - apply And.right not_contains + rw [getKey?_eq_getEntry?, getKey?_eq_getEntry?] + rw [getEntry?_insertList_of_contains_eq_false not_contains] theorem getKey?_insertList_of_mem [BEq α] [EquivBEq α] {l toInsert : List ((a : α) × β a)} {k k' : α} (k_beq : k == k') - (distinct : toInsert.Pairwise (fun a b => (a.1 == b.1) = false)) - (distinct2 : DistinctKeys l) + (distinct_l : DistinctKeys l) + (distinct_toInsert : toInsert.Pairwise (fun a b => (a.1 == b.1) = false)) (mem : k ∈ (toInsert.map Sigma.fst)) : getKey? k' (insertList l toInsert) = some k := by - rw [← DistinctKeys.def] at distinct - induction toInsert generalizing l with - | nil => simp at mem - | cons hd tl ih => - rw [getKey?_of_perm distinct2.insertList (insertList_cons_perm distinct2 distinct)] - rw [getKey?_insertEntry] - rw [List.map_cons, List.mem_cons] at mem - cases mem with - | inl mem => simp [← mem, k_beq] - | inr mem => - rw [if_neg, ih distinct.tail] - . exact distinct2 - . exact mem - . have := distinct.containsKey_eq_false - rw [containsKey_eq_false_iff] at this - rw [Bool.not_eq_true] - rw [List.mem_map] at mem - rcases mem with ⟨_, pair_mem, pair_eq⟩ - apply BEq.neq_of_neq_of_beq - . apply this; exact pair_mem - . rw [pair_eq]; exact k_beq + rw [List.mem_map] at mem + rcases mem with ⟨pair, pair_mem, pair_eq⟩ + rw [getKey?_eq_getEntry?] + have : getEntry? k' (insertList l toInsert) = getEntry? k' toInsert := by + apply getEntry?_insertList_of_contains_eq_true distinct_l distinct_toInsert + apply containsKey_of_beq _ k_beq + rw [← pair_eq] + exact containsKey_of_mem pair_mem + rw [← DistinctKeys.def] at distinct_toInsert + rw [← pair_eq] at k_beq + rw [getEntry?_of_mem distinct_toInsert k_beq pair_mem] at this + . rw [this] + simpa theorem getKey_insertList_of_contains_eq_false [BEq α] [EquivBEq α] {l toInsert : List ((a : α) × β a)} {k : α} @@ -2138,8 +2130,8 @@ theorem getKey_insertList_of_contains_eq_false [BEq α] [EquivBEq α] theorem getKey_insertList_of_mem [BEq α] [EquivBEq α] {l toInsert : List ((a : α) × β a)} {k k' : α} (k_beq : k == k') - (distinct : toInsert.Pairwise (fun a b => (a.1 == b.1) = false)) - (distinct2 : DistinctKeys l) + (distinct_l : DistinctKeys l) + (distinct_toInsert : toInsert.Pairwise (fun a b => (a.1 == b.1) = false)) (mem : k ∈ (toInsert.map Sigma.fst)) {h} : getKey k' (insertList l toInsert) h = k := by @@ -2147,8 +2139,8 @@ theorem getKey_insertList_of_mem [BEq α] [EquivBEq α] rw [← getKey?_eq_some_getKey] rw [getKey?_insertList_of_mem] . exact k_beq - . exact distinct - . exact distinct2 + . exact distinct_l + . exact distinct_toInsert . exact mem theorem getKey!_insertList_of_contains_eq_false [BEq α] [EquivBEq α] [Inhabited α] @@ -2163,16 +2155,16 @@ theorem getKey!_insertList_of_contains_eq_false [BEq α] [EquivBEq α] [Inhabite theorem getKey!_insertList_of_mem [BEq α] [EquivBEq α] [Inhabited α] {l toInsert : List ((a : α) × β a)} {k k' : α} (k_beq : k == k') - (distinct : toInsert.Pairwise (fun a b => (a.1 == b.1) = false)) - (distinct2 : DistinctKeys l) + (distinct_l : DistinctKeys l) + (distinct_toInsert : toInsert.Pairwise (fun a b => (a.1 == b.1) = false)) (mem : k ∈ (toInsert.map Sigma.fst)) : getKey! k' (insertList l toInsert) = k := by rw [getKey!_eq_getKey?] rw [getKey?_insertList_of_mem] . rw [Option.get!_some] . exact k_beq - . exact distinct - . exact distinct2 + . exact distinct_l + . exact distinct_toInsert . exact mem theorem getKeyD_insertList_of_contains_eq_false [BEq α] [EquivBEq α] @@ -2187,16 +2179,16 @@ theorem getKeyD_insertList_of_contains_eq_false [BEq α] [EquivBEq α] theorem getKeyD_insertList_of_mem [BEq α] [EquivBEq α] {l toInsert : List ((a : α) × β a)} {k k' fallback : α} (k_beq : k == k') - (distinct : toInsert.Pairwise (fun a b => (a.1 == b.1) = false)) - (distinct2 : DistinctKeys l) + (distinct_l : DistinctKeys l) + (distinct_toInsert : toInsert.Pairwise (fun a b => (a.1 == b.1) = false)) (mem : k ∈ (toInsert.map Sigma.fst)) : getKeyD k' (insertList l toInsert) fallback = k := by rw [getKeyD_eq_getKey?] rw [getKey?_insertList_of_mem] . rw [Option.getD_some] . exact k_beq - . exact distinct - . exact distinct2 + . exact distinct_l + . exact distinct_toInsert . exact mem theorem perm_insertList [BEq α] [EquivBEq α] @@ -2252,44 +2244,28 @@ variable {β : Type v} theorem getValue?_insertList_of_contains_eq_false [BEq α] [PartialEquivBEq α] {l toInsert : List ((_ : α) × β)} {k : α} - (h : (toInsert.map Sigma.fst).contains k = false) : + (not_contains : (toInsert.map Sigma.fst).contains k = false) : getValue? k (insertList l toInsert) = getValue? k l := by - rw [← containsKey_eq_contains_map_fst] at h - induction toInsert generalizing l with - | nil => simp [insertList] - | cons hd tl ih => - rw [containsKey_cons, Bool.or_eq_false_iff] at h - simp only [insertList] - rw [ih (And.right h)] - rw [getValue?_insertEntry] - simp [And.left h] + rw [← containsKey_eq_contains_map_fst] at not_contains + rw [getValue?_eq_getEntry?, getValue?_eq_getEntry?] + rw [getEntry?_insertList_of_contains_eq_false not_contains] theorem getValue?_insertList_of_mem [BEq α] [EquivBEq α] {l toInsert : List ((_ : α) × β )} {k k' : α} (k_beq : k == k') {v : β} - (distinct : toInsert.Pairwise (fun a b => (a.1 == b.1) = false)) - (distinct2 : DistinctKeys l) + (distinct_l : DistinctKeys l) + (distinct_toInsert : toInsert.Pairwise (fun a b => (a.1 == b.1) = false)) (mem : ⟨k, v⟩ ∈ toInsert) : getValue? k' (insertList l toInsert) = v := by - rw [← DistinctKeys.def] at distinct - induction toInsert generalizing l with - | nil => simp at mem - | cons hd tl ih => - rw [getValue?_of_perm distinct2.insertList (insertList_cons_perm distinct2 distinct)] - rw [getValue?_insertEntry] - rw [List.mem_cons] at mem - cases mem with - | inl mem => rw [← mem]; simp [k_beq] - | inr mem => - rw [if_neg, ih distinct.tail] - . exact distinct2 - . exact mem - . have := distinct.containsKey_eq_false - rw [containsKey_eq_false_iff] at this - rw [Bool.not_eq_true] - apply BEq.neq_of_neq_of_beq - . apply this; exact mem - . simp [k_beq] + rw [getValue?_eq_getEntry?] + have : getEntry? k' (insertList l toInsert) = getEntry? k' toInsert := by + apply getEntry?_insertList_of_contains_eq_true distinct_l distinct_toInsert + apply containsKey_of_beq _ k_beq + exact containsKey_of_mem mem + rw [← DistinctKeys.def] at distinct_toInsert + rw [getEntry?_of_mem distinct_toInsert k_beq mem] at this + rw [this] + simp theorem getValue_insertList_of_contains_eq_false [BEq α] [PartialEquivBEq α] {l toInsert : List ((_ : α) × β)} {k : α} @@ -2306,8 +2282,8 @@ theorem getValue_insertList_of_contains_eq_false [BEq α] [PartialEquivBEq α] theorem getValue_insertList_of_mem [BEq α] [EquivBEq α] {l toInsert : List ((_ : α) × β)} {k k' : α} (k_beq : k == k') {v : β} - (distinct : toInsert.Pairwise (fun a b => (a.1 == b.1) = false)) - (distinct2 : DistinctKeys l) + (distinct_l : DistinctKeys l) + (distinct_toInsert : toInsert.Pairwise (fun a b => (a.1 == b.1) = false)) (mem : ⟨k, v⟩ ∈ toInsert) {h} : getValue k' (insertList l toInsert) h = v := by @@ -2315,8 +2291,8 @@ theorem getValue_insertList_of_mem [BEq α] [EquivBEq α] rw [← getValue?_eq_some_getValue] rw [getValue?_insertList_of_mem] . exact k_beq - . exact distinct - . exact distinct2 + . exact distinct_l + . exact distinct_toInsert . exact mem theorem getValue!_insertList_of_contains_eq_false [BEq α] [PartialEquivBEq α] [Inhabited β] @@ -2329,16 +2305,16 @@ theorem getValue!_insertList_of_contains_eq_false [BEq α] [PartialEquivBEq α] theorem getValue!_insertList_of_mem [BEq α] [EquivBEq α] [Inhabited β] {l toInsert : List ((_ : α) × β)} {k k' : α} {v: β} (k_beq : k == k') - (distinct : toInsert.Pairwise (fun a b => (a.1 == b.1) = false)) - (distinct2 : DistinctKeys l) + (distinct_l : DistinctKeys l) + (distinct_toInsert : toInsert.Pairwise (fun a b => (a.1 == b.1) = false)) (mem : ⟨k, v⟩ ∈ toInsert): getValue! k' (insertList l toInsert) = v := by simp only [getValue!_eq_getValue?] rw [getValue?_insertList_of_mem] · rw [Option.get!_some] · exact k_beq - · exact distinct - . exact distinct2 + · exact distinct_l + . exact distinct_toInsert . exact mem theorem getValueD_insertList_of_contains_eq_false [BEq α] [PartialEquivBEq α] @@ -2351,19 +2327,18 @@ theorem getValueD_insertList_of_contains_eq_false [BEq α] [PartialEquivBEq α] theorem getValueD_insertList_of_mem [BEq α] [EquivBEq α] {l toInsert : List ((_ : α) × β)} {k k' : α} {v fallback: β} (k_beq : k == k') - (distinct : toInsert.Pairwise (fun a b => (a.1 == b.1) = false)) - (distinct2 : DistinctKeys l) + (distinct_l : DistinctKeys l) + (distinct_toInsert : toInsert.Pairwise (fun a b => (a.1 == b.1) = false)) (mem : ⟨k, v⟩ ∈ toInsert): getValueD k' (insertList l toInsert) fallback= v := by simp only [getValueD_eq_getValue?] rw [getValue?_insertList_of_mem] · rw [Option.getD_some] · exact k_beq - · exact distinct - . exact distinct2 + · exact distinct_l + . exact distinct_toInsert . exact mem - -- TODO: this should be unified with Mathlib... /-- Internal implementation detail of the hash map -/ def Prod.toSigma_HashMapInternal (p : α × β) : ((_ : α) × β) := ⟨p.fst, p.snd⟩ @@ -2406,15 +2381,15 @@ theorem getKey?_insertListConst_of_contains_eq_false [BEq α] [EquivBEq α] theorem getKey?_insertListConst_of_mem [BEq α] [EquivBEq α] {l : List ((_ : α) × β)} {toInsert : List (α × β)} {k k' : α} (k_beq : k == k') - (distinct : toInsert.Pairwise (fun a b => (a.1 == b.1) = false)) - (distinct2 : DistinctKeys l) + (distinct_l : DistinctKeys l) + (distinct_toInsert : toInsert.Pairwise (fun a b => (a.1 == b.1) = false)) (mem : k ∈ (toInsert.map Prod.fst)) : getKey? k' (insertListConst l toInsert) = some k := by unfold insertListConst apply getKey?_insertList_of_mem · exact k_beq + · exact distinct_l · simpa [List.pairwise_map] - · exact distinct2 · simp only [List.map_map, Prod.fst_comp_toSigma_HashMapInternal] exact mem @@ -2433,8 +2408,8 @@ theorem getKey_insertListConst_of_contains_eq_false [BEq α] [EquivBEq α] theorem getKey_insertListConst_of_mem [BEq α] [EquivBEq α] {l : List ((_ : α) × β)} {toInsert : List (α × β)} {k k' : α} (k_beq : k == k') - (distinct : toInsert.Pairwise (fun a b => (a.1 == b.1) = false)) - (distinct2 : DistinctKeys l) + (distinct_l : DistinctKeys l) + (distinct_toInsert : toInsert.Pairwise (fun a b => (a.1 == b.1) = false)) (mem : k ∈ (toInsert.map Prod.fst)) {h} : getKey k' (insertListConst l toInsert) h = k := by @@ -2442,8 +2417,8 @@ theorem getKey_insertListConst_of_mem [BEq α] [EquivBEq α] rw [← getKey?_eq_some_getKey] rw [getKey?_insertListConst_of_mem] . exact k_beq - . exact distinct - . exact distinct2 + . exact distinct_l + . exact distinct_toInsert . exact mem theorem getKey!_insertListConst_of_contains_eq_false [BEq α] [EquivBEq α] [Inhabited α] @@ -2458,16 +2433,16 @@ theorem getKey!_insertListConst_of_contains_eq_false [BEq α] [EquivBEq α] [Inh theorem getKey!_insertListConst_of_mem [BEq α] [EquivBEq α] [Inhabited α] {l : List ((_ : α) × β)} {toInsert : List (α × β)} {k k' : α} (k_beq : k == k') - (distinct : toInsert.Pairwise (fun a b => (a.1 == b.1) = false)) - (distinct2 : DistinctKeys l) + (distinct_l : DistinctKeys l) + (distinct_toInsert : toInsert.Pairwise (fun a b => (a.1 == b.1) = false)) (mem : k ∈ (toInsert.map Prod.fst)) : getKey! k' (insertListConst l toInsert) = k := by rw [getKey!_eq_getKey?] rw [getKey?_insertListConst_of_mem] . rw [Option.get!_some] . exact k_beq - . exact distinct - . exact distinct2 + . exact distinct_l + . exact distinct_toInsert . exact mem theorem getKeyD_insertListConst_of_contains_eq_false [BEq α] [EquivBEq α] @@ -2482,16 +2457,16 @@ theorem getKeyD_insertListConst_of_contains_eq_false [BEq α] [EquivBEq α] theorem getKeyD_insertListConst_of_mem [BEq α] [EquivBEq α] {l : List ((_ : α) × β)} {toInsert : List (α × β)} {k k' fallback : α} (k_beq : k == k') - (distinct : toInsert.Pairwise (fun a b => (a.1 == b.1) = false)) - (distinct2 : DistinctKeys l) + (distinct_l : DistinctKeys l) + (distinct_toInsert : toInsert.Pairwise (fun a b => (a.1 == b.1) = false)) (mem : k ∈ (toInsert.map Prod.fst)) : getKeyD k' (insertListConst l toInsert) fallback = k := by rw [getKeyD_eq_getKey?] rw [getKey?_insertListConst_of_mem] . rw [Option.getD_some] . exact k_beq - . exact distinct - . exact distinct2 + . exact distinct_l + . exact distinct_toInsert . exact mem theorem length_insertListConst [BEq α] [EquivBEq α] @@ -2535,14 +2510,14 @@ theorem getValue?_insertListConst_of_contains_eq_false [BEq α] [PartialEquivBEq theorem getValue?_insertListConst_of_mem [BEq α] [EquivBEq α] {l : List ((_ : α) × β)} {toInsert : List (α × β)} {k k' : α} (k_beq : k == k') {v : β} - (distinct : toInsert.Pairwise (fun a b => (a.1 == b.1) = false)) - (distinct2 : DistinctKeys l) + (distinct_l : DistinctKeys l) + (distinct_toInsert : toInsert.Pairwise (fun a b => (a.1 == b.1) = false)) (mem : ⟨k, v⟩ ∈ toInsert) : getValue? k' (insertListConst l toInsert) = v := by unfold insertListConst apply getValue?_insertList_of_mem (k_beq:=k_beq) + · exact distinct_l · simpa [List.pairwise_map] - · exact distinct2 . simp only [List.mem_map] exists (k, v) @@ -2561,8 +2536,8 @@ theorem getValue_insertListConst_of_contains_eq_false [BEq α] [PartialEquivBEq theorem getValue_insertListConst_of_mem [BEq α] [EquivBEq α] {l : List ((_ : α) × β)} {toInsert : List (α × β)} {k k' : α} (k_beq : k == k') {v : β} - (distinct : toInsert.Pairwise (fun a b => (a.1 == b.1) = false)) - (distinct2 : DistinctKeys l) + (distinct_l : DistinctKeys l) + (distinct_toInsert : toInsert.Pairwise (fun a b => (a.1 == b.1) = false)) (mem : ⟨k, v⟩ ∈ toInsert) {h} : getValue k' (insertListConst l toInsert) h = v := by @@ -2570,8 +2545,8 @@ theorem getValue_insertListConst_of_mem [BEq α] [EquivBEq α] rw [← getValue?_eq_some_getValue] rw [getValue?_insertListConst_of_mem (mem:=mem)] . exact k_beq - . exact distinct - . exact distinct2 + . exact distinct_l + . exact distinct_toInsert theorem getValue!_insertListConst_of_contains_eq_false [BEq α] [PartialEquivBEq α] [Inhabited β] {l : List ((_ : α) × β)} {toInsert : List (α × β)} {k : α} @@ -2583,16 +2558,16 @@ theorem getValue!_insertListConst_of_contains_eq_false [BEq α] [PartialEquivBEq theorem getValue!_insertListConst_of_mem [BEq α] [EquivBEq α] [Inhabited β] {l : List ((_ : α) × β)} {toInsert : List (α × β)} {k k' : α} {v: β} (k_beq : k == k') - (distinct : toInsert.Pairwise (fun a b => (a.1 == b.1) = false)) - (distinct2 : DistinctKeys l) + (distinct_l : DistinctKeys l) + (distinct_toInsert : toInsert.Pairwise (fun a b => (a.1 == b.1) = false)) (mem : ⟨k, v⟩ ∈ toInsert): getValue! k' (insertListConst l toInsert) = v := by simp only [getValue!_eq_getValue?] rw [getValue?_insertListConst_of_mem] · rw [Option.get!_some] · exact k_beq - · exact distinct - . exact distinct2 + · exact distinct_l + . exact distinct_toInsert . exact mem theorem getValueD_insertListConst_of_contains_eq_false [BEq α] [PartialEquivBEq α] @@ -2605,16 +2580,16 @@ theorem getValueD_insertListConst_of_contains_eq_false [BEq α] [PartialEquivBEq theorem getValueD_insertListConst_of_mem [BEq α] [EquivBEq α] {l : List ((_ : α) × β)} {toInsert : List (α × β)} {k k' : α} {v fallback: β} (k_beq : k == k') - (distinct : toInsert.Pairwise (fun a b => (a.1 == b.1) = false)) - (distinct2 : DistinctKeys l) + (distinct_l : DistinctKeys l) + (distinct_toInsert : toInsert.Pairwise (fun a b => (a.1 == b.1) = false)) (mem : ⟨k, v⟩ ∈ toInsert): getValueD k' (insertListConst l toInsert) fallback= v := by simp only [getValueD_eq_getValue?] rw [getValue?_insertListConst_of_mem] · rw [Option.getD_some] · exact k_beq - · exact distinct - . exact distinct2 + · exact distinct_l + . exact distinct_toInsert . exact mem /-- Internal implementation detail of the hash map -/ From cdc6b59eca629603f49e0d22c288c51c8da45d33 Mon Sep 17 00:00:00 2001 From: monsterkrampe Date: Wed, 11 Dec 2024 11:37:00 +0100 Subject: [PATCH 46/83] Simplify isEmpty results; Fix formatting --- .../DHashMap/Internal/List/Associative.lean | 59 ++++------------- src/Std/Data/DHashMap/Internal/RawLemmas.lean | 63 +++++++------------ 2 files changed, 37 insertions(+), 85 deletions(-) diff --git a/src/Std/Data/DHashMap/Internal/List/Associative.lean b/src/Std/Data/DHashMap/Internal/List/Associative.lean index 82952ee54074..aba78f009d9e 100644 --- a/src/Std/Data/DHashMap/Internal/List/Associative.lean +++ b/src/Std/Data/DHashMap/Internal/List/Associative.lean @@ -2195,7 +2195,7 @@ theorem perm_insertList [BEq α] [EquivBEq α] {l toInsert : List ((a : α) × β a)} (distinct_l : DistinctKeys l) (distinct_toInsert : toInsert.Pairwise (fun a b => (a.1 == b.1) = false)) - (distinct_both : ∀ (a : α), containsKey a l -> (toInsert.map Sigma.fst).contains a = false) : + (distinct_both : ∀ (a : α), containsKey a l → (toInsert.map Sigma.fst).contains a = false) : Perm (insertList l toInsert) (l ++ toInsert) := by rw [← DistinctKeys.def] at distinct_toInsert induction toInsert generalizing l with @@ -2214,29 +2214,17 @@ theorem length_insertList [BEq α] [EquivBEq α] {l toInsert : List ((a : α) × β a)} (distinct_l : DistinctKeys l) (distinct_toInsert : toInsert.Pairwise (fun a b => (a.1 == b.1) = false)) - (distinct_both : ∀ (a : α), containsKey a l -> (toInsert.map Sigma.fst).contains a = false) : + (distinct_both : ∀ (a : α), containsKey a l → (toInsert.map Sigma.fst).contains a = false) : (insertList l toInsert).length = l.length + toInsert.length := by simpa using (perm_insertList distinct_l distinct_toInsert distinct_both).length_eq -theorem isEmpty_insertList_eq_false_of_isEmpty_eq_false [BEq α] - {l toInsert : List ((a : α) × β a)} - (h : l.isEmpty = false) : - (List.insertList l toInsert).isEmpty = false := by - induction toInsert generalizing l with - | nil => simp [insertList, h] - | cons hd tl ih => - simp [insertList, ih] - theorem isEmpty_insertList [BEq α] {l toInsert : List ((a : α) × β a)} : - (List.insertList l toInsert).isEmpty ↔ l.isEmpty ∧ toInsert.isEmpty := by - induction toInsert with + (List.insertList l toInsert).isEmpty = (l.isEmpty && toInsert.isEmpty) := by + induction toInsert generalizing l with | nil => simp [insertList] | cons hd tl ih => - simp only [insertList, List.isEmpty_cons, Bool.false_eq_true, and_false, - iff_false] - apply ne_true_of_eq_false - apply isEmpty_insertList_eq_false_of_isEmpty_eq_false + rw [insertList, List.isEmpty_cons, ih, isEmpty_insertEntry] simp section @@ -2473,7 +2461,7 @@ theorem length_insertListConst [BEq α] [EquivBEq α] {l : List ((_ : α) × β)} {toInsert : List (α × β)} (distinct_l : DistinctKeys l) (distinct_toInsert : toInsert.Pairwise (fun a b => (a.1 == b.1) = false)) - (distinct_both : ∀ (a : α), containsKey a l -> (toInsert.map Prod.fst).contains a = false) : + (distinct_both : ∀ (a : α), containsKey a l → (toInsert.map Prod.fst).contains a = false) : (insertListConst l toInsert).length = l.length + toInsert.length := by unfold insertListConst rw [length_insertList] @@ -2482,22 +2470,13 @@ theorem length_insertListConst [BEq α] [EquivBEq α] · simpa [List.pairwise_map] · simpa using distinct_both -theorem isEmpty_insertListConst_eq_false_of_isEmpty_eq_false [BEq α] - {l : List ((_ : α) × β)} {toInsert : List (α × β)} - (h : l.isEmpty = false) : - (List.insertListConst l toInsert).isEmpty = false := by - induction toInsert generalizing l with - | nil => simp [insertListConst, insertList, h] - | cons hd tl ih => - unfold insertListConst at ih - simp [insertListConst, insertList, ih] - theorem isEmpty_insertListConst [BEq α] {l : List ((_ : α) × β)} {toInsert : List (α × β)} : - (List.insertListConst l toInsert).isEmpty ↔ l.isEmpty ∧ toInsert.isEmpty := by + (List.insertListConst l toInsert).isEmpty = (l.isEmpty && toInsert.isEmpty) := by unfold insertListConst rw [isEmpty_insertList] - simp only [List.isEmpty_eq_true, List.map_eq_nil_iff] + rw [Bool.eq_iff_iff] + simp theorem getValue?_insertListConst_of_contains_eq_false [BEq α] [PartialEquivBEq α] {l : List ((_ : α) × β)} {toInsert : List (α × β)} {k : α} @@ -2879,7 +2858,7 @@ theorem length_insertListIfNewUnit [BEq α] [EquivBEq α] {l : List ((_ : α) × Unit)} {toInsert : List α} (distinct_l : DistinctKeys l) (distinct_toInsert : toInsert.Pairwise (fun a b => (a == b) = false)) - (distinct_both : ∀ (a : α), containsKey a l -> toInsert.contains a = false) : + (distinct_both : ∀ (a : α), containsKey a l → toInsert.contains a = false) : (insertListIfNewUnit l toInsert).length = l.length + toInsert.length := by induction toInsert generalizing l with | nil => simp [insertListIfNewUnit] @@ -2920,25 +2899,13 @@ theorem length_insertListIfNewUnit [BEq α] [EquivBEq α] rw [Bool.or_eq_false_iff] at distinct_both apply And.right distinct_both -theorem isEmpty_insertListIfNewUnit_eq_false_of_isEmpty_eq_false [BEq α] - {l : List ((_ : α) × Unit)} {toInsert : List α} - (h : l.isEmpty = false) : - (List.insertListIfNewUnit l toInsert).isEmpty = false := by - induction toInsert generalizing l with - | nil => simp [insertListIfNewUnit, h] - | cons hd tl ih => - simp [insertListIfNewUnit, ih] - theorem isEmpty_insertListIfNewUnit [BEq α] {l : List ((_ : α) × Unit)} {toInsert : List α} : - (List.insertListIfNewUnit l toInsert).isEmpty ↔ l.isEmpty ∧ toInsert.isEmpty := by - induction toInsert with + (List.insertListIfNewUnit l toInsert).isEmpty = (l.isEmpty && toInsert.isEmpty) := by + induction toInsert generalizing l with | nil => simp [insertListIfNewUnit] | cons hd tl ih => - simp only [insertListIfNewUnit, List.isEmpty_cons, Bool.false_eq_true, and_false, - iff_false] - apply ne_true_of_eq_false - apply isEmpty_insertListIfNewUnit_eq_false_of_isEmpty_eq_false + rw [insertListIfNewUnit, List.isEmpty_cons, ih, isEmpty_insertEntryIfNew] simp theorem getValue?_list_unit [BEq α] {l : List ((_ : α) × Unit)} {k : α}: diff --git a/src/Std/Data/DHashMap/Internal/RawLemmas.lean b/src/Std/Data/DHashMap/Internal/RawLemmas.lean index 32fa01ec5c41..b896a45f63b0 100644 --- a/src/Std/Data/DHashMap/Internal/RawLemmas.lean +++ b/src/Std/Data/DHashMap/Internal/RawLemmas.lean @@ -934,7 +934,7 @@ theorem getKey?_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] (h : m. {l : List ((a : α) × β a)} {k k' : α} (k_beq : k == k') (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) - (mem : k ∈ (l.map Sigma.fst)) : + (mem : k ∈ l.map Sigma.fst) : (m.insertMany l).1.getKey? k' = some k := by simp_to_model using getKey?_insertList_of_mem @@ -950,7 +950,7 @@ theorem getKey_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] (h : m.1 {l : List ((a : α) × β a)} {k k' : α} (k_beq : k == k') (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) - (mem : k ∈ (l.map Sigma.fst)) + (mem : k ∈ l.map Sigma.fst) {h'} : (m.insertMany l).1.getKey k' h' = k := by simp_to_model using getKey_insertList_of_mem @@ -965,7 +965,7 @@ theorem getKey!_insertMany_listof_mem [EquivBEq α] [LawfulHashable α] [Inhabit {l : List ((a : α) × β a)} {k k' : α} (k_beq : k == k') (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) - (mem : k ∈ (l.map Sigma.fst)) : + (mem : k ∈ l.map Sigma.fst) : (m.insertMany l).1.getKey! k' = k := by simp_to_model using getKey!_insertList_of_mem @@ -979,24 +979,19 @@ theorem getKeyD_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] (h : m. {l : List ((a : α) × β a)} {k k' fallback : α} (k_beq : k == k') (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) - (mem : k ∈ (l.map Sigma.fst)) : + (mem : k ∈ l.map Sigma.fst) : (m.insertMany l).1.getKeyD k' fallback = k := by simp_to_model using getKeyD_insertList_of_mem theorem size_insertMany_list [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l : List ((a : α) × β a)} {distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)} : - (∀ (a : α), m.contains a -> (l.map Sigma.fst).contains a = false) → + (∀ (a : α), m.contains a → (l.map Sigma.fst).contains a = false) → (m.insertMany l).1.1.size = m.1.size + l.length := by simp_to_model using length_insertList -theorem isEmpty_insertMany_list_eq_false_of_isEmpty_eq_false [EquivBEq α] [LawfulHashable α] - (h : m.1.WF) {l : List ((a : α) × β a)}: - (m.1.isEmpty = false) → (m.insertMany l).1.1.isEmpty = false := by - simp_to_model using isEmpty_insertList_eq_false_of_isEmpty_eq_false - theorem isEmpty_insertMany_list [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l : List ((a : α) × β a)} : - (m.insertMany l).1.1.isEmpty ↔ m.1.isEmpty ∧ l.isEmpty := by + (m.insertMany l).1.1.isEmpty = (m.1.isEmpty && l.isEmpty) := by simp_to_model using isEmpty_insertList section @@ -1011,7 +1006,7 @@ theorem get?_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable theorem get?_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l : List ((_ : α) × β)} {k k' : α} (k_beq : k == k') {v : β} (mem : ⟨k, v⟩ ∈ l) - (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false )) : + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) : Const.get? (m.insertMany l).1 k' = v := by simp_to_model using getValue?_insertList_of_mem @@ -1024,7 +1019,7 @@ theorem get_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable theorem get_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l : List ((_ : α) × β)} {k k' : α} (k_beq : k == k') {v : β} (mem : ⟨k, v⟩ ∈ l) - (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false )) {h'}: + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) {h'}: Const.get (m.insertMany l).1 k' h' = v := by simp_to_model using getValue_insertList_of_mem @@ -1036,18 +1031,18 @@ theorem get!_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable theorem get!_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] [Inhabited β] (h : m.1.WF) {l : List ((_ : α) × β)} {k k' : α} (k_beq : k == k') {v : β} (mem: ⟨k, v⟩ ∈ l) - (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false )) : + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) : Const.get! (m.insertMany l).1 k' = v := by simp_to_model using getValue!_insertList_of_mem theorem getD_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l : List ((_ : α) × β)} {k : α} {fallback : β} (h' : (l.map Sigma.fst).contains k = false) : - Const.getD (m.insertMany l).1 k fallback = Const.getD m k fallback:= by + Const.getD (m.insertMany l).1 k fallback = Const.getD m k fallback := by simp_to_model using getValueD_insertList_of_contains_eq_false theorem getD_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] (h : m.1.WF) - {l : List ((_ : α) × β)} {k k' : α} (k_beq : k == k') {v fallback: β} (mem: ⟨k, v⟩ ∈ l) + {l : List ((_ : α) × β)} {k k' : α} (k_beq : k == k') {v fallback : β} (mem : ⟨k, v⟩ ∈ l) (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) : Const.getD (m.insertMany l).1 k' fallback = v := by simp_to_model using getValueD_insertList_of_mem @@ -1105,7 +1100,7 @@ theorem getKey_constInsertMany_list_of_mem [EquivBEq α] [LawfulHashable α] (h {l : List (α × β)} {k k' : α} (k_beq : k == k') (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) - (mem : k ∈ (l.map Prod.fst)) + (mem : k ∈ l.map Prod.fst) {h'} : (insertMany m l).1.getKey k' h' = k := by simp_to_model using getKey_insertListConst_of_mem @@ -1120,7 +1115,7 @@ theorem getKey!_constInsertMany_list_of_mem [EquivBEq α] [LawfulHashable α] [I {l : List (α × β)} {k k' : α} (k_beq : k == k') (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) - (mem : k ∈ (l.map Prod.fst)) : + (mem : k ∈ l.map Prod.fst) : (insertMany m l).1.getKey! k' = k := by simp_to_model using getKey!_insertListConst_of_mem @@ -1134,25 +1129,20 @@ theorem getKeyD_constInsertMany_list_of_mem [EquivBEq α] [LawfulHashable α] (h {l : List (α × β)} {k k' fallback : α} (k_beq : k == k') (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) - (mem : k ∈ (l.map Prod.fst)) : + (mem : k ∈ l.map Prod.fst) : (insertMany m l).1.getKeyD k' fallback = k := by simp_to_model using getKeyD_insertListConst_of_mem theorem size_constInsertMany_list [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l : List (α × β)} {distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)} : - (∀ (a : α), m.contains a -> (l.map Prod.fst).contains a = false) → + (∀ (a : α), m.contains a → (l.map Prod.fst).contains a = false) → (insertMany m l).1.1.size = m.1.size + l.length := by simp_to_model using length_insertListConst -theorem isEmpty_constInsertMany_list_eq_false_of_isEmpty_eq_false [EquivBEq α] [LawfulHashable α] - (h : m.1.WF) {l : List (α × β)} : - (m.1.isEmpty = false) → (insertMany m l).1.1.isEmpty = false := by - simp_to_model using isEmpty_insertListConst_eq_false_of_isEmpty_eq_false - theorem isEmpty_constInsertMany_list [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l : List (α × β)} : - (insertMany m l).1.1.isEmpty ↔ m.1.isEmpty ∧ l.isEmpty := by + (insertMany m l).1.1.isEmpty = (m.1.isEmpty && l.isEmpty) := by simp_to_model using isEmpty_insertListConst theorem get?_constInsertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] (h : m.1.WF) @@ -1163,7 +1153,7 @@ theorem get?_constInsertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHash theorem get?_constInsertMany_list_of_mem [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l : List (α × β)} {k k' : α} (k_beq : k == k') {v : β} (mem: ⟨k, v⟩ ∈ l) - (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false )) : + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) : get? (insertMany m l).1 k' = v := by simp_to_model using getValue?_insertListConst_of_mem @@ -1176,7 +1166,7 @@ theorem get_constInsertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHasha theorem get_constInsertMany_list_of_mem [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l : List (α × β)} {k k' : α} (k_beq : k == k') {v : β} (mem: ⟨k, v⟩ ∈ l) - (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false )) {h'}: + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) {h'}: get (insertMany m l).1 k' h' = v := by simp_to_model using getValue_insertListConst_of_mem @@ -1188,19 +1178,19 @@ theorem get!_constInsertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHash theorem get!_constInsertMany_list_of_mem [EquivBEq α] [LawfulHashable α] [Inhabited β] (h : m.1.WF) {l : List (α × β)} {k k' : α} (k_beq : k == k') {v : β} (mem: ⟨k, v⟩ ∈ l) - (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false )) : + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) : get! (insertMany m l).1 k' = v := by simp_to_model using getValue!_insertListConst_of_mem theorem getD_constInsertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l : List (α × β)} {k : α} {fallback : β} (h' : (l.map Prod.fst).contains k = false) : - getD (insertMany m l).1 k fallback = getD m k fallback:= by + getD (insertMany m l).1 k fallback = getD m k fallback := by simp_to_model using getValueD_insertListConst_of_contains_eq_false theorem getD_constInsertMany_list_of_mem [EquivBEq α] [LawfulHashable α] (h : m.1.WF) - {l : List (α × β)} {k k' : α} (k_beq : k == k') {v fallback: β} (mem: ⟨k, v⟩ ∈ l) - (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false )) : + {l : List (α × β)} {k k' : α} (k_beq : k == k') {v fallback : β} (mem : ⟨k, v⟩ ∈ l) + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) : getD (insertMany m l).1 k' fallback = v := by simp_to_model using getValueD_insertListConst_of_mem @@ -1316,18 +1306,13 @@ theorem getKeyD_insertManyIfNewUnit_list_of_contains_of_contains [EquivBEq α] [ theorem size_insertManyIfNewUnit_list [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l : List α} (distinct : l.Pairwise (fun a b => (a == b) = false)): - (∀ (a : α), m.contains a -> l.contains a = false) → + (∀ (a : α), m.contains a → l.contains a = false) → (insertManyIfNewUnit m l).1.1.size = m.1.size + l.length := by simp_to_model using length_insertListIfNewUnit -theorem isEmpty_insertManyIfNewUnit_list_eq_false_of_isEmpty_eq_false [EquivBEq α] [LawfulHashable α] - (h : m.1.WF) {l : List α} : m.1.isEmpty = false → - (insertManyIfNewUnit m l).1.1.isEmpty = false := by - simp_to_model using isEmpty_insertListIfNewUnit_eq_false_of_isEmpty_eq_false - theorem isEmpty_insertManyIfNewUnit_list [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l : List α} : - (insertManyIfNewUnit m l).1.1.isEmpty ↔ m.1.isEmpty ∧ l.isEmpty := by + (insertManyIfNewUnit m l).1.1.isEmpty = (m.1.isEmpty && l.isEmpty) := by simp_to_model using isEmpty_insertListIfNewUnit theorem get?_insertManyIfNewUnit_list [EquivBEq α] [LawfulHashable α] (h : m.1.WF) From 0617166d892ecdf40008dffec9a371e2c79aa0ba Mon Sep 17 00:00:00 2001 From: monsterkrampe Date: Wed, 11 Dec 2024 11:40:36 +0100 Subject: [PATCH 47/83] Fix formatting --- .../Data/DHashMap/Internal/List/Associative.lean | 16 ++++++++-------- 1 file changed, 8 insertions(+), 8 deletions(-) diff --git a/src/Std/Data/DHashMap/Internal/List/Associative.lean b/src/Std/Data/DHashMap/Internal/List/Associative.lean index aba78f009d9e..ec21672c5387 100644 --- a/src/Std/Data/DHashMap/Internal/List/Associative.lean +++ b/src/Std/Data/DHashMap/Internal/List/Associative.lean @@ -2099,7 +2099,7 @@ theorem getKey?_insertList_of_mem [BEq α] [EquivBEq α] {k k' : α} (k_beq : k == k') (distinct_l : DistinctKeys l) (distinct_toInsert : toInsert.Pairwise (fun a b => (a.1 == b.1) = false)) - (mem : k ∈ (toInsert.map Sigma.fst)) : + (mem : k ∈ toInsert.map Sigma.fst) : getKey? k' (insertList l toInsert) = some k := by rw [List.mem_map] at mem rcases mem with ⟨pair, pair_mem, pair_eq⟩ @@ -2132,7 +2132,7 @@ theorem getKey_insertList_of_mem [BEq α] [EquivBEq α] {k k' : α} (k_beq : k == k') (distinct_l : DistinctKeys l) (distinct_toInsert : toInsert.Pairwise (fun a b => (a.1 == b.1) = false)) - (mem : k ∈ (toInsert.map Sigma.fst)) + (mem : k ∈ toInsert.map Sigma.fst) {h} : getKey k' (insertList l toInsert) h = k := by rw [← Option.some_inj] @@ -2157,7 +2157,7 @@ theorem getKey!_insertList_of_mem [BEq α] [EquivBEq α] [Inhabited α] {k k' : α} (k_beq : k == k') (distinct_l : DistinctKeys l) (distinct_toInsert : toInsert.Pairwise (fun a b => (a.1 == b.1) = false)) - (mem : k ∈ (toInsert.map Sigma.fst)) : + (mem : k ∈ toInsert.map Sigma.fst) : getKey! k' (insertList l toInsert) = k := by rw [getKey!_eq_getKey?] rw [getKey?_insertList_of_mem] @@ -2181,7 +2181,7 @@ theorem getKeyD_insertList_of_mem [BEq α] [EquivBEq α] {k k' fallback : α} (k_beq : k == k') (distinct_l : DistinctKeys l) (distinct_toInsert : toInsert.Pairwise (fun a b => (a.1 == b.1) = false)) - (mem : k ∈ (toInsert.map Sigma.fst)) : + (mem : k ∈ toInsert.map Sigma.fst) : getKeyD k' (insertList l toInsert) fallback = k := by rw [getKeyD_eq_getKey?] rw [getKey?_insertList_of_mem] @@ -2371,7 +2371,7 @@ theorem getKey?_insertListConst_of_mem [BEq α] [EquivBEq α] {k k' : α} (k_beq : k == k') (distinct_l : DistinctKeys l) (distinct_toInsert : toInsert.Pairwise (fun a b => (a.1 == b.1) = false)) - (mem : k ∈ (toInsert.map Prod.fst)) : + (mem : k ∈ toInsert.map Prod.fst) : getKey? k' (insertListConst l toInsert) = some k := by unfold insertListConst apply getKey?_insertList_of_mem @@ -2398,7 +2398,7 @@ theorem getKey_insertListConst_of_mem [BEq α] [EquivBEq α] {k k' : α} (k_beq : k == k') (distinct_l : DistinctKeys l) (distinct_toInsert : toInsert.Pairwise (fun a b => (a.1 == b.1) = false)) - (mem : k ∈ (toInsert.map Prod.fst)) + (mem : k ∈ toInsert.map Prod.fst) {h} : getKey k' (insertListConst l toInsert) h = k := by rw [← Option.some_inj] @@ -2423,7 +2423,7 @@ theorem getKey!_insertListConst_of_mem [BEq α] [EquivBEq α] [Inhabited α] {k k' : α} (k_beq : k == k') (distinct_l : DistinctKeys l) (distinct_toInsert : toInsert.Pairwise (fun a b => (a.1 == b.1) = false)) - (mem : k ∈ (toInsert.map Prod.fst)) : + (mem : k ∈ toInsert.map Prod.fst) : getKey! k' (insertListConst l toInsert) = k := by rw [getKey!_eq_getKey?] rw [getKey?_insertListConst_of_mem] @@ -2447,7 +2447,7 @@ theorem getKeyD_insertListConst_of_mem [BEq α] [EquivBEq α] {k k' fallback : α} (k_beq : k == k') (distinct_l : DistinctKeys l) (distinct_toInsert : toInsert.Pairwise (fun a b => (a.1 == b.1) = false)) - (mem : k ∈ (toInsert.map Prod.fst)) : + (mem : k ∈ toInsert.map Prod.fst) : getKeyD k' (insertListConst l toInsert) fallback = k := by rw [getKeyD_eq_getKey?] rw [getKey?_insertListConst_of_mem] From 0f7fd5148b7f5c0d8e090226319d01f0cd76a3aa Mon Sep 17 00:00:00 2001 From: monsterkrampe Date: Wed, 11 Dec 2024 12:36:14 +0100 Subject: [PATCH 48/83] Remove one theorem variant for Const --- .../DHashMap/Internal/List/Associative.lean | 194 +++++------------- src/Std/Data/DHashMap/Internal/RawLemmas.lean | 92 ++------- 2 files changed, 77 insertions(+), 209 deletions(-) diff --git a/src/Std/Data/DHashMap/Internal/List/Associative.lean b/src/Std/Data/DHashMap/Internal/List/Associative.lean index ec21672c5387..a428a23fd2da 100644 --- a/src/Std/Data/DHashMap/Internal/List/Associative.lean +++ b/src/Std/Data/DHashMap/Internal/List/Associative.lean @@ -2169,12 +2169,12 @@ theorem getKey!_insertList_of_mem [BEq α] [EquivBEq α] [Inhabited α] theorem getKeyD_insertList_of_contains_eq_false [BEq α] [EquivBEq α] {l toInsert : List ((a : α) × β a)} {k fallback : α} - (h : (toInsert.map Sigma.fst).contains k = false) : + (not_contains : (toInsert.map Sigma.fst).contains k = false) : getKeyD k (insertList l toInsert) fallback = getKeyD k l fallback := by rw [getKeyD_eq_getKey?] rw [getKey?_insertList_of_contains_eq_false] . rw [getKeyD_eq_getKey?] - . exact h + . exact not_contains theorem getKeyD_insertList_of_mem [BEq α] [EquivBEq α] {l toInsert : List ((a : α) × β a)} @@ -2230,103 +2230,6 @@ theorem isEmpty_insertList [BEq α] section variable {β : Type v} -theorem getValue?_insertList_of_contains_eq_false [BEq α] [PartialEquivBEq α] - {l toInsert : List ((_ : α) × β)} {k : α} - (not_contains : (toInsert.map Sigma.fst).contains k = false) : - getValue? k (insertList l toInsert) = getValue? k l := by - rw [← containsKey_eq_contains_map_fst] at not_contains - rw [getValue?_eq_getEntry?, getValue?_eq_getEntry?] - rw [getEntry?_insertList_of_contains_eq_false not_contains] - -theorem getValue?_insertList_of_mem [BEq α] [EquivBEq α] - {l toInsert : List ((_ : α) × β )} - {k k' : α} (k_beq : k == k') {v : β} - (distinct_l : DistinctKeys l) - (distinct_toInsert : toInsert.Pairwise (fun a b => (a.1 == b.1) = false)) - (mem : ⟨k, v⟩ ∈ toInsert) : - getValue? k' (insertList l toInsert) = v := by - rw [getValue?_eq_getEntry?] - have : getEntry? k' (insertList l toInsert) = getEntry? k' toInsert := by - apply getEntry?_insertList_of_contains_eq_true distinct_l distinct_toInsert - apply containsKey_of_beq _ k_beq - exact containsKey_of_mem mem - rw [← DistinctKeys.def] at distinct_toInsert - rw [getEntry?_of_mem distinct_toInsert k_beq mem] at this - rw [this] - simp - -theorem getValue_insertList_of_contains_eq_false [BEq α] [PartialEquivBEq α] - {l toInsert : List ((_ : α) × β)} {k : α} - (h' : (toInsert.map Sigma.fst).contains k = false) - {h} : - getValue k (insertList l toInsert) h = - getValue k l (containsKey_of_containsKey_insertList h h') := by - rw [← Option.some_inj] - rw [← getValue?_eq_some_getValue] - rw [← getValue?_eq_some_getValue] - rw [getValue?_insertList_of_contains_eq_false] - exact h' - -theorem getValue_insertList_of_mem [BEq α] [EquivBEq α] - {l toInsert : List ((_ : α) × β)} - {k k' : α} (k_beq : k == k') {v : β} - (distinct_l : DistinctKeys l) - (distinct_toInsert : toInsert.Pairwise (fun a b => (a.1 == b.1) = false)) - (mem : ⟨k, v⟩ ∈ toInsert) - {h} : - getValue k' (insertList l toInsert) h = v := by - rw [← Option.some_inj] - rw [← getValue?_eq_some_getValue] - rw [getValue?_insertList_of_mem] - . exact k_beq - . exact distinct_l - . exact distinct_toInsert - . exact mem - -theorem getValue!_insertList_of_contains_eq_false [BEq α] [PartialEquivBEq α] [Inhabited β] - {l toInsert : List ((_ : α) × β)} {k : α} - (h : (toInsert.map Sigma.fst).contains k = false) : - getValue! k (insertList l toInsert) = getValue! k l := by - simp only [getValue!_eq_getValue?] - rw [getValue?_insertList_of_contains_eq_false] - exact h - -theorem getValue!_insertList_of_mem [BEq α] [EquivBEq α] [Inhabited β] - {l toInsert : List ((_ : α) × β)} {k k' : α} {v: β} (k_beq : k == k') - (distinct_l : DistinctKeys l) - (distinct_toInsert : toInsert.Pairwise (fun a b => (a.1 == b.1) = false)) - (mem : ⟨k, v⟩ ∈ toInsert): - getValue! k' (insertList l toInsert) = v := by - simp only [getValue!_eq_getValue?] - rw [getValue?_insertList_of_mem] - · rw [Option.get!_some] - · exact k_beq - · exact distinct_l - . exact distinct_toInsert - . exact mem - -theorem getValueD_insertList_of_contains_eq_false [BEq α] [PartialEquivBEq α] - {l toInsert : List ((_ : α) × β)} {k : α} {fallback : β} - (h : (toInsert.map Sigma.fst).contains k = false) : - getValueD k (insertList l toInsert) fallback = getValueD k l fallback := by - simp only [getValueD_eq_getValue?] - rw [getValue?_insertList_of_contains_eq_false] - exact h - -theorem getValueD_insertList_of_mem [BEq α] [EquivBEq α] - {l toInsert : List ((_ : α) × β)} {k k' : α} {v fallback: β} (k_beq : k == k') - (distinct_l : DistinctKeys l) - (distinct_toInsert : toInsert.Pairwise (fun a b => (a.1 == b.1) = false)) - (mem : ⟨k, v⟩ ∈ toInsert): - getValueD k' (insertList l toInsert) fallback= v := by - simp only [getValueD_eq_getValue?] - rw [getValue?_insertList_of_mem] - · rw [Option.getD_some] - · exact k_beq - · exact distinct_l - . exact distinct_toInsert - . exact mem - -- TODO: this should be unified with Mathlib... /-- Internal implementation detail of the hash map -/ def Prod.toSigma_HashMapInternal (p : α × β) : ((_ : α) × β) := ⟨p.fst, p.snd⟩ @@ -2360,7 +2263,7 @@ theorem containsKey_of_containsKey_insertListConst [BEq α] [PartialEquivBEq α] theorem getKey?_insertListConst_of_contains_eq_false [BEq α] [EquivBEq α] {l : List ((_ : α) × β)} {toInsert : List (α × β)} {k : α} - (h : (toInsert.map Prod.fst).contains k = false) : + (not_contains : (toInsert.map Prod.fst).contains k = false) : getKey? k (insertListConst l toInsert) = getKey? k l := by unfold insertListConst apply getKey?_insertList_of_contains_eq_false @@ -2383,15 +2286,15 @@ theorem getKey?_insertListConst_of_mem [BEq α] [EquivBEq α] theorem getKey_insertListConst_of_contains_eq_false [BEq α] [EquivBEq α] {l : List ((_ : α) × β)} {toInsert : List (α × β)} {k : α} - (h' : (toInsert.map Prod.fst).contains k = false) + (not_contains : (toInsert.map Prod.fst).contains k = false) {h} : getKey k (insertListConst l toInsert) h = - getKey k l (containsKey_of_containsKey_insertListConst h h') := by + getKey k l (containsKey_of_containsKey_insertListConst h not_contains) := by rw [← Option.some_inj] rw [← getKey?_eq_some_getKey] rw [getKey?_insertListConst_of_contains_eq_false] . rw [getKey?_eq_some_getKey] - . exact h' + . exact not_contains theorem getKey_insertListConst_of_mem [BEq α] [EquivBEq α] {l : List ((_ : α) × β)} {toInsert : List (α × β)} @@ -2411,12 +2314,12 @@ theorem getKey_insertListConst_of_mem [BEq α] [EquivBEq α] theorem getKey!_insertListConst_of_contains_eq_false [BEq α] [EquivBEq α] [Inhabited α] {l : List ((_ : α) × β)} {toInsert : List (α × β)} {k : α} - (h : (toInsert.map Prod.fst).contains k = false) : + (not_contains : (toInsert.map Prod.fst).contains k = false) : getKey! k (insertListConst l toInsert) = getKey! k l := by rw [getKey!_eq_getKey?] rw [getKey?_insertListConst_of_contains_eq_false] . rw [getKey!_eq_getKey?] - . exact h + . exact not_contains theorem getKey!_insertListConst_of_mem [BEq α] [EquivBEq α] [Inhabited α] {l : List ((_ : α) × β)} {toInsert : List (α × β)} @@ -2435,12 +2338,12 @@ theorem getKey!_insertListConst_of_mem [BEq α] [EquivBEq α] [Inhabited α] theorem getKeyD_insertListConst_of_contains_eq_false [BEq α] [EquivBEq α] {l : List ((_ : α) × β)} {toInsert : List (α × β)} {k fallback : α} - (h : (toInsert.map Prod.fst).contains k = false) : + (not_contains : (toInsert.map Prod.fst).contains k = false) : getKeyD k (insertListConst l toInsert) fallback = getKeyD k l fallback := by rw [getKeyD_eq_getKey?] rw [getKey?_insertListConst_of_contains_eq_false] . rw [getKeyD_eq_getKey?] - . exact h + . exact not_contains theorem getKeyD_insertListConst_of_mem [BEq α] [EquivBEq α] {l : List ((_ : α) × β)} {toInsert : List (α × β)} @@ -2480,10 +2383,12 @@ theorem isEmpty_insertListConst [BEq α] theorem getValue?_insertListConst_of_contains_eq_false [BEq α] [PartialEquivBEq α] {l : List ((_ : α) × β)} {toInsert : List (α × β)} {k : α} - (h : (toInsert.map Prod.fst).contains k = false) : + (not_contains : (toInsert.map Prod.fst).contains k = false) : getValue? k (insertListConst l toInsert) = getValue? k l := by unfold insertListConst - rw [getValue?_insertList_of_contains_eq_false] + rw [getValue?_eq_getEntry?, getValue?_eq_getEntry?] + rw [getEntry?_insertList_of_contains_eq_false] + rw [containsKey_eq_contains_map_fst] simpa theorem getValue?_insertListConst_of_mem [BEq α] [EquivBEq α] @@ -2494,23 +2399,34 @@ theorem getValue?_insertListConst_of_mem [BEq α] [EquivBEq α] (mem : ⟨k, v⟩ ∈ toInsert) : getValue? k' (insertListConst l toInsert) = v := by unfold insertListConst - apply getValue?_insertList_of_mem (k_beq:=k_beq) - · exact distinct_l - · simpa [List.pairwise_map] + rw [getValue?_eq_getEntry?] + have : getEntry? k' (insertList l (toInsert.map Prod.toSigma_HashMapInternal)) = getEntry? k' (toInsert.map Prod.toSigma_HashMapInternal) := by + apply getEntry?_insertList_of_contains_eq_true distinct_l (by simpa [List.pairwise_map]) + apply containsKey_of_beq _ k_beq + rw [containsKey_eq_contains_map_fst, List.map_map, Prod.fst_comp_toSigma_HashMapInternal, List.contains_iff_exists_mem_beq] + exists k + rw [List.mem_map] + constructor + . exists ⟨k, v⟩ + . simp + rw [this] + rw [getEntry?_of_mem _ k_beq _] + . rfl + · simpa [DistinctKeys.def, List.pairwise_map] . simp only [List.mem_map] exists (k, v) theorem getValue_insertListConst_of_contains_eq_false [BEq α] [PartialEquivBEq α] {l : List ((_ : α) × β)} {toInsert : List (α × β)} {k : α} - {h : (toInsert.map Prod.fst).contains k = false} - {h'} : - getValue k (insertListConst l toInsert) h' = - getValue k l (containsKey_of_containsKey_insertListConst h' h) := by + {not_contains : (toInsert.map Prod.fst).contains k = false} + {h} : + getValue k (insertListConst l toInsert) h = + getValue k l (containsKey_of_containsKey_insertListConst h not_contains) := by rw [← Option.some_inj] rw [← getValue?_eq_some_getValue] rw [← getValue?_eq_some_getValue] rw [getValue?_insertListConst_of_contains_eq_false] - exact h + exact not_contains theorem getValue_insertListConst_of_mem [BEq α] [EquivBEq α] {l : List ((_ : α) × β)} {toInsert : List (α × β)} @@ -2529,11 +2445,11 @@ theorem getValue_insertListConst_of_mem [BEq α] [EquivBEq α] theorem getValue!_insertListConst_of_contains_eq_false [BEq α] [PartialEquivBEq α] [Inhabited β] {l : List ((_ : α) × β)} {toInsert : List (α × β)} {k : α} - (h : (toInsert.map Prod.fst).contains k = false) : + (not_contains : (toInsert.map Prod.fst).contains k = false) : getValue! k (insertListConst l toInsert) = getValue! k l := by simp only [getValue!_eq_getValue?] rw [getValue?_insertListConst_of_contains_eq_false] - exact h + exact not_contains theorem getValue!_insertListConst_of_mem [BEq α] [EquivBEq α] [Inhabited β] {l : List ((_ : α) × β)} {toInsert : List (α × β)} {k k' : α} {v: β} (k_beq : k == k') @@ -2551,11 +2467,11 @@ theorem getValue!_insertListConst_of_mem [BEq α] [EquivBEq α] [Inhabited β] theorem getValueD_insertListConst_of_contains_eq_false [BEq α] [PartialEquivBEq α] {l : List ((_ : α) × β)} {toInsert : List (α × β)} {k : α} {fallback : β} - (not_mem : (toInsert.map Prod.fst).contains k = false) : + (not_contains : (toInsert.map Prod.fst).contains k = false) : getValueD k (insertListConst l toInsert) fallback = getValueD k l fallback := by simp only [getValueD_eq_getValue?] rw [getValue?_insertListConst_of_contains_eq_false] - exact not_mem + exact not_contains theorem getValueD_insertListConst_of_mem [BEq α] [EquivBEq α] {l : List ((_ : α) × β)} {toInsert : List (α × β)} {k k' : α} {v fallback: β} (k_beq : k == k') @@ -2642,7 +2558,7 @@ theorem containsKey_of_containsKey_insertListIfNewUnit [BEq α] [PartialEquivBEq theorem getKey?_insertListIfNewUnit_of_contains_eq_false [BEq α] [EquivBEq α] {l : List ((_ : α) × Unit)} {toInsert : List α} {k : α} - (not_mem : toInsert.contains k = false) : + (not_contains : toInsert.contains k = false) : getKey? k (insertListIfNewUnit l toInsert) = getKey? k l := by induction toInsert generalizing l with | nil => simp [insertListIfNewUnit] @@ -2652,15 +2568,15 @@ theorem getKey?_insertListIfNewUnit_of_contains_eq_false [BEq α] [EquivBEq α] · rw [getKey?_insertEntryIfNew] split · rename_i hd_k - simp at not_mem + simp at not_contains exfalso rcases hd_k with ⟨h,_⟩ - rcases not_mem with ⟨h',_⟩ + rcases not_contains with ⟨h',_⟩ rw [BEq.symm h] at h' simp at h' · rfl - · simp only [List.contains_cons, Bool.or_eq_false_iff] at not_mem - apply And.right not_mem + · simp only [List.contains_cons, Bool.or_eq_false_iff] at not_contains + apply And.right not_contains theorem getKey?_insertListIfNewUnit_of_mem_of_contains_eq_false [BEq α] [EquivBEq α] {l : List ((_ : α) × Unit)} {toInsert : List α} @@ -2742,15 +2658,15 @@ theorem getKey?_insertListIfNewUnit_of_contains_of_contains [BEq α] [EquivBEq theorem getKey_insertListIfNewUnit_of_contains_eq_false [BEq α] [EquivBEq α] {l : List ((_ : α) × Unit)} {toInsert : List α} {k : α} - {not_mem : toInsert.contains k = false} + {not_contains : toInsert.contains k = false} {h} : getKey k (insertListIfNewUnit l toInsert) h = - getKey k l (containsKey_of_containsKey_insertListIfNewUnit not_mem h) := by + getKey k l (containsKey_of_containsKey_insertListIfNewUnit not_contains h) := by rw [← Option.some_inj] rw [← getKey?_eq_some_getKey] rw [getKey?_insertListIfNewUnit_of_contains_eq_false] . rw [getKey?_eq_some_getKey] - . exact not_mem + . exact not_contains theorem getKey_insertListIfNewUnit_of_mem_of_contains_eq_false [BEq α] [EquivBEq α] {l : List ((_ : α) × Unit)} {toInsert : List α} @@ -2784,18 +2700,18 @@ theorem getKey_insertListIfNewUnit_of_contains_of_contains [BEq α] [EquivBEq α theorem getKey!_insertListIfNewUnit_of_contains_eq_false [BEq α] [EquivBEq α] [Inhabited α] {l : List ((_ : α) × Unit)} {toInsert : List α} {k : α} - (h : toInsert.contains k = false) : + (not_contains : toInsert.contains k = false) : getKey! k (insertListIfNewUnit l toInsert) = getKey! k l := by rw [getKey!_eq_getKey?] rw [getKey?_insertListIfNewUnit_of_contains_eq_false] . rw [getKey!_eq_getKey?] - . exact h + . exact not_contains theorem getKey!_insertListIfNewUnit_of_mem_of_contains_eq_false [BEq α] [EquivBEq α] [Inhabited α] {l : List ((_ : α) × Unit)} {toInsert : List α} {k k' : α} (k_beq : k == k') (distinct : toInsert.Pairwise (fun a b => (a == b) = false)) - (mem : k ∈ toInsert) (h': containsKey k l = false) : + (mem : k ∈ toInsert) (h : containsKey k l = false) : getKey! k' (insertListIfNewUnit l toInsert) = k := by rw [getKey!_eq_getKey?] rw [getKey?_insertListIfNewUnit_of_mem_of_contains_eq_false] @@ -2803,35 +2719,35 @@ theorem getKey!_insertListIfNewUnit_of_mem_of_contains_eq_false [BEq α] [EquivB . exact k_beq . exact distinct . exact mem - . exact h' + . exact h theorem getKey!_insertListIfNewUnit_of_contains_of_contains [BEq α] [EquivBEq α] [Inhabited α] {l : List ((_ : α) × Unit)} {toInsert : List α} {k : α} (distinct : toInsert.Pairwise (fun a b => (a == b) = false)) - (mem : toInsert.contains k) (h' : containsKey k l = true) : + (mem : toInsert.contains k) (h : containsKey k l = true) : getKey! k (insertListIfNewUnit l toInsert) = getKey! k l := by rw [getKey!_eq_getKey?] rw [getKey?_insertListIfNewUnit_of_contains_of_contains] . rw [getKey!_eq_getKey?] · exact distinct · exact mem - · exact h' + · exact h theorem getKeyD_insertListIfNewUnit_of_contains_eq_false [BEq α] [EquivBEq α] {l : List ((_ : α) × Unit)} {toInsert : List α} {k fallback : α} - (h : toInsert.contains k = false) : + (not_contains : toInsert.contains k = false) : getKeyD k (insertListIfNewUnit l toInsert) fallback = getKeyD k l fallback := by rw [getKeyD_eq_getKey?] rw [getKey?_insertListIfNewUnit_of_contains_eq_false] . rw [getKeyD_eq_getKey?] - . exact h + . exact not_contains theorem getKeyD_insertListIfNewUnit_of_mem_of_contains_eq_false [BEq α] [EquivBEq α] {l : List ((_ : α) × Unit)} {toInsert : List α} {k k' fallback : α} (k_beq : k == k') (distinct : toInsert.Pairwise (fun a b => (a == b) = false)) - {mem : k ∈ toInsert } {h': containsKey k l = false} : + {mem : k ∈ toInsert } {h : containsKey k l = false} : getKeyD k' (insertListIfNewUnit l toInsert) fallback = k := by rw [getKeyD_eq_getKey?] rw [getKey?_insertListIfNewUnit_of_mem_of_contains_eq_false] @@ -2839,7 +2755,7 @@ theorem getKeyD_insertListIfNewUnit_of_mem_of_contains_eq_false [BEq α] [EquivB . exact k_beq . exact distinct . exact mem - . exact h' + . exact h theorem getKeyD_insertListIfNewUnit_of_contains_of_contains [BEq α] [EquivBEq α] {l : List ((_ : α) × Unit)} {toInsert : List α} diff --git a/src/Std/Data/DHashMap/Internal/RawLemmas.lean b/src/Std/Data/DHashMap/Internal/RawLemmas.lean index b896a45f63b0..a0c76f82347c 100644 --- a/src/Std/Data/DHashMap/Internal/RawLemmas.lean +++ b/src/Std/Data/DHashMap/Internal/RawLemmas.lean @@ -999,54 +999,6 @@ variable {β : Type v} (m : Raw₀ α (fun _ => β)) namespace Const -theorem get?_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] (h : m.1.WF) - {l : List ((_ : α) × β)} {k : α} (h : (l.map Sigma.fst).contains k = false) : - Const.get? (m.insertMany l).1 k = Const.get? m k := by - simp_to_model using getValue?_insertList_of_contains_eq_false - -theorem get?_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] (h : m.1.WF) - {l : List ((_ : α) × β)} {k k' : α} (k_beq : k == k') {v : β} (mem : ⟨k, v⟩ ∈ l) - (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) : - Const.get? (m.insertMany l).1 k' = v := by - simp_to_model using getValue?_insertList_of_mem - -theorem get_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] (h : m.1.WF) - {l : List ((_ : α) × β)} {k : α} - (h₁ : (l.map Sigma.fst).contains k = false) {h'}: - Const.get (m.insertMany l).1 k h' = - Const.get m k (contains_of_contains_insertMany_list _ h h' h₁) := by - simp_to_model using getValue_insertList_of_contains_eq_false - -theorem get_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] (h : m.1.WF) - {l : List ((_ : α) × β)} {k k' : α} (k_beq : k == k') {v : β} (mem : ⟨k, v⟩ ∈ l) - (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) {h'}: - Const.get (m.insertMany l).1 k' h' = v := by - simp_to_model using getValue_insertList_of_mem - -theorem get!_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] - [Inhabited β] (h : m.1.WF) {l : List ((_ : α) × β)} {k : α} - (h' : (l.map Sigma.fst).contains k = false) : - Const.get! (m.insertMany l).1 k = Const.get! m k := by - simp_to_model using getValue!_insertList_of_contains_eq_false - -theorem get!_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] [Inhabited β] (h : m.1.WF) - {l : List ((_ : α) × β)} {k k' : α} (k_beq : k == k') {v : β} (mem: ⟨k, v⟩ ∈ l) - (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) : - Const.get! (m.insertMany l).1 k' = v := by - simp_to_model using getValue!_insertList_of_mem - -theorem getD_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] (h : m.1.WF) - {l : List ((_ : α) × β)} {k : α} {fallback : β} - (h' : (l.map Sigma.fst).contains k = false) : - Const.getD (m.insertMany l).1 k fallback = Const.getD m k fallback := by - simp_to_model using getValueD_insertList_of_contains_eq_false - -theorem getD_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] (h : m.1.WF) - {l : List ((_ : α) × β)} {k k' : α} (k_beq : k == k') {v fallback : β} (mem : ⟨k, v⟩ ∈ l) - (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) : - Const.getD (m.insertMany l).1 k' fallback = v := by - simp_to_model using getValueD_insertList_of_mem - @[simp] theorem insertMany_nil : insertMany m [] = m := by @@ -1064,23 +1016,23 @@ theorem insertMany_cons {l : List (α × β)} {k : α} {v : β} : | nil => simp [insertListₘ] | cons hd tl => simp [insertListₘ] -theorem contains_constInsertMany_list [EquivBEq α] [LawfulHashable α] (h : m.1.WF) +theorem contains_insertMany_list [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l : List (α × β)} {k : α} : (Const.insertMany m l).1.contains k = (m.contains k || (l.map Prod.fst).contains k) := by simp_to_model using containsKey_insertListConst -theorem contains_of_contains_constInsertMany_list [EquivBEq α] [LawfulHashable α] (h : m.1.WF) +theorem contains_of_contains_insertMany_list [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l : List ( α × β )} {k : α} : (insertMany m l).1.contains k → (l.map Prod.fst).contains k = false → m.contains k := by simp_to_model using containsKey_of_containsKey_insertListConst -theorem getKey?_constInsertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] (h : m.1.WF) +theorem getKey?_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l : List (α × β)} {k : α} (h' : (l.map Prod.fst).contains k = false) : (insertMany m l).1.getKey? k = m.getKey? k := by simp_to_model using getKey?_insertListConst_of_contains_eq_false -theorem getKey?_constInsertMany_list_of_mem [EquivBEq α] [LawfulHashable α] (h : m.1.WF) +theorem getKey?_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l : List (α × β)} {k k' : α} (k_beq : k == k') (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) @@ -1088,15 +1040,15 @@ theorem getKey?_constInsertMany_list_of_mem [EquivBEq α] [LawfulHashable α] (h (insertMany m l).1.getKey? k' = some k := by simp_to_model using getKey?_insertListConst_of_mem -theorem getKey_constInsertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] (h : m.1.WF) +theorem getKey_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l : List (α × β)} {k : α} {h₁ : (l.map Prod.fst).contains k = false} {h'} : (insertMany m l).1.getKey k h' = - m.getKey k (contains_of_contains_constInsertMany_list _ h h' h₁) := by + m.getKey k (contains_of_contains_insertMany_list _ h h' h₁) := by simp_to_model using getKey_insertListConst_of_contains_eq_false -theorem getKey_constInsertMany_list_of_mem [EquivBEq α] [LawfulHashable α] (h : m.1.WF) +theorem getKey_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l : List (α × β)} {k k' : α} (k_beq : k == k') (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) @@ -1105,13 +1057,13 @@ theorem getKey_constInsertMany_list_of_mem [EquivBEq α] [LawfulHashable α] (h (insertMany m l).1.getKey k' h' = k := by simp_to_model using getKey_insertListConst_of_mem -theorem getKey!_constInsertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] [Inhabited α] +theorem getKey!_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] [Inhabited α] (h : m.1.WF) {l : List (α × β)} {k : α} (h' : (l.map Prod.fst).contains k = false) : (insertMany m l).1.getKey! k = m.getKey! k := by simp_to_model using getKey!_insertListConst_of_contains_eq_false -theorem getKey!_constInsertMany_list_of_mem [EquivBEq α] [LawfulHashable α] [Inhabited α] (h : m.1.WF) +theorem getKey!_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] [Inhabited α] (h : m.1.WF) {l : List (α × β)} {k k' : α} (k_beq : k == k') (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) @@ -1119,13 +1071,13 @@ theorem getKey!_constInsertMany_list_of_mem [EquivBEq α] [LawfulHashable α] [I (insertMany m l).1.getKey! k' = k := by simp_to_model using getKey!_insertListConst_of_mem -theorem getKeyD_constInsertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] (h : m.1.WF) +theorem getKeyD_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l : List (α × β)} {k fallback : α} (h' : (l.map Prod.fst).contains k = false) : (insertMany m l).1.getKeyD k fallback = m.getKeyD k fallback := by simp_to_model using getKeyD_insertListConst_of_contains_eq_false -theorem getKeyD_constInsertMany_list_of_mem [EquivBEq α] [LawfulHashable α] (h : m.1.WF) +theorem getKeyD_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l : List (α × β)} {k k' fallback : α} (k_beq : k == k') (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) @@ -1133,62 +1085,62 @@ theorem getKeyD_constInsertMany_list_of_mem [EquivBEq α] [LawfulHashable α] (h (insertMany m l).1.getKeyD k' fallback = k := by simp_to_model using getKeyD_insertListConst_of_mem -theorem size_constInsertMany_list [EquivBEq α] [LawfulHashable α] (h : m.1.WF) +theorem size_insertMany_list [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l : List (α × β)} {distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)} : (∀ (a : α), m.contains a → (l.map Prod.fst).contains a = false) → (insertMany m l).1.1.size = m.1.size + l.length := by simp_to_model using length_insertListConst -theorem isEmpty_constInsertMany_list [EquivBEq α] [LawfulHashable α] (h : m.1.WF) +theorem isEmpty_insertMany_list [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l : List (α × β)} : (insertMany m l).1.1.isEmpty = (m.1.isEmpty && l.isEmpty) := by simp_to_model using isEmpty_insertListConst -theorem get?_constInsertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] (h : m.1.WF) +theorem get?_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l : List (α × β)} {k : α} (h' : (l.map Prod.fst).contains k = false) : get? (insertMany m l).1 k = get? m k := by simp_to_model using getValue?_insertListConst_of_contains_eq_false -theorem get?_constInsertMany_list_of_mem [EquivBEq α] [LawfulHashable α] (h : m.1.WF) +theorem get?_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l : List (α × β)} {k k' : α} (k_beq : k == k') {v : β} (mem: ⟨k, v⟩ ∈ l) (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) : get? (insertMany m l).1 k' = v := by simp_to_model using getValue?_insertListConst_of_mem -theorem get_constInsertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] (h : m.1.WF) +theorem get_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l : List (α × β)} {k : α} (h₁ : (l.map Prod.fst).contains k = false) {h'}: - get (insertMany m l).1 k h' = get m k (contains_of_contains_constInsertMany_list _ h h' h₁) := by + get (insertMany m l).1 k h' = get m k (contains_of_contains_insertMany_list _ h h' h₁) := by simp_to_model using getValue_insertListConst_of_contains_eq_false -theorem get_constInsertMany_list_of_mem [EquivBEq α] [LawfulHashable α] (h : m.1.WF) +theorem get_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l : List (α × β)} {k k' : α} (k_beq : k == k') {v : β} (mem: ⟨k, v⟩ ∈ l) (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) {h'}: get (insertMany m l).1 k' h' = v := by simp_to_model using getValue_insertListConst_of_mem -theorem get!_constInsertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] +theorem get!_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] [Inhabited β] (h : m.1.WF) {l : List (α × β)} {k : α} (h' : (l.map Prod.fst).contains k = false) : get! (insertMany m l).1 k = get! m k := by simp_to_model using getValue!_insertListConst_of_contains_eq_false -theorem get!_constInsertMany_list_of_mem [EquivBEq α] [LawfulHashable α] [Inhabited β] (h : m.1.WF) +theorem get!_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] [Inhabited β] (h : m.1.WF) {l : List (α × β)} {k k' : α} (k_beq : k == k') {v : β} (mem: ⟨k, v⟩ ∈ l) (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) : get! (insertMany m l).1 k' = v := by simp_to_model using getValue!_insertListConst_of_mem -theorem getD_constInsertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] (h : m.1.WF) +theorem getD_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l : List (α × β)} {k : α} {fallback : β} (h' : (l.map Prod.fst).contains k = false) : getD (insertMany m l).1 k fallback = getD m k fallback := by simp_to_model using getValueD_insertListConst_of_contains_eq_false -theorem getD_constInsertMany_list_of_mem [EquivBEq α] [LawfulHashable α] (h : m.1.WF) +theorem getD_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l : List (α × β)} {k k' : α} (k_beq : k == k') {v fallback : β} (mem : ⟨k, v⟩ ∈ l) (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) : getD (insertMany m l).1 k' fallback = v := by From fe60013e1916a32db78c1bb47511a08ab0d52799 Mon Sep 17 00:00:00 2001 From: jt0202 Date: Thu, 12 Dec 2024 18:48:50 +0100 Subject: [PATCH 49/83] Style fix for Internal list --- .../DHashMap/Internal/List/Associative.lean | 114 +++++++++++------- 1 file changed, 68 insertions(+), 46 deletions(-) diff --git a/src/Std/Data/DHashMap/Internal/List/Associative.lean b/src/Std/Data/DHashMap/Internal/List/Associative.lean index a428a23fd2da..18ed12a7f4e9 100644 --- a/src/Std/Data/DHashMap/Internal/List/Associative.lean +++ b/src/Std/Data/DHashMap/Internal/List/Associative.lean @@ -240,7 +240,6 @@ theorem containsKey_eq_false_iff [BEq α][PartialEquivBEq α] {l : List ((a : α · intro h apply PartialEquivBEq.symm h - theorem containsKey_cons_eq_false [BEq α] {l : List ((a : α) × β a)} {k a : α} {v : β k} : (containsKey a (⟨k, v⟩ :: l) = false) ↔ ((k == a) = false) ∧ (containsKey a l = false) := by simp [containsKey_cons, not_or] @@ -274,6 +273,15 @@ theorem containsKey_eq_isSome_getEntry? [BEq α] {l : List ((a : α) × β a)} { · simp [getEntry?_cons_of_false h, h, ih] · simp [getEntry?_cons_of_true h, h] +theorem containsKey_eq_contains_map_fst [BEq α] [PartialEquivBEq α] {l : List ((a : α) × β a)} + {k : α} : containsKey k l = (l.map Sigma.fst).contains k := by + induction l with + | nil => simp + | cons hd tl ih => + rw [containsKey_cons, ih] + simp only [List.map_cons, List.contains_cons] + rw [BEq.comm] + theorem isEmpty_eq_false_of_containsKey [BEq α] {l : List ((a : α) × β a)} {a : α} (h : containsKey a l = true) : l.isEmpty = false := by cases l <;> simp_all @@ -984,6 +992,19 @@ theorem DistinctKeys.replaceEntry [BEq α] [PartialEquivBEq α] {l : List ((a : refine ⟨h.1, ?_⟩ simpa [containsKey_congr (BEq.symm hk'k)] using h.2 +theorem getEntry?_of_mem [BEq α] [PartialEquivBEq α] + {l : List ((a : α) × β a)} (hl : DistinctKeys l) + {k k' : α} (hk : k == k') {v : β k} (hkv : ⟨k, v⟩ ∈ l) : + getEntry? k' l = some ⟨k, v⟩ := by + induction l using assoc_induction with + | nil => simp at hkv + | cons k₁ v₁ t ih => + obtain ⟨⟨⟩⟩|hkv := List.mem_cons.1 hkv + · rw [getEntry?_cons_of_true hk] + · rw [getEntry?_cons_of_false, ih hl.tail hkv] + exact BEq.neq_of_neq_of_beq (containsKey_eq_false_iff.1 hl.containsKey_eq_false _ hkv) hk + + /-- Internal implementation detail of the hash map -/ def insertEntry [BEq α] (k : α) (v : β k) (l : List ((a : α) × β a)) : List ((a : α) × β a) := bif containsKey k l then replaceEntry k v l else ⟨k, v⟩ :: l @@ -1946,27 +1967,6 @@ theorem getEntry?_insertList_of_contains_eq_true [BEq α] [EquivBEq α] rw [containsKey_eq_isSome_getEntry?] at contains exact Option.or_of_isSome contains -theorem getEntry?_of_mem [BEq α] [PartialEquivBEq α] - {l : List ((a : α) × β a)} (hl : DistinctKeys l) - {k k' : α} (hk : k == k') {v : β k} (hkv : ⟨k, v⟩ ∈ l) : - getEntry? k' l = some ⟨k, v⟩ := by - induction l using assoc_induction with - | nil => simp at hkv - | cons k₁ v₁ t ih => - obtain ⟨⟨⟩⟩|hkv := List.mem_cons.1 hkv - · rw [getEntry?_cons_of_true hk] - · rw [getEntry?_cons_of_false, ih hl.tail hkv] - exact BEq.neq_of_neq_of_beq (containsKey_eq_false_iff.1 hl.containsKey_eq_false _ hkv) hk - -theorem containsKey_eq_contains_map_fst [BEq α] [PartialEquivBEq α] {l : List ((a : α) × β a)} - {k : α} : containsKey k l = (l.map Sigma.fst).contains k := by - induction l with - | nil => simp - | cons hd tl ih => - rw [containsKey_cons, ih] - simp only [List.map_cons, List.contains_cons] - rw [BEq.comm] - theorem containsKey_insertList [BEq α] [PartialEquivBEq α] {l toInsert : List ((a : α) × β a)} {k : α} : containsKey k (List.insertList l toInsert) = (containsKey k l || (toInsert.map Sigma.fst).contains k) := by @@ -2101,19 +2101,10 @@ theorem getKey?_insertList_of_mem [BEq α] [EquivBEq α] (distinct_toInsert : toInsert.Pairwise (fun a b => (a.1 == b.1) = false)) (mem : k ∈ toInsert.map Sigma.fst) : getKey? k' (insertList l toInsert) = some k := by - rw [List.mem_map] at mem - rcases mem with ⟨pair, pair_mem, pair_eq⟩ - rw [getKey?_eq_getEntry?] - have : getEntry? k' (insertList l toInsert) = getEntry? k' toInsert := by - apply getEntry?_insertList_of_contains_eq_true distinct_l distinct_toInsert - apply containsKey_of_beq _ k_beq - rw [← pair_eq] - exact containsKey_of_mem pair_mem - rw [← DistinctKeys.def] at distinct_toInsert - rw [← pair_eq] at k_beq - rw [getEntry?_of_mem distinct_toInsert k_beq pair_mem] at this - . rw [this] - simpa + rcases List.mem_map.1 mem with ⟨⟨k, v⟩, pair_mem, rfl⟩ + rw [getKey?_eq_getEntry?, getEntry?_insertList distinct_l distinct_toInsert, + getEntry?_of_mem (DistinctKeys.def.2 distinct_toInsert) k_beq pair_mem, Option.some_or, + Option.map_some'] theorem getKey_insertList_of_contains_eq_false [BEq α] [EquivBEq α] {l toInsert : List ((a : α) × β a)} {k : α} @@ -2218,6 +2209,18 @@ theorem length_insertList [BEq α] [EquivBEq α] (insertList l toInsert).length = l.length + toInsert.length := by simpa using (perm_insertList distinct_l distinct_toInsert distinct_both).length_eq +theorem length_insertList_le [BEq α] + {l toInsert : List ((a : α) × β a)} : + (insertList l toInsert).length ≤ l.length + toInsert.length := by + induction toInsert generalizing l with + | nil => simp only [insertList, List.length_nil, Nat.add_zero, Nat.le_refl] + | cons hd tl ih => + simp only [insertList, List.length_cons] + apply Nat.le_trans ih + rw [Nat.add_comm tl.length 1, ← Nat.add_assoc] + apply Nat.add_le_add _ (Nat.le_refl _) + apply length_insertEntry_le + theorem isEmpty_insertList [BEq α] {l toInsert : List ((a : α) × β a)} : (List.insertList l toInsert).isEmpty = (l.isEmpty && toInsert.isEmpty) := by @@ -2230,20 +2233,18 @@ theorem isEmpty_insertList [BEq α] section variable {β : Type v} --- TODO: this should be unified with Mathlib... /-- Internal implementation detail of the hash map -/ -def Prod.toSigma_HashMapInternal (p : α × β) : ((_ : α) × β) := ⟨p.fst, p.snd⟩ +def Prod.toSigma (p : α × β) : ((_ : α) × β) := ⟨p.fst, p.snd⟩ @[simp] -theorem Prod.fst_comp_toSigma_HashMapInternal : - Sigma.fst ∘ Prod.toSigma_HashMapInternal = @Prod.fst α β := by +theorem Prod.fst_comp_toSigma : + Sigma.fst ∘ Prod.toSigma = @Prod.fst α β := by apply funext - simp [Prod.toSigma_HashMapInternal] + simp [Prod.toSigma] --- TODO: maybe insertList_prod would be a better name? /-- Internal implementation detail of the hash map -/ -def insertListConst [BEq α] (l: List ((_ : α) × β)) (toInsert: List (α × β)) : List ((_ : α) × β) := - insertList l (toInsert.map Prod.toSigma_HashMapInternal) +def insertListConst [BEq α] (l : List ((_ : α) × β)) (toInsert : List (α × β)) : List ((_ : α) × β) := + insertList l (toInsert.map Prod.toSigma) theorem containsKey_insertListConst [BEq α] [PartialEquivBEq α] {l : List ((_ : α) × β)} {toInsert : List (α × β)} {k : α} : @@ -2281,7 +2282,7 @@ theorem getKey?_insertListConst_of_mem [BEq α] [EquivBEq α] · exact k_beq · exact distinct_l · simpa [List.pairwise_map] - · simp only [List.map_map, Prod.fst_comp_toSigma_HashMapInternal] + · simp only [List.map_map, Prod.fst_comp_toSigma] exact mem theorem getKey_insertListConst_of_contains_eq_false [BEq α] [EquivBEq α] @@ -2373,6 +2374,13 @@ theorem length_insertListConst [BEq α] [EquivBEq α] · simpa [List.pairwise_map] · simpa using distinct_both +theorem length_insertListConst_le [BEq α] + {l : List ((_ : α) × β)} {toInsert : List (α × β)} : + (insertListConst l toInsert).length ≤ l.length + toInsert.length := by + unfold insertListConst + rw [← List.length_map toInsert Prod.toSigma] + apply length_insertList_le + theorem isEmpty_insertListConst [BEq α] {l : List ((_ : α) × β)} {toInsert : List (α × β)} : (List.insertListConst l toInsert).isEmpty = (l.isEmpty && toInsert.isEmpty) := by @@ -2400,10 +2408,12 @@ theorem getValue?_insertListConst_of_mem [BEq α] [EquivBEq α] getValue? k' (insertListConst l toInsert) = v := by unfold insertListConst rw [getValue?_eq_getEntry?] - have : getEntry? k' (insertList l (toInsert.map Prod.toSigma_HashMapInternal)) = getEntry? k' (toInsert.map Prod.toSigma_HashMapInternal) := by + have : getEntry? k' (insertList l (toInsert.map Prod.toSigma)) = + getEntry? k' (toInsert.map Prod.toSigma) := by apply getEntry?_insertList_of_contains_eq_true distinct_l (by simpa [List.pairwise_map]) apply containsKey_of_beq _ k_beq - rw [containsKey_eq_contains_map_fst, List.map_map, Prod.fst_comp_toSigma_HashMapInternal, List.contains_iff_exists_mem_beq] + rw [containsKey_eq_contains_map_fst, List.map_map, Prod.fst_comp_toSigma, + List.contains_iff_exists_mem_beq] exists k rw [List.mem_map] constructor @@ -2815,6 +2825,18 @@ theorem length_insertListIfNewUnit [BEq α] [EquivBEq α] rw [Bool.or_eq_false_iff] at distinct_both apply And.right distinct_both +theorem length_insertListIfNewUnit_le [BEq α] [EquivBEq α] + {l : List ((_ : α) × Unit)} {toInsert : List α}: + (insertListIfNewUnit l toInsert).length ≤ l.length + toInsert.length := by + induction toInsert generalizing l with + | nil => simp only [insertListIfNewUnit, List.length_nil, Nat.add_zero, Nat.le_refl] + | cons hd tl ih => + simp only [insertListIfNewUnit, List.length_cons] + apply Nat.le_trans ih + rw [Nat.add_comm tl.length 1, ← Nat.add_assoc] + apply Nat.add_le_add _ (Nat.le_refl _) + apply length_insertEntryIfNew_le + theorem isEmpty_insertListIfNewUnit [BEq α] {l : List ((_ : α) × Unit)} {toInsert : List α} : (List.insertListIfNewUnit l toInsert).isEmpty = (l.isEmpty && toInsert.isEmpty) := by From 4b79e10f8567380ec454f7bb58890a175d8aede5 Mon Sep 17 00:00:00 2001 From: jt0202 Date: Thu, 12 Dec 2024 19:19:21 +0100 Subject: [PATCH 50/83] Added size_insertMany_list_le & further style fixes --- src/Std/Data/DHashMap/Internal/RawLemmas.lean | 49 ++++++++++++------- 1 file changed, 32 insertions(+), 17 deletions(-) diff --git a/src/Std/Data/DHashMap/Internal/RawLemmas.lean b/src/Std/Data/DHashMap/Internal/RawLemmas.lean index a0c76f82347c..b5a9b905af91 100644 --- a/src/Std/Data/DHashMap/Internal/RawLemmas.lean +++ b/src/Std/Data/DHashMap/Internal/RawLemmas.lean @@ -871,7 +871,7 @@ theorem contains_of_contains_insertMany_list [EquivBEq α] [LawfulHashable α] ( theorem get?_insertMany_list_of_contains_eq_false [LawfulBEq α] (h : m.1.WF) {l : List ((a : α) × β a)} {k : α} - (h' : (l.map (Sigma.fst)).contains k = false) : + (h' : (l.map Sigma.fst).contains k = false) : (m.insertMany l).1.get? k = m.get? k := by simp_to_model using getValueCast?_insertList_of_contains_eq_false @@ -884,7 +884,7 @@ theorem get?_insertMany_list_of_mem [LawfulBEq α] (h : m.1.WF) theorem get_insertMany_list_of_contains_eq_false [LawfulBEq α] (h : m.1.WF) {l : List ((a : α) × β a)} {k : α} - (contains : (l.map (Sigma.fst)).contains k = false) + (contains : (l.map Sigma.fst).contains k = false) {h'} : (m.insertMany l).1.get k h' = m.get k (contains_of_contains_insertMany_list _ h h' contains) := by @@ -900,7 +900,7 @@ theorem get_insertMany_list_of_mem [LawfulBEq α] (h : m.1.WF) theorem get!_insertMany_list_of_contains_eq_false [LawfulBEq α] (h : m.1.WF) {l : List ((a : α) × β a)} {k : α} [Inhabited (β k)] - (h' : (l.map (Sigma.fst)).contains k = false) : + (h' : (l.map Sigma.fst).contains k = false) : (m.insertMany l).1.get! k = m.get! k := by simp_to_model using getValueCast!_insertList_of_contains_eq_false @@ -913,7 +913,7 @@ theorem get!_insertMany_list_of_mem [LawfulBEq α] (h : m.1.WF) theorem getD_insertMany_list_of_contains_eq_false [LawfulBEq α] (h : m.1.WF) {l : List ((a : α) × β a)} {k : α} {fallback : β k} - (not_mem : (l.map (Sigma.fst)).contains k = false) : + (not_mem : (l.map Sigma.fst).contains k = false) : (m.insertMany l).1.getD k fallback = m.getD k fallback := by simp_to_model using getValueCastD_insertList_of_contains_eq_false @@ -989,6 +989,11 @@ theorem size_insertMany_list [EquivBEq α] [LawfulHashable α] (h : m.1.WF) (m.insertMany l).1.1.size = m.1.size + l.length := by simp_to_model using length_insertList +theorem size_insertMany_list_le [EquivBEq α] [LawfulHashable α] (h : m.1.WF) + {l : List ((a : α) × β a)} : + (m.insertMany l).1.1.size ≤ m.1.size + l.length := by + simp_to_model using length_insertList_le + theorem isEmpty_insertMany_list [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l : List ((a : α) × β a)} : (m.insertMany l).1.1.isEmpty = (m.1.isEmpty && l.isEmpty) := by @@ -1092,6 +1097,11 @@ theorem size_insertMany_list [EquivBEq α] [LawfulHashable α] (h : m.1.WF) (insertMany m l).1.1.size = m.1.size + l.length := by simp_to_model using length_insertListConst +theorem size_insertMany_list_le [EquivBEq α] [LawfulHashable α] (h : m.1.WF) + {l : List (α × β)} : + (insertMany m l).1.1.size ≤ m.1.size + l.length := by + simp_to_model using length_insertListConst_le + theorem isEmpty_insertMany_list [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l : List (α × β)} : (insertMany m l).1.1.isEmpty = (m.1.isEmpty && l.isEmpty) := by @@ -1104,7 +1114,7 @@ theorem get?_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable simp_to_model using getValue?_insertListConst_of_contains_eq_false theorem get?_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] (h : m.1.WF) - {l : List (α × β)} {k k' : α} (k_beq : k == k') {v : β} (mem: ⟨k, v⟩ ∈ l) + {l : List (α × β)} {k k' : α} (k_beq : k == k') {v : β} (mem : ⟨k, v⟩ ∈ l) (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) : get? (insertMany m l).1 k' = v := by simp_to_model using getValue?_insertListConst_of_mem @@ -1112,13 +1122,13 @@ theorem get?_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] (h : m.1.W theorem get_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l : List (α × β)} {k : α} (h₁ : (l.map Prod.fst).contains k = false) - {h'}: + {h'} : get (insertMany m l).1 k h' = get m k (contains_of_contains_insertMany_list _ h h' h₁) := by simp_to_model using getValue_insertListConst_of_contains_eq_false theorem get_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] (h : m.1.WF) - {l : List (α × β)} {k k' : α} (k_beq : k == k') {v : β} (mem: ⟨k, v⟩ ∈ l) - (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) {h'}: + {l : List (α × β)} {k k' : α} (k_beq : k == k') {v : β} (mem : ⟨k, v⟩ ∈ l) + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) {h'} : get (insertMany m l).1 k' h' = v := by simp_to_model using getValue_insertListConst_of_mem @@ -1129,7 +1139,7 @@ theorem get!_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable simp_to_model using getValue!_insertListConst_of_contains_eq_false theorem get!_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] [Inhabited β] (h : m.1.WF) - {l : List (α × β)} {k k' : α} (k_beq : k == k') {v : β} (mem: ⟨k, v⟩ ∈ l) + {l : List (α × β)} {k k' : α} (k_beq : k == k') {v : β} (mem : ⟨k, v⟩ ∈ l) (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) : get! (insertMany m l).1 k' = v := by simp_to_model using getValue!_insertListConst_of_mem @@ -1146,7 +1156,7 @@ theorem getD_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] (h : m.1.W getD (insertMany m l).1 k' fallback = v := by simp_to_model using getValueD_insertListConst_of_mem -variable (m: Raw₀ α (fun _ => Unit)) +variable (m : Raw₀ α (fun _ => Unit)) @[simp] theorem insertManyIfNewUnit_nil : @@ -1166,7 +1176,7 @@ theorem insertManyIfNewUnit_cons {l : List α} {k : α} : | cons hd tl => simp [insertListIfNewUnitₘ] theorem contains_insertManyIfNewUnit_list [EquivBEq α] [LawfulHashable α] (h : m.1.WF) - {l : List α} {k : α}: + {l : List α} {k : α} : (insertManyIfNewUnit m l).1.contains k = (m.contains k || l.contains k) := by simp_to_model using containsKey_insertListIfNewUnit @@ -1189,7 +1199,7 @@ theorem getKey?_insertManyIfNewUnit_list_of_mem_of_contains_eq_false [EquivBEq theorem getKey?_insertManyIfNewUnit_list_of_contains_of_contains [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l : List α} {k : α} (distinct : l.Pairwise (fun a b => (a == b) = false)) - (h': l.contains k) : + (h' : l.contains k) : m.contains k → getKey? (insertManyIfNewUnit m l).1 k = getKey? m k := by simp_to_model using getKey?_insertListIfNewUnit_of_contains_of_contains @@ -1211,7 +1221,7 @@ theorem getKey_insertManyIfNewUnit_list_of_mem_of_contains_eq_false [EquivBEq α theorem getKey_insertManyIfNewUnit_list_mem_of_contains_of_contains [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l : List α} {k : α} (distinct : l.Pairwise (fun a b => (a == b) = false)) - (h': l.contains k) {mem' : m.contains k} {h'} : + (h' : l.contains k) {mem' : m.contains k} {h'} : getKey (insertManyIfNewUnit m l).1 k h' = getKey m k mem' := by simp_to_model using getKey_insertListIfNewUnit_of_contains_of_contains @@ -1231,13 +1241,13 @@ theorem getKey!_insertManyIfNewUnit_list_of_mem_of_contains_eq_false [EquivBEq theorem getKey!_insertManyIfNewUnit_list_mem_of_contains_of_contains [EquivBEq α] [LawfulHashable α] [Inhabited α] (h : m.1.WF) {l : List α} {k : α} (distinct : l.Pairwise (fun a b => (a == b) = false)) - (h': l.contains k) : m.contains k → + (h' : l.contains k) : m.contains k → getKey! (insertManyIfNewUnit m l).1 k = getKey! m k := by simp_to_model using getKey!_insertListIfNewUnit_of_contains_of_contains theorem getKeyD_insertManyIfNewUnit_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l : List α} {k fallback : α} - (h' : l.contains k = false): + (h' : l.contains k = false) : getKeyD (insertManyIfNewUnit m l).1 k fallback = getKeyD m k fallback := by simp_to_model using getKeyD_insertListIfNewUnit_of_contains_eq_false @@ -1251,17 +1261,22 @@ theorem getKeyD_insertManyIfNewUnit_list_of_mem_of_contains_eq_false [EquivBEq theorem getKeyD_insertManyIfNewUnit_list_of_contains_of_contains [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l : List α} {k fallback : α} (distinct : l.Pairwise (fun a b => (a == b) = false)) - (h': l.contains k) : m.contains k → + (h' : l.contains k) : m.contains k → getKeyD (insertManyIfNewUnit m l).1 k fallback = getKeyD m k fallback := by simp_to_model using getKeyD_insertListIfNewUnit_of_contains_of_contains theorem size_insertManyIfNewUnit_list [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l : List α} - (distinct : l.Pairwise (fun a b => (a == b) = false)): + (distinct : l.Pairwise (fun a b => (a == b) = false)) : (∀ (a : α), m.contains a → l.contains a = false) → (insertManyIfNewUnit m l).1.1.size = m.1.size + l.length := by simp_to_model using length_insertListIfNewUnit +theorem size_insertManyIfNewUnit_list_le [EquivBEq α] [LawfulHashable α] (h : m.1.WF) + {l : List α} : + (insertManyIfNewUnit m l).1.1.size ≤ m.1.size + l.length := by + simp_to_model using length_insertListIfNewUnit_le + theorem isEmpty_insertManyIfNewUnit_list [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l : List α} : (insertManyIfNewUnit m l).1.1.isEmpty = (m.1.isEmpty && l.isEmpty) := by From b233bc5dd8b6655099d40047c4e297e0bf733196 Mon Sep 17 00:00:00 2001 From: jt0202 Date: Thu, 12 Dec 2024 21:13:51 +0100 Subject: [PATCH 51/83] Moved results up to Raw --- src/Std/Data/DHashMap/Internal/Raw.lean | 18 + src/Std/Data/DHashMap/Internal/RawLemmas.lean | 4 +- src/Std/Data/DHashMap/RawLemmas.lean | 465 +++++++++++++++++- 3 files changed, 483 insertions(+), 4 deletions(-) diff --git a/src/Std/Data/DHashMap/Internal/Raw.lean b/src/Std/Data/DHashMap/Internal/Raw.lean index a735d30d7dd3..5c9b4c7606cb 100644 --- a/src/Std/Data/DHashMap/Internal/Raw.lean +++ b/src/Std/Data/DHashMap/Internal/Raw.lean @@ -211,6 +211,24 @@ section variable {β : Type v} +theorem Const.insertMany_eq [BEq α] [Hashable α] {m : Raw α (fun _ => β)} (h : m.WF) {ρ : Type w} [ForIn Id ρ (α × β)] {l : ρ} : + Raw.Const.insertMany m l = Raw₀.Const.insertMany ⟨m, h.size_buckets_pos⟩ l := by + simp [Raw.Const.insertMany, h.size_buckets_pos] + +theorem Const.insertMany_val [BEq α][Hashable α] {m : Raw₀ α (fun _ => β)} {ρ : Type w} [ForIn Id ρ (α × β)] {l : ρ} : + Raw.Const.insertMany m.val l = Raw₀.Const.insertMany m l := by + simp [Raw.Const.insertMany, m.2] + +theorem Const.insertManyIfNewUnit_eq {ρ : Type w} [ForIn Id ρ α] [BEq α] [Hashable α] + {m : Raw α (fun _ => Unit)} {l : ρ} (h : m.WF): + Raw.Const.insertManyIfNewUnit m l = Raw₀.Const.insertManyIfNewUnit ⟨m, h.size_buckets_pos⟩ l := by + simp [Raw.Const.insertManyIfNewUnit, h.size_buckets_pos] + +theorem Const.insertManyIfNewUnit_val {ρ : Type w} [ForIn Id ρ α] [BEq α] [Hashable α] + {m : Raw₀ α (fun _ => Unit)} {l : ρ} : + Raw.Const.insertManyIfNewUnit m.val l = Raw₀.Const.insertManyIfNewUnit m l := by + simp [Raw.Const.insertManyIfNewUnit, m.2] + theorem Const.get?_eq [BEq α] [Hashable α] {m : Raw α (fun _ => β)} (h : m.WF) {a : α} : Raw.Const.get? m a = Raw₀.Const.get? ⟨m, h.size_buckets_pos⟩ a := by simp [Raw.Const.get?, h.size_buckets_pos] diff --git a/src/Std/Data/DHashMap/Internal/RawLemmas.lean b/src/Std/Data/DHashMap/Internal/RawLemmas.lean index b5a9b905af91..060155162be9 100644 --- a/src/Std/Data/DHashMap/Internal/RawLemmas.lean +++ b/src/Std/Data/DHashMap/Internal/RawLemmas.lean @@ -999,10 +999,9 @@ theorem isEmpty_insertMany_list [EquivBEq α] [LawfulHashable α] (h : m.1.WF) (m.insertMany l).1.1.isEmpty = (m.1.isEmpty && l.isEmpty) := by simp_to_model using isEmpty_insertList -section -variable {β : Type v} (m : Raw₀ α (fun _ => β)) namespace Const +variable {β : Type v} (m : Raw₀ α (fun _ => β)) @[simp] theorem insertMany_nil : @@ -1304,7 +1303,6 @@ theorem getValueD_insertManyIfNewUnit_list simp end Const -end end Raw₀ end Std.DHashMap.Internal diff --git a/src/Std/Data/DHashMap/RawLemmas.lean b/src/Std/Data/DHashMap/RawLemmas.lean index 8b8bb67a68db..134f2e05d24a 100644 --- a/src/Std/Data/DHashMap/RawLemmas.lean +++ b/src/Std/Data/DHashMap/RawLemmas.lean @@ -58,7 +58,10 @@ private def baseNames : Array Name := ``getKey?_eq, ``getKey?_val, ``getKey_eq, ``getKey_val, ``getKey!_eq, ``getKey!_val, - ``getKeyD_eq, ``getKeyD_val] + ``getKeyD_eq, ``getKeyD_val, + ``insertMany_eq, ``insertMany_val, + ``Const.insertMany_eq, ``Const.insertMany_val, + ``Const.insertManyIfNewUnit_eq, ``Const.insertManyIfNewUnit_val] /-- Internal implementation detail of the hash map -/ scoped syntax "simp_to_raw" ("using" term)? : tactic @@ -1047,6 +1050,466 @@ theorem distinct_keys [EquivBEq α] [LawfulHashable α] (h : m.WF) : m.keys.Pairwise (fun a b => (a == b) = false) := by simp_to_raw using Raw₀.distinct_keys ⟨m, h.size_buckets_pos⟩ h +@[simp] +theorem insertMany_nil [EquivBEq α] [LawfulHashable α] (h : m.WF) : + m.insertMany [] = m := by + simp_to_raw + rw [Raw₀.insertMany_nil] + +@[simp] +theorem insertMany_list_singleton {k : α} {v : β k} [EquivBEq α] [LawfulHashable α] (h : m.WF) : + m.insertMany [⟨k, v⟩] = m.insert k v := by + simp_to_raw + rw [Raw₀.insertMany_list_singleton] + +theorem insertMany_cons {l : List ((a : α) × β a)} {k : α} {v : β k} [EquivBEq α] [LawfulHashable α] (h : m.WF) : + (m.insertMany (⟨k, v⟩ :: l)).1 = ((m.insert k v).insertMany l).1 := by + simp_to_raw + rw [Raw₀.insertMany_cons] + +theorem contains_insertMany_list [EquivBEq α] [LawfulHashable α] (h : m.WF) + {l : List ((a : α) × β a)} {k : α} : + (m.insertMany l).contains k = (m.contains k || (l.map Sigma.fst).contains k) := by + simp_to_raw using Raw₀.contains_insertMany_list + +theorem contains_of_contains_insertMany_list [EquivBEq α] [LawfulHashable α] (h : m.WF) + {l : List ((a : α) × β a)} {k : α} : + (m.insertMany l).contains k → (l.map Sigma.fst).contains k = false → m.contains k := by + simp_to_raw using Raw₀.contains_of_contains_insertMany_list + +theorem get?_insertMany_list_of_contains_eq_false [LawfulBEq α] (h : m.WF) + {l : List ((a : α) × β a)} {k : α} + (h' : (l.map Sigma.fst).contains k = false) : + (m.insertMany l).get? k = m.get? k := by + simp_to_raw using Raw₀.get?_insertMany_list_of_contains_eq_false + +theorem get?_insertMany_list_of_mem [LawfulBEq α] (h : m.WF) + {l : List ((a : α) × β a)} {k k' : α} (k_beq : k == k') {v : β k} + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) + (mem : ⟨k, v⟩ ∈ l) : + (m.insertMany l).get? k' = some (cast (by congr; apply LawfulBEq.eq_of_beq k_beq) v) := by + simp_to_raw using Raw₀.get?_insertMany_list_of_mem + +theorem get_insertMany_list_of_contains_eq_false [LawfulBEq α] (h : m.WF) + {l : List ((a : α) × β a)} {k : α} + (contains : (l.map Sigma.fst).contains k = false) + {h'} : + (m.insertMany l).get k h' = + m.get k (contains_of_contains_insertMany_list h h' contains) := by + simp_to_raw using Raw₀.get_insertMany_list_of_contains_eq_false + +theorem get_insertMany_list_of_mem [LawfulBEq α] (h : m.WF) + {l : List ((a : α) × β a)} {k k' : α} (k_beq : k == k') {v : β k} + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) + (mem : ⟨k, v⟩ ∈ l) + {h'} : + (m.insertMany l).get k' h' = cast (by congr; apply LawfulBEq.eq_of_beq k_beq) v := by + simp_to_raw using Raw₀.get_insertMany_list_of_mem + +theorem get!_insertMany_list_of_contains_eq_false [LawfulBEq α] (h : m.WF) + {l : List ((a : α) × β a)} {k : α} [Inhabited (β k)] + (h' : (l.map Sigma.fst).contains k = false) : + (m.insertMany l).get! k = m.get! k := by + simp_to_raw using Raw₀.get!_insertMany_list_of_contains_eq_false + +theorem get!_insertMany_list_of_mem [LawfulBEq α] (h : m.WF) + {l : List ((a : α) × β a)} {k k' : α} (k_beq : k == k') {v : β k} [Inhabited (β k')] + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) + (mem : ⟨k, v⟩ ∈ l) : + (m.insertMany l).get! k' = cast (by congr; apply LawfulBEq.eq_of_beq k_beq) v := by + simp_to_raw using Raw₀.get!_insertMany_list_of_mem + +theorem getD_insertMany_list_of_contains_eq_false [LawfulBEq α] (h : m.WF) + {l : List ((a : α) × β a)} {k : α} {fallback : β k} + (not_mem : (l.map Sigma.fst).contains k = false) : + (m.insertMany l).getD k fallback = m.getD k fallback := by + simp_to_raw using Raw₀.getD_insertMany_list_of_contains_eq_false + +theorem getD_insertMany_list_of_mem [LawfulBEq α] (h : m.WF) + {l : List ((a : α) × β a)} {k k' : α} (k_beq : k == k') {v : β k} {fallback : β k'} + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) + (mem : ⟨k, v⟩ ∈ l) : + (m.insertMany l).getD k' fallback = cast (by congr; apply LawfulBEq.eq_of_beq k_beq) v := by + simp_to_raw using Raw₀.getD_insertMany_list_of_mem + +theorem getKey?_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] (h : m.WF) + {l : List ((a : α) × β a)} {k : α} + (h : (l.map Sigma.fst).contains k = false) : + (m.insertMany l).getKey? k = m.getKey? k := by + simp_to_raw using Raw₀.getKey?_insertMany_list_of_contains_eq_false + +theorem getKey?_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] (h : m.WF) + {l : List ((a : α) × β a)} + {k k' : α} (k_beq : k == k') + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) + (mem : k ∈ l.map Sigma.fst) : + (m.insertMany l).getKey? k' = some k := by + simp_to_raw using Raw₀.getKey?_insertMany_list_of_mem + +theorem getKey_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] (h : m.WF) + {l : List ((a : α) × β a)} {k : α} + (h₁ : (l.map Sigma.fst).contains k = false) + {h'} : + (m.insertMany l).getKey k h' = + m.getKey k (contains_of_contains_insertMany_list h h' h₁) := by + simp_to_raw using Raw₀.getKey_insertMany_list_of_contains_eq_false + +theorem getKey_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] (h : m.WF) + {l : List ((a : α) × β a)} + {k k' : α} (k_beq : k == k') + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) + (mem : k ∈ l.map Sigma.fst) + {h'} : + (m.insertMany l).getKey k' h' = k := by + simp_to_raw using Raw₀.getKey_insertMany_list_of_mem + +theorem getKey!_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] [Inhabited α] + (h : m.WF) {l : List ((a : α) × β a)} {k : α} + (h' : (l.map Sigma.fst).contains k = false) : + (m.insertMany l).getKey! k = m.getKey! k := by + simp_to_raw using Raw₀.getKey!_insertMany_list_of_contains_eq_false + +theorem getKey!_insertMany_listof_mem [EquivBEq α] [LawfulHashable α] [Inhabited α] (h : m.WF) + {l : List ((a : α) × β a)} + {k k' : α} (k_beq : k == k') + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) + (mem : k ∈ l.map Sigma.fst) : + (m.insertMany l).getKey! k' = k := by + simp_to_raw using Raw₀.getKey!_insertMany_listof_mem + +theorem getKeyD_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] (h : m.WF) + {l : List ((a : α) × β a)} {k fallback : α} + (h' : (l.map Sigma.fst).contains k = false) : + (m.insertMany l).getKeyD k fallback = m.getKeyD k fallback := by + simp_to_raw using Raw₀.getKeyD_insertMany_list_of_contains_eq_false + +theorem getKeyD_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] (h : m.WF) + {l : List ((a : α) × β a)} + {k k' fallback : α} (k_beq : k == k') + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) + (mem : k ∈ l.map Sigma.fst) : + (m.insertMany l).getKeyD k' fallback = k := by + simp_to_raw using Raw₀.getKeyD_insertMany_list_of_mem + +theorem size_insertMany_list [EquivBEq α] [LawfulHashable α] (h : m.WF) + {l : List ((a : α) × β a)} {distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)} : + (∀ (a : α), m.contains a → (l.map Sigma.fst).contains a = false) → + (m.insertMany l).size = m.size + l.length := by + simp_to_raw using Raw₀.size_insertMany_list + +theorem size_insertMany_list_le [EquivBEq α] [LawfulHashable α] (h : m.WF) + {l : List ((a : α) × β a)} : + (m.insertMany l).size ≤ m.size + l.length := by + simp_to_raw using Raw₀.size_insertMany_list_le + +theorem isEmpty_insertMany_list [EquivBEq α] [LawfulHashable α] (h : m.WF) + {l : List ((a : α) × β a)} : + (m.insertMany l).isEmpty = (m.isEmpty && l.isEmpty) := by + simp_to_raw using Raw₀.isEmpty_insertMany_list + +namespace Const +variable {β : Type v} (m : Raw α (fun _ => β)) + +@[simp] +theorem insertMany_nil [EquivBEq α] [LawfulHashable α] (h : m.WF) : + insertMany m [] = m := by + simp_to_raw + rw [Raw₀.Const.insertMany_nil] + +@[simp] +theorem insertMany_list_singleton {k : α} {v : β} [EquivBEq α] [LawfulHashable α] (h : m.WF) : + insertMany m [⟨k, v⟩] = m.insert k v := by + simp_to_raw + rw [Raw₀.Const.insertMany_list_singleton] + +theorem insertMany_cons {l : List (α × β)} {k : α} {v : β} [EquivBEq α] [LawfulHashable α] (h : m.WF) : + (insertMany m (⟨k, v⟩ :: l)).1 = (insertMany (m.insert k v) l).1 := by + simp_to_raw + rw [Raw₀.Const.insertMany_cons] + +theorem contains_insertMany_list [EquivBEq α] [LawfulHashable α] (h : m.WF) + {l : List (α × β)} {k : α} : + (insertMany m l).contains k = (m.contains k || (l.map Prod.fst).contains k) := by + simp_to_raw using Raw₀.Const.contains_insertMany_list + +theorem contains_of_contains_insertMany_list [EquivBEq α] [LawfulHashable α] (h : m.WF) + {l : List ( α × β )} {k : α} : + (insertMany m l).contains k → (l.map Prod.fst).contains k = false → m.contains k := by + simp_to_raw using Raw₀.Const.contains_of_contains_insertMany_list + +theorem getKey?_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] (h : m.WF) + {l : List (α × β)} {k : α} + (h' : (l.map Prod.fst).contains k = false) : + (insertMany m l).getKey? k = m.getKey? k := by + simp_to_raw using Raw₀.Const.getKey?_insertMany_list_of_contains_eq_false + +theorem getKey?_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] (h : m.WF) + {l : List (α × β)} + {k k' : α} (k_beq : k == k') + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) + (mem : k ∈ l.map Prod.fst) : + (insertMany m l).getKey? k' = some k := by + simp_to_raw using Raw₀.Const.getKey?_insertMany_list_of_mem + +theorem getKey_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] (h : m.WF) + {l : List (α × β)} {k : α} + {h₁ : (l.map Prod.fst).contains k = false} + {h'} : + (insertMany m l).getKey k h' = + m.getKey k (contains_of_contains_insertMany_list _ h h' h₁) := by + simp_to_raw using Raw₀.Const.getKey_insertMany_list_of_contains_eq_false + +theorem getKey_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] (h : m.WF) + {l : List (α × β)} + {k k' : α} (k_beq : k == k') + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) + (mem : k ∈ l.map Prod.fst) + {h'} : + (insertMany m l).getKey k' h' = k := by + simp_to_raw using Raw₀.Const.getKey_insertMany_list_of_mem + +theorem getKey!_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] [Inhabited α] + (h : m.WF) {l : List (α × β)} {k : α} + (h' : (l.map Prod.fst).contains k = false) : + (insertMany m l).getKey! k = m.getKey! k := by + simp_to_raw using Raw₀.Const.getKey!_insertMany_list_of_contains_eq_false + +theorem getKey!_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] [Inhabited α] (h : m.WF) + {l : List (α × β)} + {k k' : α} (k_beq : k == k') + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) + (mem : k ∈ l.map Prod.fst) : + (insertMany m l).getKey! k' = k := by + simp_to_raw using Raw₀.Const.getKey!_insertMany_list_of_mem + +theorem getKeyD_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] (h : m.WF) + {l : List (α × β)} {k fallback : α} + (h' : (l.map Prod.fst).contains k = false) : + (insertMany m l).getKeyD k fallback = m.getKeyD k fallback := by + simp_to_raw using Raw₀.Const.getKeyD_insertMany_list_of_contains_eq_false + +theorem getKeyD_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] (h : m.WF) + {l : List (α × β)} + {k k' fallback : α} (k_beq : k == k') + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) + (mem : k ∈ l.map Prod.fst) : + (insertMany m l).getKeyD k' fallback = k := by + simp_to_raw using Raw₀.Const.getKeyD_insertMany_list_of_mem + +theorem size_insertMany_list [EquivBEq α] [LawfulHashable α] (h : m.WF) + {l : List (α × β)} + {distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)} : + (∀ (a : α), m.contains a → (l.map Prod.fst).contains a = false) → + (insertMany m l).size = m.size + l.length := by + simp_to_raw using Raw₀.Const.size_insertMany_list + +theorem size_insertMany_list_le [EquivBEq α] [LawfulHashable α] (h : m.WF) + {l : List (α × β)} : + (insertMany m l).size ≤ m.size + l.length := by + simp_to_raw using Raw₀.Const.size_insertMany_list_le + +theorem isEmpty_insertMany_list [EquivBEq α] [LawfulHashable α] (h : m.WF) + {l : List (α × β)} : + (insertMany m l).isEmpty = (m.isEmpty && l.isEmpty) := by + simp_to_raw using Raw₀.Const.isEmpty_insertMany_list + +theorem get?_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] (h : m.WF) + {l : List (α × β)} {k : α} + (h' : (l.map Prod.fst).contains k = false) : + get? (insertMany m l) k = get? m k := by + simp_to_raw using Raw₀.Const.get?_insertMany_list_of_contains_eq_false + +theorem get?_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] (h : m.WF) + {l : List (α × β)} {k k' : α} (k_beq : k == k') {v : β} (mem : ⟨k, v⟩ ∈ l) + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) : + get? (insertMany m l) k' = v := by + simp_to_raw using Raw₀.Const.get?_insertMany_list_of_mem + +theorem get_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] (h : m.WF) + {l : List (α × β)} {k : α} + (h₁ : (l.map Prod.fst).contains k = false) + {h'} : + get (insertMany m l) k h' = get m k (contains_of_contains_insertMany_list _ h h' h₁) := by + simp_to_raw using Raw₀.Const.get_insertMany_list_of_contains_eq_false + +theorem get_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] (h : m.WF) + {l : List (α × β)} {k k' : α} (k_beq : k == k') {v : β} (mem : ⟨k, v⟩ ∈ l) + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) {h'} : + get (insertMany m l) k' h' = v := by + simp_to_raw using Raw₀.Const.get_insertMany_list_of_mem + +theorem get!_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] + [Inhabited β] (h : m.WF) {l : List (α × β)} {k : α} + (h' : (l.map Prod.fst).contains k = false) : + get! (insertMany m l) k = get! m k := by + simp_to_raw using Raw₀.Const.get!_insertMany_list_of_contains_eq_false + +theorem get!_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] [Inhabited β] (h : m.WF) + {l : List (α × β)} {k k' : α} (k_beq : k == k') {v : β} (mem : ⟨k, v⟩ ∈ l) + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) : + get! (insertMany m l) k' = v := by + simp_to_raw using Raw₀.Const.get!_insertMany_list_of_mem + +theorem getD_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] (h : m.WF) + {l : List (α × β)} {k : α} {fallback : β} + (h' : (l.map Prod.fst).contains k = false) : + getD (insertMany m l) k fallback = getD m k fallback := by + simp_to_raw using Raw₀.Const.getD_insertMany_list_of_contains_eq_false + +theorem getD_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] (h : m.WF) + {l : List (α × β)} {k k' : α} (k_beq : k == k') {v fallback : β} (mem : ⟨k, v⟩ ∈ l) + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) : + getD (insertMany m l) k' fallback = v := by + simp_to_raw using Raw₀.Const.getD_insertMany_list_of_mem + +variable (m : Raw α (fun _ => Unit)) + +@[simp] +theorem insertManyIfNewUnit_nil [EquivBEq α] [LawfulHashable α] (h : m.WF) : + insertManyIfNewUnit m [] = m := by + simp_to_raw + rw [Raw₀.Const.insertManyIfNewUnit_nil] + +@[simp] +theorem insertManyIfNewUnit_list_singleton [EquivBEq α] [LawfulHashable α] {k : α} (h : m.WF) : + insertManyIfNewUnit m [k] = m.insertIfNew k () := by + simp_to_raw + rw [Raw₀.Const.insertManyIfNewUnit_list_singleton] + +theorem insertManyIfNewUnit_cons [EquivBEq α] [LawfulHashable α] (h : m.WF) {l : List α} {k : α} : + (insertManyIfNewUnit m (k :: l)).1 = (insertManyIfNewUnit (m.insertIfNew k ()) l).1 := by + simp_to_raw + rw [Raw₀.Const.insertManyIfNewUnit_cons] + +theorem contains_insertManyIfNewUnit_list [EquivBEq α] [LawfulHashable α] (h : m.WF) + {l : List α} {k : α} : + (insertManyIfNewUnit m l).contains k = (m.contains k || l.contains k) := by + simp_to_raw using Raw₀.Const.contains_insertManyIfNewUnit_list + +theorem contains_of_contains_insertManyIfNewUnit_list [EquivBEq α] [LawfulHashable α] (h : m.WF) + {l : List α} {k : α} (h₂ : l.contains k = false) : + (insertManyIfNewUnit m l).contains k → m.contains k := by + simp_to_raw using Raw₀.Const.contains_of_contains_insertManyIfNewUnit_list + +theorem getKey?_insertManyIfNewUnit_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] + (h : m.WF) {l : List α} {k : α} (h' : l.contains k = false) : + getKey? (insertManyIfNewUnit m l) k = getKey? m k := by + simp_to_raw using Raw₀.Const.getKey?_insertManyIfNewUnit_list_of_contains_eq_false + +theorem getKey?_insertManyIfNewUnit_list_of_mem_of_contains_eq_false [EquivBEq α] [LawfulHashable α] + (h : m.WF) {l : List α} {k k' : α} (k_beq : k == k') + (distinct : l.Pairwise (fun a b => (a == b) = false)) (mem : k ∈ l) : + m.contains k = false → getKey? (insertManyIfNewUnit m l) k' = some k := by + simp_to_raw using Raw₀.Const.getKey?_insertManyIfNewUnit_list_of_mem_of_contains_eq_false + +theorem getKey?_insertManyIfNewUnit_list_of_contains_of_contains [EquivBEq α] [LawfulHashable α] + (h : m.WF) {l : List α} {k : α} + (distinct : l.Pairwise (fun a b => (a == b) = false)) + (h' : l.contains k) : + m.contains k → getKey? (insertManyIfNewUnit m l) k = getKey? m k := by + simp_to_raw using Raw₀.Const.getKey?_insertManyIfNewUnit_list_of_contains_of_contains + +theorem getKey_insertManyIfNewUnit_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] + (h : m.WF) {l : List α} {k : α} (h₁ : l.contains k = false) {h'} : + getKey (insertManyIfNewUnit m l) k h' = + getKey m k (contains_of_contains_insertManyIfNewUnit_list m h h₁ h') := by + simp_to_raw using Raw₀.Const.getKey_insertManyIfNewUnit_list_of_contains_eq_false + +theorem getKey_insertManyIfNewUnit_list_of_mem_of_contains_eq_false [EquivBEq α] [LawfulHashable α] + (h : m.WF) {l : List α} + {k k' : α} (k_beq : k == k') + (distinct : l.Pairwise (fun a b => (a == b) = false)) + (mem : k ∈ l) {h'} : + m.contains k = false → + getKey (insertManyIfNewUnit m l) k' h' = k := by + simp_to_raw using Raw₀.Const.getKey_insertManyIfNewUnit_list_of_mem_of_contains_eq_false + +theorem getKey_insertManyIfNewUnit_list_mem_of_contains_of_contains [EquivBEq α] [LawfulHashable α] + (h : m.WF) {l : List α} {k : α} + (distinct : l.Pairwise (fun a b => (a == b) = false)) + (h' : l.contains k) {mem' : m.contains k} {h'} : + getKey (insertManyIfNewUnit m l) k h' = getKey m k mem' := by + simp_to_raw using Raw₀.Const.getKey_insertManyIfNewUnit_list_mem_of_contains_of_contains + +theorem getKey!_insertManyIfNewUnit_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] + [Inhabited α] (h : m.WF) {l : List α} {k : α} + (h' : l.contains k = false) : + getKey! (insertManyIfNewUnit m l) k = getKey! m k := by + simp_to_raw using Raw₀.Const.getKey!_insertManyIfNewUnit_list_of_contains_eq_false + +theorem getKey!_insertManyIfNewUnit_list_of_mem_of_contains_eq_false [EquivBEq α] [LawfulHashable α] + [Inhabited α] (h : m.WF) {l : List α} {k k' : α} (k_beq : k == k') + (distinct : l.Pairwise (fun a b => (a == b) = false)) + (mem : k ∈ l) : contains m k = false → + getKey! (insertManyIfNewUnit m l) k' = k := by + simp_to_raw using Raw₀.Const.getKey!_insertManyIfNewUnit_list_of_mem_of_contains_eq_false + +theorem getKey!_insertManyIfNewUnit_list_mem_of_contains_of_contains [EquivBEq α] [LawfulHashable α] + [Inhabited α] (h : m.WF) {l : List α} {k : α} + (distinct : l.Pairwise (fun a b => (a == b) = false)) + (h' : l.contains k) : m.contains k → + getKey! (insertManyIfNewUnit m l) k = getKey! m k := by + simp_to_raw using Raw₀.Const.getKey!_insertManyIfNewUnit_list_mem_of_contains_of_contains + +theorem getKeyD_insertManyIfNewUnit_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] + (h : m.WF) {l : List α} {k fallback : α} + (h' : l.contains k = false) : + getKeyD (insertManyIfNewUnit m l) k fallback = getKeyD m k fallback := by + simp_to_raw using Raw₀.Const.getKeyD_insertManyIfNewUnit_list_of_contains_eq_false + +theorem getKeyD_insertManyIfNewUnit_list_of_mem_of_contains_eq_false [EquivBEq α] [LawfulHashable α] + (h : m.WF) {l : List α} {k k' fallback : α} (k_beq : k == k') + (distinct : l.Pairwise (fun a b => (a == b) = false)) + {mem : k ∈ l } : m.contains k = false → + getKeyD (insertManyIfNewUnit m l) k' fallback = k := by + simp_to_raw using Raw₀.Const.getKeyD_insertManyIfNewUnit_list_of_mem_of_contains_eq_false + +theorem getKeyD_insertManyIfNewUnit_list_of_contains_of_contains [EquivBEq α] [LawfulHashable α] + (h : m.WF) {l : List α} {k fallback : α} + (distinct : l.Pairwise (fun a b => (a == b) = false)) + (h' : l.contains k) : m.contains k → + getKeyD (insertManyIfNewUnit m l) k fallback = getKeyD m k fallback := by + simp_to_raw using Raw₀.Const.getKeyD_insertManyIfNewUnit_list_of_contains_of_contains + +theorem size_insertManyIfNewUnit_list [EquivBEq α] [LawfulHashable α] (h : m.WF) + {l : List α} + (distinct : l.Pairwise (fun a b => (a == b) = false)) : + (∀ (a : α), m.contains a → l.contains a = false) → + (insertManyIfNewUnit m l).size = m.size + l.length := by + simp_to_raw using Raw₀.Const.size_insertManyIfNewUnit_list + +theorem size_insertManyIfNewUnit_list_le [EquivBEq α] [LawfulHashable α] (h : m.WF) + {l : List α} : + (insertManyIfNewUnit m l).size ≤ m.size + l.length := by + simp_to_raw using Raw₀.Const.size_insertManyIfNewUnit_list_le + +theorem isEmpty_insertManyIfNewUnit_list [EquivBEq α] [LawfulHashable α] (h : m.WF) + {l : List α} : + (insertManyIfNewUnit m l).isEmpty = (m.isEmpty && l.isEmpty) := by + simp_to_raw using Raw₀.Const.isEmpty_insertManyIfNewUnit_list + +theorem get?_insertManyIfNewUnit_list [EquivBEq α] [LawfulHashable α] (h : m.WF) + {l : List α} {k : α} : + get? (insertManyIfNewUnit m l) k = + if m.contains k ∨ l.contains k then some () else none := by + simp_to_raw using Raw₀.Const.get?_insertManyIfNewUnit_list + +theorem getValue_insertManyIfNewUnit_list + {l : List α} {k : α} {h} : + get (insertManyIfNewUnit m l) k h = () := by + simp + +theorem getValue!_insertManyIfNewUnit_list + {l : List α} {k : α} : + get! (insertManyIfNewUnit m l) k = () := by + simp + +theorem getValueD_insertManyIfNewUnit_list + {l : List α} {k : α} {fallback : Unit} : + getD (insertManyIfNewUnit m l) k fallback = () := by + simp + +end Const end Raw end Std.DHashMap From 41acb69105553dc3aa3d91ddc202e35706d6a3c8 Mon Sep 17 00:00:00 2001 From: jt0202 Date: Thu, 12 Dec 2024 22:45:30 +0100 Subject: [PATCH 52/83] ofList lemmas introduced --- src/Std/Data/DHashMap/Internal/RawLemmas.lean | 2 +- src/Std/Data/DHashMap/RawLemmas.lean | 475 +++++++++++++++++- 2 files changed, 455 insertions(+), 22 deletions(-) diff --git a/src/Std/Data/DHashMap/Internal/RawLemmas.lean b/src/Std/Data/DHashMap/Internal/RawLemmas.lean index 060155162be9..6247c117aebd 100644 --- a/src/Std/Data/DHashMap/Internal/RawLemmas.lean +++ b/src/Std/Data/DHashMap/Internal/RawLemmas.lean @@ -961,7 +961,7 @@ theorem getKey!_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashab (m.insertMany l).1.getKey! k = m.getKey! k := by simp_to_model using getKey!_insertList_of_contains_eq_false -theorem getKey!_insertMany_listof_mem [EquivBEq α] [LawfulHashable α] [Inhabited α] (h : m.1.WF) +theorem getKey!_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] [Inhabited α] (h : m.1.WF) {l : List ((a : α) × β a)} {k k' : α} (k_beq : k == k') (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) diff --git a/src/Std/Data/DHashMap/RawLemmas.lean b/src/Std/Data/DHashMap/RawLemmas.lean index 134f2e05d24a..40cd2c5e91f5 100644 --- a/src/Std/Data/DHashMap/RawLemmas.lean +++ b/src/Std/Data/DHashMap/RawLemmas.lean @@ -1063,7 +1063,7 @@ theorem insertMany_list_singleton {k : α} {v : β k} [EquivBEq α] [LawfulHasha rw [Raw₀.insertMany_list_singleton] theorem insertMany_cons {l : List ((a : α) × β a)} {k : α} {v : β k} [EquivBEq α] [LawfulHashable α] (h : m.WF) : - (m.insertMany (⟨k, v⟩ :: l)).1 = ((m.insert k v).insertMany l).1 := by + m.insertMany (⟨k, v⟩ :: l) = (m.insert k v).insertMany l := by simp_to_raw rw [Raw₀.insertMany_cons] @@ -1121,7 +1121,7 @@ theorem get!_insertMany_list_of_mem [LawfulBEq α] (h : m.WF) theorem getD_insertMany_list_of_contains_eq_false [LawfulBEq α] (h : m.WF) {l : List ((a : α) × β a)} {k : α} {fallback : β k} - (not_mem : (l.map Sigma.fst).contains k = false) : + (contains_eq_false : (l.map Sigma.fst).contains k = false) : (m.insertMany l).getD k fallback = m.getD k fallback := by simp_to_raw using Raw₀.getD_insertMany_list_of_contains_eq_false @@ -1169,13 +1169,13 @@ theorem getKey!_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashab (m.insertMany l).getKey! k = m.getKey! k := by simp_to_raw using Raw₀.getKey!_insertMany_list_of_contains_eq_false -theorem getKey!_insertMany_listof_mem [EquivBEq α] [LawfulHashable α] [Inhabited α] (h : m.WF) +theorem getKey!_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] [Inhabited α] (h : m.WF) {l : List ((a : α) × β a)} {k k' : α} (k_beq : k == k') (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) (mem : k ∈ l.map Sigma.fst) : (m.insertMany l).getKey! k' = k := by - simp_to_raw using Raw₀.getKey!_insertMany_listof_mem + simp_to_raw using Raw₀.getKey!_insertMany_list_of_mem theorem getKeyD_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] (h : m.WF) {l : List ((a : α) × β a)} {k fallback : α} @@ -1192,7 +1192,7 @@ theorem getKeyD_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] (h : m. simp_to_raw using Raw₀.getKeyD_insertMany_list_of_mem theorem size_insertMany_list [EquivBEq α] [LawfulHashable α] (h : m.WF) - {l : List ((a : α) × β a)} {distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)} : + {l : List ((a : α) × β a)} (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) : (∀ (a : α), m.contains a → (l.map Sigma.fst).contains a = false) → (m.insertMany l).size = m.size + l.length := by simp_to_raw using Raw₀.size_insertMany_list @@ -1208,7 +1208,7 @@ theorem isEmpty_insertMany_list [EquivBEq α] [LawfulHashable α] (h : m.WF) simp_to_raw using Raw₀.isEmpty_insertMany_list namespace Const -variable {β : Type v} (m : Raw α (fun _ => β)) +variable {β : Type v} {m : Raw α (fun _ => β)} @[simp] theorem insertMany_nil [EquivBEq α] [LawfulHashable α] (h : m.WF) : @@ -1223,7 +1223,7 @@ theorem insertMany_list_singleton {k : α} {v : β} [EquivBEq α] [LawfulHashabl rw [Raw₀.Const.insertMany_list_singleton] theorem insertMany_cons {l : List (α × β)} {k : α} {v : β} [EquivBEq α] [LawfulHashable α] (h : m.WF) : - (insertMany m (⟨k, v⟩ :: l)).1 = (insertMany (m.insert k v) l).1 := by + insertMany m (⟨k, v⟩ :: l) = insertMany (m.insert k v) l := by simp_to_raw rw [Raw₀.Const.insertMany_cons] @@ -1256,7 +1256,7 @@ theorem getKey_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashabl {h₁ : (l.map Prod.fst).contains k = false} {h'} : (insertMany m l).getKey k h' = - m.getKey k (contains_of_contains_insertMany_list _ h h' h₁) := by + m.getKey k (contains_of_contains_insertMany_list h h' h₁) := by simp_to_raw using Raw₀.Const.getKey_insertMany_list_of_contains_eq_false theorem getKey_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] (h : m.WF) @@ -1298,7 +1298,7 @@ theorem getKeyD_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] (h : m. theorem size_insertMany_list [EquivBEq α] [LawfulHashable α] (h : m.WF) {l : List (α × β)} - {distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)} : + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) : (∀ (a : α), m.contains a → (l.map Prod.fst).contains a = false) → (insertMany m l).size = m.size + l.length := by simp_to_raw using Raw₀.Const.size_insertMany_list @@ -1320,8 +1320,8 @@ theorem get?_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable simp_to_raw using Raw₀.Const.get?_insertMany_list_of_contains_eq_false theorem get?_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] (h : m.WF) - {l : List (α × β)} {k k' : α} (k_beq : k == k') {v : β} (mem : ⟨k, v⟩ ∈ l) - (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) : + {l : List (α × β)} {k k' : α} (k_beq : k == k') {v : β} + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) (mem : ⟨k, v⟩ ∈ l) : get? (insertMany m l) k' = v := by simp_to_raw using Raw₀.Const.get?_insertMany_list_of_mem @@ -1329,12 +1329,12 @@ theorem get_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable {l : List (α × β)} {k : α} (h₁ : (l.map Prod.fst).contains k = false) {h'} : - get (insertMany m l) k h' = get m k (contains_of_contains_insertMany_list _ h h' h₁) := by + get (insertMany m l) k h' = get m k (contains_of_contains_insertMany_list h h' h₁) := by simp_to_raw using Raw₀.Const.get_insertMany_list_of_contains_eq_false theorem get_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] (h : m.WF) - {l : List (α × β)} {k k' : α} (k_beq : k == k') {v : β} (mem : ⟨k, v⟩ ∈ l) - (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) {h'} : + {l : List (α × β)} {k k' : α} (k_beq : k == k') {v : β} + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) (mem : ⟨k, v⟩ ∈ l){h'} : get (insertMany m l) k' h' = v := by simp_to_raw using Raw₀.Const.get_insertMany_list_of_mem @@ -1345,8 +1345,8 @@ theorem get!_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable simp_to_raw using Raw₀.Const.get!_insertMany_list_of_contains_eq_false theorem get!_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] [Inhabited β] (h : m.WF) - {l : List (α × β)} {k k' : α} (k_beq : k == k') {v : β} (mem : ⟨k, v⟩ ∈ l) - (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) : + {l : List (α × β)} {k k' : α} (k_beq : k == k') {v : β} + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) (mem : ⟨k, v⟩ ∈ l) : get! (insertMany m l) k' = v := by simp_to_raw using Raw₀.Const.get!_insertMany_list_of_mem @@ -1357,12 +1357,12 @@ theorem getD_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable simp_to_raw using Raw₀.Const.getD_insertMany_list_of_contains_eq_false theorem getD_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] (h : m.WF) - {l : List (α × β)} {k k' : α} (k_beq : k == k') {v fallback : β} (mem : ⟨k, v⟩ ∈ l) - (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) : + {l : List (α × β)} {k k' : α} (k_beq : k == k') {v fallback : β} + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) (mem : ⟨k, v⟩ ∈ l) : getD (insertMany m l) k' fallback = v := by simp_to_raw using Raw₀.Const.getD_insertMany_list_of_mem -variable (m : Raw α (fun _ => Unit)) +variable {m : Raw α (fun _ => Unit)} @[simp] theorem insertManyIfNewUnit_nil [EquivBEq α] [LawfulHashable α] (h : m.WF) : @@ -1412,7 +1412,7 @@ theorem getKey?_insertManyIfNewUnit_list_of_contains_of_contains [EquivBEq α] [ theorem getKey_insertManyIfNewUnit_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] (h : m.WF) {l : List α} {k : α} (h₁ : l.contains k = false) {h'} : getKey (insertManyIfNewUnit m l) k h' = - getKey m k (contains_of_contains_insertManyIfNewUnit_list m h h₁ h') := by + getKey m k (contains_of_contains_insertManyIfNewUnit_list h h₁ h') := by simp_to_raw using Raw₀.Const.getKey_insertManyIfNewUnit_list_of_contains_eq_false theorem getKey_insertManyIfNewUnit_list_of_mem_of_contains_eq_false [EquivBEq α] [LawfulHashable α] @@ -1460,7 +1460,7 @@ theorem getKeyD_insertManyIfNewUnit_list_of_contains_eq_false [EquivBEq α] [Law theorem getKeyD_insertManyIfNewUnit_list_of_mem_of_contains_eq_false [EquivBEq α] [LawfulHashable α] (h : m.WF) {l : List α} {k k' fallback : α} (k_beq : k == k') (distinct : l.Pairwise (fun a b => (a == b) = false)) - {mem : k ∈ l } : m.contains k = false → + (mem : k ∈ l) : m.contains k = false → getKeyD (insertManyIfNewUnit m l) k' fallback = k := by simp_to_raw using Raw₀.Const.getKeyD_insertManyIfNewUnit_list_of_mem_of_contains_eq_false @@ -1512,4 +1512,437 @@ theorem getValueD_insertManyIfNewUnit_list end Const end Raw +namespace Raw +variable [BEq α] [Hashable α] +@[simp] +theorem ofList_nil [EquivBEq α] [LawfulHashable α] : + ofList ([] : List ((a : α) × (β a))) = empty := by + simp only [ofList, WF.emptyc, insertMany_nil] + rfl + +@[simp] +theorem ofList_singleton {k : α} {v : β k} [EquivBEq α] [LawfulHashable α] : + ofList [⟨k, v⟩] = empty.insert k v := by + simp only [ofList, WF.emptyc, insertMany_list_singleton] + rfl + +theorem contains_ofList [EquivBEq α] [LawfulHashable α] + {l : List ((a : α) × β a)} {k : α} : + (ofList l).contains k = (l.map Sigma.fst).contains k := by + simp only [ofList, WF.emptyc, contains_insertMany_list, contains_emptyc, Bool.false_or] + +theorem get?_ofList_of_contains_eq_false [LawfulBEq α] + {l : List ((a : α) × β a)} {k : α} + (h : (l.map Sigma.fst).contains k = false) : + (ofList l).get? k = none := by + simp only [ofList] + rw [get?_insertMany_list_of_contains_eq_false WF.emptyc h] + apply get?_emptyc + +theorem get?_ofList_of_mem [LawfulBEq α] + {l : List ((a : α) × β a)} {k k' : α} (k_beq : k == k') {v : β k} + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) + (mem : ⟨k, v⟩ ∈ l) : + (ofList l).get? k' = some (cast (by congr; apply LawfulBEq.eq_of_beq k_beq) v) := by + simp only [ofList] + rw [get?_insertMany_list_of_mem WF.emptyc k_beq distinct mem] + +theorem get_ofList_of_mem [LawfulBEq α] + {l : List ((a : α) × β a)} {k k' : α} (k_beq : k == k') {v : β k} + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) + (mem : ⟨k, v⟩ ∈ l) + {h} : + (ofList l).get k' h = cast (by congr; apply LawfulBEq.eq_of_beq k_beq) v := by + simp only [ofList] + rw [get_insertMany_list_of_mem WF.emptyc k_beq distinct mem] + +theorem get!_ofList_of_contains_eq_false [LawfulBEq α] + {l : List ((a : α) × β a)} {k : α} [Inhabited (β k)] + (h : (l.map Sigma.fst).contains k = false) : + (ofList l).get! k = default := by + simp only [ofList] + rw [get!_insertMany_list_of_contains_eq_false WF.emptyc h] + apply get!_emptyc + +theorem get!_ofList_of_mem [LawfulBEq α] + {l : List ((a : α) × β a)} {k k' : α} (k_beq : k == k') {v : β k} [Inhabited (β k')] + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) + (mem : ⟨k, v⟩ ∈ l) : + (ofList l).get! k' = cast (by congr; apply LawfulBEq.eq_of_beq k_beq) v := by + simp only [ofList] + rw [get!_insertMany_list_of_mem WF.emptyc k_beq distinct mem] + +theorem getD_ofList_of_contains_eq_false [LawfulBEq α] + {l : List ((a : α) × β a)} {k : α} {fallback : β k} + (contains_eq_false : (l.map Sigma.fst).contains k = false) : + (ofList l).getD k fallback = fallback := by + simp only [ofList] + rw [getD_insertMany_list_of_contains_eq_false WF.emptyc contains_eq_false] + apply getD_emptyc + +theorem getD_ofList_of_mem [LawfulBEq α] + {l : List ((a : α) × β a)} {k k' : α} (k_beq : k == k') {v : β k} {fallback : β k'} + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) + (mem : ⟨k, v⟩ ∈ l) : + (ofList l).getD k' fallback = cast (by congr; apply LawfulBEq.eq_of_beq k_beq) v := by + simp only [ofList] + rw [getD_insertMany_list_of_mem WF.emptyc k_beq distinct mem] + +theorem getKey?_ofList_of_contains_eq_false [EquivBEq α] [LawfulHashable α] + {l : List ((a : α) × β a)} {k : α} + (h : (l.map Sigma.fst).contains k = false) : + (ofList l).getKey? k = none := by + simp only [ofList] + rw [getKey?_insertMany_list_of_contains_eq_false WF.emptyc h] + apply getKey?_emptyc + +theorem getKey?_ofList_of_mem [EquivBEq α] [LawfulHashable α] + {l : List ((a : α) × β a)} + {k k' : α} (k_beq : k == k') + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) + (mem : k ∈ l.map Sigma.fst) : + (ofList l).getKey? k' = some k := by + simp only [ofList] + rw [getKey?_insertMany_list_of_mem WF.emptyc k_beq distinct mem] + +theorem getKey_ofList_of_mem [EquivBEq α] [LawfulHashable α] + {l : List ((a : α) × β a)} + {k k' : α} (k_beq : k == k') + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) + (mem : k ∈ l.map Sigma.fst) + {h'} : + (ofList l).getKey k' h' = k := by + simp only [ofList] + rw [getKey_insertMany_list_of_mem WF.emptyc k_beq distinct mem] + +theorem getKey!_ofList_of_contains_eq_false [EquivBEq α] [LawfulHashable α] [Inhabited α] + {l : List ((a : α) × β a)} {k : α} + (h : (l.map Sigma.fst).contains k = false) : + (ofList l).getKey! k = default := by + simp only [ofList] + rw [getKey!_insertMany_list_of_contains_eq_false WF.emptyc h] + apply getKey!_emptyc + +theorem getKey!_ofList_of_mem [EquivBEq α] [LawfulHashable α] [Inhabited α] + {l : List ((a : α) × β a)} + {k k' : α} (k_beq : k == k') + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) + (mem : k ∈ l.map Sigma.fst) : + (ofList l).getKey! k' = k := by + simp only [ofList] + rw [getKey!_insertMany_list_of_mem WF.emptyc k_beq distinct mem] + +theorem getKeyD_ofList_of_contains_eq_false [EquivBEq α] [LawfulHashable α] + {l : List ((a : α) × β a)} {k fallback : α} + (h : (l.map Sigma.fst).contains k = false) : + (ofList l).getKeyD k fallback = fallback := by + simp only [ofList] + rw [getKeyD_insertMany_list_of_contains_eq_false WF.emptyc h] + apply getKeyD_emptyc + +theorem getKeyD_ofList_of_mem [EquivBEq α] [LawfulHashable α] + {l : List ((a : α) × β a)} + {k k' fallback : α} (k_beq : k == k') + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) + (mem : k ∈ l.map Sigma.fst) : + (ofList l).getKeyD k' fallback = k := by + simp only [ofList] + rw [getKeyD_insertMany_list_of_mem WF.emptyc k_beq distinct mem] + +theorem size_ofList [EquivBEq α] [LawfulHashable α] + {l : List ((a : α) × β a)} (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) : + (ofList l).size = l.length := by + simp only [ofList] + rw [size_insertMany_list WF.emptyc distinct] + · simp only [size_emptyc, Nat.zero_add] + · simp only [contains_emptyc, Bool.false_eq_true, false_implies, implies_true] + +theorem size_ofList_le [EquivBEq α] [LawfulHashable α] + {l : List ((a : α) × β a)} : + (ofList l).size ≤ l.length := by + simp only [ofList] + rw [← Nat.zero_add l.length, ← size_emptyc] + apply (size_insertMany_list_le WF.emptyc) + +theorem isEmpty_ofList [EquivBEq α] [LawfulHashable α] + {l : List ((a : α) × β a)} : + (ofList l).isEmpty = l.isEmpty := by + simp only [ofList] + rw [isEmpty_insertMany_list WF.emptyc] + simp + +namespace Const +variable {β : Type v} + +@[simp] +theorem ofList_nil [EquivBEq α] [LawfulHashable α] : + ofList ([] : List (α × β)) = empty := by + simp only [ofList, WF.emptyc, insertMany_nil] + rfl + +@[simp] +theorem ofList_singleton {k : α} {v : β} [EquivBEq α] [LawfulHashable α] : + ofList [⟨k, v⟩] = empty.insert k v := by + simp only [ofList, WF.emptyc, insertMany_list_singleton] + rfl + +theorem contains_ofList [EquivBEq α] [LawfulHashable α] + {l : List (α × β)} {k : α} : + (ofList l).contains k = (l.map Prod.fst).contains k := by + simp only [ofList, WF.emptyc, contains_insertMany_list, contains_emptyc, Bool.false_or] + +theorem get?_ofList_of_contains_eq_false [LawfulBEq α] + {l : List (α × β)} {k : α} + (h : (l.map Prod.fst).contains k = false) : + get? (ofList l) k = none := by + simp only [ofList] + rw [get?_insertMany_list_of_contains_eq_false WF.emptyc h] + apply get?_emptyc + +theorem get?_ofList_of_mem [LawfulBEq α] + {l : List (α × β)} {k k' : α} (k_beq : k == k') {v : β} + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) + (mem : ⟨k, v⟩ ∈ l) : + get? (ofList l) k' = some v := by + simp only [ofList] + rw [get?_insertMany_list_of_mem WF.emptyc k_beq distinct mem] + +theorem get_ofList_of_mem [LawfulBEq α] + {l : List (α × β)} {k k' : α} (k_beq : k == k') {v : β} + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) + (mem : ⟨k, v⟩ ∈ l) + {h} : + get (ofList l) k' h = v := by + simp only [ofList] + rw [get_insertMany_list_of_mem WF.emptyc k_beq distinct mem] + +theorem get!_ofList_of_contains_eq_false [LawfulBEq α] + {l : List (α × β)} {k : α} [Inhabited β] + (h : (l.map Prod.fst).contains k = false) : + get! (ofList l) k = default := by + simp only [ofList] + rw [get!_insertMany_list_of_contains_eq_false WF.emptyc h] + apply get!_emptyc + +theorem get!_ofList_of_mem [LawfulBEq α] + {l : List (α × β)} {k k' : α} (k_beq : k == k') {v : β} [Inhabited β] + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) + (mem : ⟨k, v⟩ ∈ l) : + get! (ofList l) k' = v := by + simp only [ofList] + rw [get!_insertMany_list_of_mem WF.emptyc k_beq distinct mem] + +theorem getD_ofList_of_contains_eq_false [LawfulBEq α] + {l : List (α × β)} {k : α} {fallback : β} + (contains_eq_false : (l.map Prod.fst).contains k = false) : + getD (ofList l) k fallback = fallback := by + simp only [ofList] + rw [getD_insertMany_list_of_contains_eq_false WF.emptyc contains_eq_false] + apply getD_emptyc + +theorem getD_ofList_of_mem [LawfulBEq α] + {l : List (α × β)} {k k' : α} (k_beq : k == k') {v : β} {fallback : β} + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) + (mem : ⟨k, v⟩ ∈ l) : + getD (ofList l) k' fallback = v := by + simp only [ofList] + rw [getD_insertMany_list_of_mem WF.emptyc k_beq distinct mem] + +theorem getKey?_ofList_of_contains_eq_false [EquivBEq α] [LawfulHashable α] + {l : List (α × β)} {k : α} + (h : (l.map Prod.fst).contains k = false) : + (ofList l).getKey? k = none := by + simp only [ofList] + rw [getKey?_insertMany_list_of_contains_eq_false WF.emptyc h] + apply getKey?_emptyc + +theorem getKey?_ofList_of_mem [EquivBEq α] [LawfulHashable α] + {l : List (α × β)} + {k k' : α} (k_beq : k == k') + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) + (mem : k ∈ l.map Prod.fst) : + (ofList l).getKey? k' = some k := by + simp only [ofList] + rw [getKey?_insertMany_list_of_mem WF.emptyc k_beq distinct mem] + +theorem getKey_ofList_of_mem [EquivBEq α] [LawfulHashable α] + {l : List (α × β)} + {k k' : α} (k_beq : k == k') + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) + (mem : k ∈ l.map Prod.fst) + {h'} : + (ofList l).getKey k' h' = k := by + simp only [ofList] + rw [getKey_insertMany_list_of_mem WF.emptyc k_beq distinct mem] + +theorem getKey!_ofList_of_contains_eq_false [EquivBEq α] [LawfulHashable α] [Inhabited α] + {l : List (α × β)} {k : α} + (h : (l.map Prod.fst).contains k = false) : + (ofList l).getKey! k = default := by + simp only [ofList] + rw [getKey!_insertMany_list_of_contains_eq_false WF.emptyc h] + apply getKey!_emptyc + +theorem getKey!_ofList_of_mem [EquivBEq α] [LawfulHashable α] [Inhabited α] + {l : List (α × β)} + {k k' : α} (k_beq : k == k') + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) + (mem : k ∈ l.map Prod.fst) : + (ofList l).getKey! k' = k := by + simp only [ofList] + rw [getKey!_insertMany_list_of_mem WF.emptyc k_beq distinct mem] + +theorem getKeyD_ofList_of_contains_eq_false [EquivBEq α] [LawfulHashable α] + {l : List (α × β)} {k fallback : α} + (h : (l.map Prod.fst).contains k = false) : + (ofList l).getKeyD k fallback = fallback := by + simp only [ofList] + rw [getKeyD_insertMany_list_of_contains_eq_false WF.emptyc h] + apply getKeyD_emptyc + +theorem getKeyD_ofList_of_mem [EquivBEq α] [LawfulHashable α] + {l : List (α × β)} + {k k' fallback : α} (k_beq : k == k') + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) + (mem : k ∈ l.map Prod.fst) : + (ofList l).getKeyD k' fallback = k := by + simp only [ofList] + rw [getKeyD_insertMany_list_of_mem WF.emptyc k_beq distinct mem] + +theorem size_ofList [EquivBEq α] [LawfulHashable α] + {l : List (α × β)} (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) : + (ofList l).size = l.length := by + simp only [ofList] + rw [size_insertMany_list WF.emptyc distinct] + · simp only [size_emptyc, Nat.zero_add] + · simp only [contains_emptyc, Bool.false_eq_true, false_implies, implies_true] + +theorem size_ofList_le [EquivBEq α] [LawfulHashable α] + {l : List (α × β)} : + (ofList l).size ≤ l.length := by + simp only [ofList] + rw [← Nat.zero_add l.length, ← size_emptyc] + apply (size_insertMany_list_le WF.emptyc) + +theorem isEmpty_ofList [EquivBEq α] [LawfulHashable α] + {l : List (α × β)} : + (ofList l).isEmpty = l.isEmpty := by + simp [ofList, isEmpty_insertMany_list WF.emptyc] + +@[simp] +theorem unitOfList_nil [EquivBEq α] [LawfulHashable α] : + unitOfList ([] : List α) = empty := by + simp only [unitOfList, insertManyIfNewUnit_nil WF.emptyc] + rfl + +@[simp] +theorem unitOfList_singleton [EquivBEq α] [LawfulHashable α] {k : α} : + unitOfList [k] = empty.insertIfNew k () := by + simp only [unitOfList, insertManyIfNewUnit_list_singleton WF.emptyc] + rfl + +theorem contains_unitOfList [EquivBEq α] [LawfulHashable α] + {l : List α} {k : α} : + (unitOfList l).contains k = l.contains k := by + simp [unitOfList, contains_insertManyIfNewUnit_list WF.emptyc] + +theorem getKey?_unitOfList_of_contains_eq_false [EquivBEq α] [LawfulHashable α] + {l : List α} {k : α} (h' : l.contains k = false) : + getKey? (unitOfList l) k = none := by + simp [unitOfList, getKey?_insertManyIfNewUnit_list_of_contains_eq_false WF.emptyc h'] + +theorem getKey?_unitOfList_of_mem [EquivBEq α] [LawfulHashable α] + {l : List α} {k k' : α} (k_beq : k == k') + (distinct : l.Pairwise (fun a b => (a == b) = false)) (mem : k ∈ l) : + getKey? (unitOfList l) k' = some k := by + simp [unitOfList, getKey?_insertManyIfNewUnit_list_of_mem_of_contains_eq_false WF.emptyc k_beq + distinct mem contains_emptyc] + +theorem getKey_unitOfList_of_mem [EquivBEq α] [LawfulHashable α] + {l : List α} + {k k' : α} (k_beq : k == k') + (distinct : l.Pairwise (fun a b => (a == b) = false)) + (mem : k ∈ l) {h'} : + getKey (unitOfList l) k' h' = k := by + simp [unitOfList, getKey_insertManyIfNewUnit_list_of_mem_of_contains_eq_false WF.emptyc k_beq + distinct mem contains_emptyc] + +theorem getKey!_unitOfList_of_contains_eq_false [EquivBEq α] [LawfulHashable α] + [Inhabited α] {l : List α} {k : α} + (h' : l.contains k = false) : + getKey! (unitOfList l) k = default := by + simp [unitOfList, getKey!_insertManyIfNewUnit_list_of_contains_eq_false WF.emptyc h'] + +theorem getKey!_unitOfList_of_mem [EquivBEq α] [LawfulHashable α] + [Inhabited α] {l : List α} {k k' : α} (k_beq : k == k') + (distinct : l.Pairwise (fun a b => (a == b) = false)) + (mem : k ∈ l) : + getKey! (unitOfList l) k' = k := by + simp [unitOfList, getKey!_insertManyIfNewUnit_list_of_mem_of_contains_eq_false WF.emptyc k_beq + distinct mem contains_emptyc] + +theorem getKeyD_unitOfList_of_contains_eq_false [EquivBEq α] [LawfulHashable α] + {l : List α} {k fallback : α} + (h' : l.contains k = false) : + getKeyD (unitOfList l) k fallback = fallback := by + simp [unitOfList, getKeyD_insertManyIfNewUnit_list_of_contains_eq_false WF.emptyc h'] + +theorem getKeyD_unitOfList_of_mem [EquivBEq α] [LawfulHashable α] + {l : List α} {k k' fallback : α} (k_beq : k == k') + (distinct : l.Pairwise (fun a b => (a == b) = false)) + {mem : k ∈ l } : + getKeyD (unitOfList l) k' fallback = k := by + simp [unitOfList, getKeyD_insertManyIfNewUnit_list_of_mem_of_contains_eq_false WF.emptyc k_beq + distinct mem contains_emptyc] + +theorem size_unitOfList [EquivBEq α] [LawfulHashable α] + {l : List α} + (distinct : l.Pairwise (fun a b => (a == b) = false)) : + (unitOfList l).size = l.length := by + simp only [unitOfList] + rw [size_insertManyIfNewUnit_list WF.emptyc distinct] + · simp + · simp + +theorem size_unitOfList_le [EquivBEq α] [LawfulHashable α] + {l : List α} : + (unitOfList l).size ≤ l.length := by + simp only [unitOfList] + apply Nat.le_trans (size_insertManyIfNewUnit_list_le WF.emptyc) + simp + + +theorem isEmpty_unitOfList [EquivBEq α] [LawfulHashable α] + {l : List α} : + (unitOfList l).isEmpty = l.isEmpty := by + simp only [unitOfList] + rw [isEmpty_insertManyIfNewUnit_list WF.emptyc] + simp + +theorem get?_unitOfList [EquivBEq α] [LawfulHashable α] + {l : List α} {k : α} : + get? (unitOfList l) k = + if l.contains k then some () else none := by + simp only [unitOfList] + rw [get?_insertManyIfNewUnit_list WF.emptyc] + simp + +theorem getValue_unitOfList + {l : List α} {k : α} {h} : + get (unitOfList l) k h = () := by + simp + +theorem getValue!_unitOfList + {l : List α} {k : α} : + get! (unitOfList l) k = () := by + simp + +theorem getValueD_unitOfList + {l : List α} {k : α} {fallback : Unit} : + getD (unitOfList l) k fallback = () := by + simp + +end Const +end Raw + end Std.DHashMap From cd58f46cf86753beb8dd4f7813ea9148dddf7448 Mon Sep 17 00:00:00 2001 From: jt0202 Date: Thu, 12 Dec 2024 23:01:55 +0100 Subject: [PATCH 53/83] Small name correction --- src/Std/Data/DHashMap/Internal/RawLemmas.lean | 6 +++--- src/Std/Data/DHashMap/RawLemmas.lean | 12 ++++++------ 2 files changed, 9 insertions(+), 9 deletions(-) diff --git a/src/Std/Data/DHashMap/Internal/RawLemmas.lean b/src/Std/Data/DHashMap/Internal/RawLemmas.lean index 6247c117aebd..1ce5dd2b5fc8 100644 --- a/src/Std/Data/DHashMap/Internal/RawLemmas.lean +++ b/src/Std/Data/DHashMap/Internal/RawLemmas.lean @@ -1287,17 +1287,17 @@ theorem get?_insertManyIfNewUnit_list [EquivBEq α] [LawfulHashable α] (h : m.1 if m.contains k ∨ l.contains k then some () else none := by simp_to_model using getValue?_insertListIfNewUnit -theorem getValue_insertManyIfNewUnit_list +theorem get_insertManyIfNewUnit_list {l : List α} {k : α} {h} : get (insertManyIfNewUnit m l).1 k h = () := by simp -theorem getValue!_insertManyIfNewUnit_list +theorem get!_insertManyIfNewUnit_list {l : List α} {k : α} : get! (insertManyIfNewUnit m l).1 k = () := by simp -theorem getValueD_insertManyIfNewUnit_list +theorem getD_insertManyIfNewUnit_list {l : List α} {k : α} {fallback : Unit} : getD (insertManyIfNewUnit m l).1 k fallback = () := by simp diff --git a/src/Std/Data/DHashMap/RawLemmas.lean b/src/Std/Data/DHashMap/RawLemmas.lean index 40cd2c5e91f5..fdceda2aed81 100644 --- a/src/Std/Data/DHashMap/RawLemmas.lean +++ b/src/Std/Data/DHashMap/RawLemmas.lean @@ -1494,17 +1494,17 @@ theorem get?_insertManyIfNewUnit_list [EquivBEq α] [LawfulHashable α] (h : m.W if m.contains k ∨ l.contains k then some () else none := by simp_to_raw using Raw₀.Const.get?_insertManyIfNewUnit_list -theorem getValue_insertManyIfNewUnit_list +theorem get_insertManyIfNewUnit_list {l : List α} {k : α} {h} : get (insertManyIfNewUnit m l) k h = () := by simp -theorem getValue!_insertManyIfNewUnit_list +theorem get!_insertManyIfNewUnit_list {l : List α} {k : α} : get! (insertManyIfNewUnit m l) k = () := by simp -theorem getValueD_insertManyIfNewUnit_list +theorem getD_insertManyIfNewUnit_list {l : List α} {k : α} {fallback : Unit} : getD (insertManyIfNewUnit m l) k fallback = () := by simp @@ -1927,17 +1927,17 @@ theorem get?_unitOfList [EquivBEq α] [LawfulHashable α] rw [get?_insertManyIfNewUnit_list WF.emptyc] simp -theorem getValue_unitOfList +theorem get_unitOfList {l : List α} {k : α} {h} : get (unitOfList l) k h = () := by simp -theorem getValue!_unitOfList +theorem get!_unitOfList {l : List α} {k : α} : get! (unitOfList l) k = () := by simp -theorem getValueD_unitOfList +theorem getD_unitOfList {l : List α} {k : α} {fallback : Unit} : getD (unitOfList l) k fallback = () := by simp From e1234aaad1406d32060ca7e5ca430c92096fcc17 Mon Sep 17 00:00:00 2001 From: jt0202 Date: Sun, 15 Dec 2024 20:17:31 +0100 Subject: [PATCH 54/83] Added cons for ofList --- src/Std/Data/DHashMap/RawLemmas.lean | 23 +++++++++++++++++++---- 1 file changed, 19 insertions(+), 4 deletions(-) diff --git a/src/Std/Data/DHashMap/RawLemmas.lean b/src/Std/Data/DHashMap/RawLemmas.lean index fdceda2aed81..3c4f8c8a5b95 100644 --- a/src/Std/Data/DHashMap/RawLemmas.lean +++ b/src/Std/Data/DHashMap/RawLemmas.lean @@ -1377,7 +1377,7 @@ theorem insertManyIfNewUnit_list_singleton [EquivBEq α] [LawfulHashable α] {k rw [Raw₀.Const.insertManyIfNewUnit_list_singleton] theorem insertManyIfNewUnit_cons [EquivBEq α] [LawfulHashable α] (h : m.WF) {l : List α} {k : α} : - (insertManyIfNewUnit m (k :: l)).1 = (insertManyIfNewUnit (m.insertIfNew k ()) l).1 := by + (insertManyIfNewUnit m (k :: l)) = (insertManyIfNewUnit (m.insertIfNew k ()) l) := by simp_to_raw rw [Raw₀.Const.insertManyIfNewUnit_cons] @@ -1526,6 +1526,11 @@ theorem ofList_singleton {k : α} {v : β k} [EquivBEq α] [LawfulHashable α] : simp only [ofList, WF.emptyc, insertMany_list_singleton] rfl +theorem ofList_cons [EquivBEq α] [LawfulHashable α] {hd : (a : α) × (β a)} {tl : List ((a : α) × (β a))} : + ofList (hd::tl) = ((∅ : Raw α β).insert hd.1 hd.2).insertMany tl := by + simp only [ofList] + rw [insertMany_cons WF.emptyc] + theorem contains_ofList [EquivBEq α] [LawfulHashable α] {l : List ((a : α) × β a)} {k : α} : (ofList l).contains k = (l.map Sigma.fst).contains k := by @@ -1686,6 +1691,11 @@ theorem ofList_singleton {k : α} {v : β} [EquivBEq α] [LawfulHashable α] : simp only [ofList, WF.emptyc, insertMany_list_singleton] rfl +theorem ofList_cons [EquivBEq α] [LawfulHashable α] {hd : α × β} {tl : List (α × β)} : + ofList (hd::tl) = insertMany ((∅ : Raw α (fun _ => β)).insert hd.1 hd.2) tl := by + simp only [ofList] + rw [insertMany_cons WF.emptyc] + theorem contains_ofList [EquivBEq α] [LawfulHashable α] {l : List (α × β)} {k : α} : (ofList l).contains k = (l.map Prod.fst).contains k := by @@ -1841,6 +1851,11 @@ theorem unitOfList_singleton [EquivBEq α] [LawfulHashable α] {k : α} : simp only [unitOfList, insertManyIfNewUnit_list_singleton WF.emptyc] rfl +theorem unitOfList_cons [EquivBEq α] [LawfulHashable α] {hd : α} {tl : List α} : + unitOfList (hd::tl) = insertManyIfNewUnit ((∅ : Raw α (fun _ => Unit)).insertIfNew hd ()) tl := by + simp only [unitOfList] + rw [insertManyIfNewUnit_cons WF.emptyc] + theorem contains_unitOfList [EquivBEq α] [LawfulHashable α] {l : List α} {k : α} : (unitOfList l).contains k = l.contains k := by @@ -1865,7 +1880,7 @@ theorem getKey_unitOfList_of_mem [EquivBEq α] [LawfulHashable α] (mem : k ∈ l) {h'} : getKey (unitOfList l) k' h' = k := by simp [unitOfList, getKey_insertManyIfNewUnit_list_of_mem_of_contains_eq_false WF.emptyc k_beq - distinct mem contains_emptyc] + distinct mem contains_emptyc] theorem getKey!_unitOfList_of_contains_eq_false [EquivBEq α] [LawfulHashable α] [Inhabited α] {l : List α} {k : α} @@ -1879,7 +1894,7 @@ theorem getKey!_unitOfList_of_mem [EquivBEq α] [LawfulHashable α] (mem : k ∈ l) : getKey! (unitOfList l) k' = k := by simp [unitOfList, getKey!_insertManyIfNewUnit_list_of_mem_of_contains_eq_false WF.emptyc k_beq - distinct mem contains_emptyc] + distinct mem contains_emptyc] theorem getKeyD_unitOfList_of_contains_eq_false [EquivBEq α] [LawfulHashable α] {l : List α} {k fallback : α} @@ -1893,7 +1908,7 @@ theorem getKeyD_unitOfList_of_mem [EquivBEq α] [LawfulHashable α] {mem : k ∈ l } : getKeyD (unitOfList l) k' fallback = k := by simp [unitOfList, getKeyD_insertManyIfNewUnit_list_of_mem_of_contains_eq_false WF.emptyc k_beq - distinct mem contains_emptyc] + distinct mem contains_emptyc] theorem size_unitOfList [EquivBEq α] [LawfulHashable α] {l : List α} From 62f5a53b8e946675f029a16241ca3c92359aaa2f Mon Sep 17 00:00:00 2001 From: jt0202 Date: Sat, 28 Dec 2024 16:24:51 +0100 Subject: [PATCH 55/83] Small style fixes --- src/Std/Data/DHashMap/Internal/RawLemmas.lean | 4 ++-- src/Std/Data/DHashMap/RawLemmas.lean | 6 ++++-- 2 files changed, 6 insertions(+), 4 deletions(-) diff --git a/src/Std/Data/DHashMap/Internal/RawLemmas.lean b/src/Std/Data/DHashMap/Internal/RawLemmas.lean index 1ce5dd2b5fc8..00849a9c3282 100644 --- a/src/Std/Data/DHashMap/Internal/RawLemmas.lean +++ b/src/Std/Data/DHashMap/Internal/RawLemmas.lean @@ -926,7 +926,7 @@ theorem getD_insertMany_list_of_mem [LawfulBEq α] (h : m.1.WF) theorem getKey?_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l : List ((a : α) × β a)} {k : α} - (h : (l.map Sigma.fst).contains k = false) : + (h' : (l.map Sigma.fst).contains k = false) : (m.insertMany l).1.getKey? k = m.getKey? k := by simp_to_model using getKey?_insertList_of_contains_eq_false @@ -984,7 +984,7 @@ theorem getKeyD_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] (h : m. simp_to_model using getKeyD_insertList_of_mem theorem size_insertMany_list [EquivBEq α] [LawfulHashable α] (h : m.1.WF) - {l : List ((a : α) × β a)} {distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)} : + {l : List ((a : α) × β a)} (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) : (∀ (a : α), m.contains a → (l.map Sigma.fst).contains a = false) → (m.insertMany l).1.1.size = m.1.size + l.length := by simp_to_model using length_insertList diff --git a/src/Std/Data/DHashMap/RawLemmas.lean b/src/Std/Data/DHashMap/RawLemmas.lean index 3c4f8c8a5b95..4a323de994ce 100644 --- a/src/Std/Data/DHashMap/RawLemmas.lean +++ b/src/Std/Data/DHashMap/RawLemmas.lean @@ -1217,12 +1217,14 @@ theorem insertMany_nil [EquivBEq α] [LawfulHashable α] (h : m.WF) : rw [Raw₀.Const.insertMany_nil] @[simp] -theorem insertMany_list_singleton {k : α} {v : β} [EquivBEq α] [LawfulHashable α] (h : m.WF) : +theorem insertMany_list_singleton [EquivBEq α] [LawfulHashable α] (h : m.WF) + {k : α} {v : β} : insertMany m [⟨k, v⟩] = m.insert k v := by simp_to_raw rw [Raw₀.Const.insertMany_list_singleton] -theorem insertMany_cons {l : List (α × β)} {k : α} {v : β} [EquivBEq α] [LawfulHashable α] (h : m.WF) : +theorem insertMany_cons [EquivBEq α] [LawfulHashable α] (h : m.WF) {l : List (α × β)} + {k : α} {v : β} : insertMany m (⟨k, v⟩ :: l) = insertMany (m.insert k v) l := by simp_to_raw rw [Raw₀.Const.insertMany_cons] From ee522d326f0396a938ebf0f0ff8eb0430e8ed5f3 Mon Sep 17 00:00:00 2001 From: jt0202 Date: Sat, 28 Dec 2024 18:02:00 +0100 Subject: [PATCH 56/83] Further small fixes --- src/Std/Data/DHashMap/Internal/RawLemmas.lean | 8 ++++---- 1 file changed, 4 insertions(+), 4 deletions(-) diff --git a/src/Std/Data/DHashMap/Internal/RawLemmas.lean b/src/Std/Data/DHashMap/Internal/RawLemmas.lean index 00849a9c3282..809577ae19b2 100644 --- a/src/Std/Data/DHashMap/Internal/RawLemmas.lean +++ b/src/Std/Data/DHashMap/Internal/RawLemmas.lean @@ -1046,7 +1046,7 @@ theorem getKey?_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] (h : m. theorem getKey_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l : List (α × β)} {k : α} - {h₁ : (l.map Prod.fst).contains k = false} + (h₁ : (l.map Prod.fst).contains k = false) {h'} : (insertMany m l).1.getKey k h' = m.getKey k (contains_of_contains_insertMany_list _ h h' h₁) := by @@ -1091,7 +1091,7 @@ theorem getKeyD_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] (h : m. theorem size_insertMany_list [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l : List (α × β)} - {distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)} : + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) : (∀ (a : α), m.contains a → (l.map Prod.fst).contains a = false) → (insertMany m l).1.1.size = m.1.size + l.length := by simp_to_model using length_insertListConst @@ -1220,7 +1220,7 @@ theorem getKey_insertManyIfNewUnit_list_of_mem_of_contains_eq_false [EquivBEq α theorem getKey_insertManyIfNewUnit_list_mem_of_contains_of_contains [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l : List α} {k : α} (distinct : l.Pairwise (fun a b => (a == b) = false)) - (h' : l.contains k) {mem' : m.contains k} {h'} : + (h₁ : l.contains k) {mem' : m.contains k} {h'} : getKey (insertManyIfNewUnit m l).1 k h' = getKey m k mem' := by simp_to_model using getKey_insertListIfNewUnit_of_contains_of_contains @@ -1253,7 +1253,7 @@ theorem getKeyD_insertManyIfNewUnit_list_of_contains_eq_false [EquivBEq α] [Law theorem getKeyD_insertManyIfNewUnit_list_of_mem_of_contains_eq_false [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l : List α} {k k' fallback : α} (k_beq : k == k') (distinct : l.Pairwise (fun a b => (a == b) = false)) - {mem : k ∈ l } : m.contains k = false → + (mem : k ∈ l ) : m.contains k = false → getKeyD (insertManyIfNewUnit m l).1 k' fallback = k := by simp_to_model using getKeyD_insertListIfNewUnit_of_mem_of_contains_eq_false From 57f188c1cda2c41f1878e39b1d1997f80212e9cf Mon Sep 17 00:00:00 2001 From: jt0202 Date: Sat, 28 Dec 2024 18:02:40 +0100 Subject: [PATCH 57/83] Experimental: Pulled results up to DHashMap --- src/Std/Data/DHashMap/Lemmas.lean | 917 +++++++++++++++++++++++++++++- 1 file changed, 915 insertions(+), 2 deletions(-) diff --git a/src/Std/Data/DHashMap/Lemmas.lean b/src/Std/Data/DHashMap/Lemmas.lean index e6016179028c..8c3f37a3b1c2 100644 --- a/src/Std/Data/DHashMap/Lemmas.lean +++ b/src/Std/Data/DHashMap/Lemmas.lean @@ -960,12 +960,925 @@ theorem contains_keys [EquivBEq α] [LawfulHashable α] {k : α} : @[simp] theorem mem_keys [LawfulBEq α] [LawfulHashable α] {k : α} : - k ∈ m.keys ↔ k ∈ m := by + k ∈ m.keys ↔ k ∈ m := by rw [mem_iff_contains] exact Raw₀.mem_keys ⟨m.1, _⟩ m.2 theorem distinct_keys [EquivBEq α] [LawfulHashable α] : - m.keys.Pairwise (fun a b => (a == b) = false) := + m.keys.Pairwise (fun a b => (a == b) = false) := Raw₀.distinct_keys ⟨m.1, m.2.size_buckets_pos⟩ m.2 +@[simp] +theorem insertMany_nil : + m.insertMany [] = m := by + apply Subtype.eq + congr + +@[simp] +theorem insertMany_list_singleton {k : α} {v : β k} : + m.insertMany [⟨k, v⟩] = m.insert k v := by + apply Subtype.eq + congr + +theorem insertMany_cons {l : List ((a : α) × β a)} {k : α} {v : β k} : + (m.insertMany (⟨k, v⟩ :: l)) = ((m.insert k v).insertMany l) := by + apply Subtype.eq + simp [insertMany] + rw [Raw₀.insertMany_cons ⟨m.1, _⟩] + rfl + +theorem contains_insertMany_list [EquivBEq α] [LawfulHashable α] + {l : List ((a : α) × β a)} {k : α} : + (m.insertMany l).contains k = (m.contains k || (l.map Sigma.fst).contains k) := + Raw₀.contains_insertMany_list ⟨m.1, _⟩ m.2 + +theorem contains_of_contains_insertMany_list [EquivBEq α] [LawfulHashable α] + {l : List ((a : α) × β a)} {k : α} : + (m.insertMany l).contains k → (l.map Sigma.fst).contains k = false → m.contains k := + Raw₀.contains_of_contains_insertMany_list ⟨m.1, _⟩ m.2 + +theorem get?_insertMany_list_of_contains_eq_false [LawfulBEq α] + {l : List ((a : α) × β a)} {k : α} + (h' : (l.map Sigma.fst).contains k = false) : + (m.insertMany l).get? k = m.get? k := + Raw₀.get?_insertMany_list_of_contains_eq_false ⟨m.1, _⟩ m.2 h' + +theorem get?_insertMany_list_of_mem [LawfulBEq α] + {l : List ((a : α) × β a)} {k k' : α} (k_beq : k == k') {v : β k} + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) + (mem : ⟨k, v⟩ ∈ l) : + (m.insertMany l).get? k' = some (cast (by congr; apply LawfulBEq.eq_of_beq k_beq) v) := + Raw₀.get?_insertMany_list_of_mem ⟨m.1, _⟩ m.2 k_beq distinct mem + +theorem get_insertMany_list_of_contains_eq_false [LawfulBEq α] + {l : List ((a : α) × β a)} {k : α} + (contains : (l.map Sigma.fst).contains k = false) + {h'} : + (m.insertMany l).get k h' = + m.get k (contains_of_contains_insertMany_list h' contains) := + Raw₀.get_insertMany_list_of_contains_eq_false ⟨m.1, _⟩ m.2 contains + +theorem get_insertMany_list_of_mem [LawfulBEq α] + {l : List ((a : α) × β a)} {k k' : α} (k_beq : k == k') {v : β k} + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) + (mem : ⟨k, v⟩ ∈ l) + {h'} : + (m.insertMany l).get k' h' = cast (by congr; apply LawfulBEq.eq_of_beq k_beq) v := + Raw₀.get_insertMany_list_of_mem ⟨m.1, _⟩ m.2 k_beq distinct mem + +theorem get!_insertMany_list_of_contains_eq_false [LawfulBEq α] + {l : List ((a : α) × β a)} {k : α} [Inhabited (β k)] + (h' : (l.map Sigma.fst).contains k = false) : + (m.insertMany l).get! k = m.get! k := + Raw₀.get!_insertMany_list_of_contains_eq_false ⟨m.1, _⟩ m.2 h' + +theorem get!_insertMany_list_of_mem [LawfulBEq α] + {l : List ((a : α) × β a)} {k k' : α} (k_beq : k == k') {v : β k} [Inhabited (β k')] + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) + (mem : ⟨k, v⟩ ∈ l) : + (m.insertMany l).get! k' = cast (by congr; apply LawfulBEq.eq_of_beq k_beq) v := + Raw₀.get!_insertMany_list_of_mem ⟨m.1, _⟩ m.2 k_beq distinct mem + +theorem getD_insertMany_list_of_contains_eq_false [LawfulBEq α] + {l : List ((a : α) × β a)} {k : α} {fallback : β k} + (not_mem : (l.map Sigma.fst).contains k = false) : + (m.insertMany l).getD k fallback = m.getD k fallback := + Raw₀.getD_insertMany_list_of_contains_eq_false ⟨m.1, _⟩ m.2 not_mem + +theorem getD_insertMany_list_of_mem [LawfulBEq α] + {l : List ((a : α) × β a)} {k k' : α} (k_beq : k == k') {v : β k} {fallback : β k'} + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) + (mem : ⟨k, v⟩ ∈ l) : + (m.insertMany l).getD k' fallback = cast (by congr; apply LawfulBEq.eq_of_beq k_beq) v := + Raw₀.getD_insertMany_list_of_mem ⟨m.1, _⟩ m.2 k_beq distinct mem + +theorem getKey?_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] + {l : List ((a : α) × β a)} {k : α} + (h' : (l.map Sigma.fst).contains k = false) : + (m.insertMany l).getKey? k = m.getKey? k := + Raw₀.getKey?_insertMany_list_of_contains_eq_false ⟨m.1, _⟩ m.2 h' + +theorem getKey?_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] + {l : List ((a : α) × β a)} + {k k' : α} (k_beq : k == k') + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) + (mem : k ∈ l.map Sigma.fst) : + (m.insertMany l).getKey? k' = some k := + Raw₀.getKey?_insertMany_list_of_mem ⟨m.1, _⟩ m.2 k_beq distinct mem + +theorem getKey_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] + {l : List ((a : α) × β a)} {k : α} + (h₁ : (l.map Sigma.fst).contains k = false) + {h'} : + (m.insertMany l).getKey k h' = + m.getKey k (contains_of_contains_insertMany_list h' h₁) := + Raw₀.getKey_insertMany_list_of_contains_eq_false ⟨m.1, _⟩ m.2 h₁ + +theorem getKey_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] + {l : List ((a : α) × β a)} + {k k' : α} (k_beq : k == k') + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) + (mem : k ∈ l.map Sigma.fst) + {h'} : + (m.insertMany l).getKey k' h' = k := + Raw₀.getKey_insertMany_list_of_mem ⟨m.1, _⟩ m.2 k_beq distinct mem + +theorem getKey!_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] [Inhabited α] + {l : List ((a : α) × β a)} {k : α} + (h' : (l.map Sigma.fst).contains k = false) : + (m.insertMany l).getKey! k = m.getKey! k := + Raw₀.getKey!_insertMany_list_of_contains_eq_false ⟨m.1, _⟩ m.2 h' + +theorem getKey!_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] [Inhabited α] + {l : List ((a : α) × β a)} + {k k' : α} (k_beq : k == k') + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) + (mem : k ∈ l.map Sigma.fst) : + (m.insertMany l).getKey! k' = k := + Raw₀.getKey!_insertMany_list_of_mem ⟨m.1, _⟩ m.2 k_beq distinct mem + +theorem getKeyD_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] + {l : List ((a : α) × β a)} {k fallback : α} + (h' : (l.map Sigma.fst).contains k = false) : + (m.insertMany l).getKeyD k fallback = m.getKeyD k fallback := + Raw₀.getKeyD_insertMany_list_of_contains_eq_false ⟨m.1, _⟩ m.2 h' + +theorem getKeyD_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] + {l : List ((a : α) × β a)} + {k k' fallback : α} (k_beq : k == k') + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) + (mem : k ∈ l.map Sigma.fst) : + (m.insertMany l).getKeyD k' fallback = k := + Raw₀.getKeyD_insertMany_list_of_mem ⟨m.1, _⟩ m.2 k_beq distinct mem + +theorem size_insertMany_list [EquivBEq α] [LawfulHashable α] + {l : List ((a : α) × β a)} (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) : + (∀ (a : α), m.contains a → (l.map Sigma.fst).contains a = false) → + (m.insertMany l).size = m.size + l.length := + Raw₀.size_insertMany_list ⟨m.1, _⟩ m.2 distinct + +theorem size_insertMany_list_le [EquivBEq α] [LawfulHashable α] + {l : List ((a : α) × β a)} : + (m.insertMany l).size ≤ m.size + l.length := + Raw₀.size_insertMany_list_le ⟨m.1, _⟩ m.2 + +theorem isEmpty_insertMany_list [EquivBEq α] [LawfulHashable α] + {l : List ((a : α) × β a)} : + (m.insertMany l).isEmpty = (m.isEmpty && l.isEmpty) := + Raw₀.isEmpty_insertMany_list ⟨m.1, _⟩ m.2 + +namespace Const +variable {β : Type v} (m : DHashMap α (fun _ => β)) + +@[simp] +theorem insertMany_nil : + insertMany m [] = m := by + apply Subtype.eq + congr + +@[simp] +theorem insertMany_list_singleton {k : α} {v : β} : + insertMany m [⟨k, v⟩] = m.insert k v := by + apply Subtype.eq + congr + +theorem insertMany_cons {l : List (α × β)} {k : α} {v : β} : + (insertMany m (⟨k, v⟩ :: l)) = (insertMany (m.insert k v) l) := by + apply Subtype.eq + simp [insertMany] + rw [Raw₀.Const.insertMany_cons ⟨m.1, _⟩] + rfl +theorem contains_insertMany_list [EquivBEq α] [LawfulHashable α] + {l : List (α × β)} {k : α} : + (Const.insertMany m l).contains k = (m.contains k || (l.map Prod.fst).contains k) := + Raw₀.Const.contains_insertMany_list ⟨m.1, _⟩ m.2 + +theorem contains_of_contains_insertMany_list [EquivBEq α] [LawfulHashable α] + {l : List ( α × β )} {k : α} : + (insertMany m l).contains k → (l.map Prod.fst).contains k = false → m.contains k := + Raw₀.Const.contains_of_contains_insertMany_list ⟨m.1, _⟩ m.2 + +theorem getKey?_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] + {l : List (α × β)} {k : α} + (h' : (l.map Prod.fst).contains k = false) : + (insertMany m l).getKey? k = m.getKey? k := + Raw₀.Const.getKey?_insertMany_list_of_contains_eq_false ⟨m.1, _⟩ m.2 h' + +theorem getKey?_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] + {l : List (α × β)} + {k k' : α} (k_beq : k == k') + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) + (mem : k ∈ l.map Prod.fst) : + (insertMany m l).getKey? k' = some k := + Raw₀.Const.getKey?_insertMany_list_of_mem ⟨m.1, _⟩ m.2 k_beq distinct mem + +theorem getKey_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] + {l : List (α × β)} {k : α} + (h₁ : (l.map Prod.fst).contains k = false) + {h'} : + (insertMany m l).getKey k h' = + m.getKey k (contains_of_contains_insertMany_list _ h' h₁) := + Raw₀.Const.getKey_insertMany_list_of_contains_eq_false ⟨m.1, _⟩ m.2 h₁ + +theorem getKey_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] + {l : List (α × β)} + {k k' : α} (k_beq : k == k') + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) + (mem : k ∈ l.map Prod.fst) + {h'} : + (insertMany m l).getKey k' h' = k := + Raw₀.Const.getKey_insertMany_list_of_mem ⟨m.1, _⟩ m.2 k_beq distinct mem + +theorem getKey!_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] [Inhabited α] + {l : List (α × β)} {k : α} + (h' : (l.map Prod.fst).contains k = false) : + (insertMany m l).getKey! k = m.getKey! k := + Raw₀.Const.getKey!_insertMany_list_of_contains_eq_false ⟨m.1, _⟩ m.2 h' + +theorem getKey!_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] [Inhabited α] + {l : List (α × β)} + {k k' : α} (k_beq : k == k') + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) + (mem : k ∈ l.map Prod.fst) : + (insertMany m l).getKey! k' = k := + Raw₀.Const.getKey!_insertMany_list_of_mem ⟨m.1, _⟩ m.2 k_beq distinct mem + +theorem getKeyD_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] + {l : List (α × β)} {k fallback : α} + (h' : (l.map Prod.fst).contains k = false) : + (insertMany m l).getKeyD k fallback = m.getKeyD k fallback := + Raw₀.Const.getKeyD_insertMany_list_of_contains_eq_false ⟨m.1, _⟩ m.2 h' + +theorem getKeyD_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] + {l : List (α × β)} + {k k' fallback : α} (k_beq : k == k') + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) + (mem : k ∈ l.map Prod.fst) : + (insertMany m l).getKeyD k' fallback = k := + Raw₀.Const.getKeyD_insertMany_list_of_mem ⟨m.1, _⟩ m.2 k_beq distinct mem + +theorem size_insertMany_list [EquivBEq α] [LawfulHashable α] + {l : List (α × β)} + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) : + (∀ (a : α), m.contains a → (l.map Prod.fst).contains a = false) → + (insertMany m l).size = m.size + l.length := + Raw₀.Const.size_insertMany_list ⟨m.1, _⟩ m.2 distinct + +theorem size_insertMany_list_le [EquivBEq α] [LawfulHashable α] + {l : List (α × β)} : + (insertMany m l).size ≤ m.size + l.length := + Raw₀.Const.size_insertMany_list_le ⟨m.1, _⟩ m.2 + +theorem isEmpty_insertMany_list [EquivBEq α] [LawfulHashable α] + {l : List (α × β)} : + (insertMany m l).isEmpty = (m.isEmpty && l.isEmpty) := + Raw₀.Const.isEmpty_insertMany_list ⟨m.1, _⟩ m.2 + +theorem get?_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] + {l : List (α × β)} {k : α} + (h' : (l.map Prod.fst).contains k = false) : + get? (insertMany m l) k = get? m k := + Raw₀.Const.get?_insertMany_list_of_contains_eq_false ⟨m.1, _⟩ m.2 h' + +theorem get?_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] + {l : List (α × β)} {k k' : α} (k_beq : k == k') {v : β} (mem : ⟨k, v⟩ ∈ l) + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) : + get? (insertMany m l) k' = v := + Raw₀.Const.get?_insertMany_list_of_mem ⟨m.1, _⟩ m.2 k_beq mem distinct + +theorem get_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] + {l : List (α × β)} {k : α} + (h₁ : (l.map Prod.fst).contains k = false) + {h'} : + get (insertMany m l) k h' = get m k (contains_of_contains_insertMany_list _ h' h₁) := + Raw₀.Const.get_insertMany_list_of_contains_eq_false ⟨m.1, _⟩ m.2 h₁ + +theorem get_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] + {l : List (α × β)} {k k' : α} (k_beq : k == k') {v : β} (mem : ⟨k, v⟩ ∈ l) + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) {h'} : + get (insertMany m l) k' h' = v := + Raw₀.Const.get_insertMany_list_of_mem ⟨m.1, _⟩ m.2 k_beq mem distinct + +theorem get!_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] + [Inhabited β] {l : List (α × β)} {k : α} + (h' : (l.map Prod.fst).contains k = false) : + get! (insertMany m l) k = get! m k := + Raw₀.Const.get!_insertMany_list_of_contains_eq_false ⟨m.1, _⟩ m.2 h' + +theorem get!_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] [Inhabited β] + {l : List (α × β)} {k k' : α} (k_beq : k == k') {v : β} (mem : ⟨k, v⟩ ∈ l) + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) : + get! (insertMany m l) k' = v := + Raw₀.Const.get!_insertMany_list_of_mem ⟨m.1, _⟩ m.2 k_beq mem distinct + +theorem getD_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] + {l : List (α × β)} {k : α} {fallback : β} + (h' : (l.map Prod.fst).contains k = false) : + getD (insertMany m l) k fallback = getD m k fallback := + Raw₀.Const.getD_insertMany_list_of_contains_eq_false ⟨m.1, _⟩ m.2 h' + +theorem getD_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] + {l : List (α × β)} {k k' : α} (k_beq : k == k') {v fallback : β} (mem : ⟨k, v⟩ ∈ l) + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) : + getD (insertMany m l) k' fallback = v := + Raw₀.Const.getD_insertMany_list_of_mem ⟨m.1, _⟩ m.2 k_beq mem distinct + +variable (m : DHashMap α (fun _ => Unit)) + +@[simp] +theorem insertManyIfNewUnit_nil : + insertManyIfNewUnit m [] = m := by + apply Subtype.eq + congr + +@[simp] +theorem insertManyIfNewUnit_list_singleton [EquivBEq α] [LawfulHashable α] {k : α} : + insertManyIfNewUnit m [k] = m.insertIfNew k () := by + apply Subtype.eq + congr + +theorem insertManyIfNewUnit_cons {l : List α} {k : α} : + (insertManyIfNewUnit m (k :: l)) = (insertManyIfNewUnit (m.insertIfNew k ()) l) := by + apply Subtype.eq + simp [insertManyIfNewUnit] + rw [Raw₀.Const.insertManyIfNewUnit_cons ⟨m.1, _⟩] + rfl + +theorem contains_insertManyIfNewUnit_list [EquivBEq α] [LawfulHashable α] + {l : List α} {k : α} : + (insertManyIfNewUnit m l).contains k = (m.contains k || l.contains k) := + Raw₀.Const.contains_insertManyIfNewUnit_list ⟨m.1, _⟩ m.2 + +theorem contains_of_contains_insertManyIfNewUnit_list [EquivBEq α] [LawfulHashable α] + {l : List α} {k : α} (h₂ : l.contains k = false) : + (insertManyIfNewUnit m l).contains k → m.contains k := + Raw₀.Const.contains_of_contains_insertManyIfNewUnit_list ⟨m.1, _⟩ m.2 h₂ + +theorem getKey?_insertManyIfNewUnit_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] + {l : List α} {k : α} (h' : l.contains k = false) : + getKey? (insertManyIfNewUnit m l) k = getKey? m k := + Raw₀.Const.getKey?_insertManyIfNewUnit_list_of_contains_eq_false ⟨m.1, _⟩ m.2 h' + +theorem getKey?_insertManyIfNewUnit_list_of_mem_of_contains_eq_false [EquivBEq α] [LawfulHashable α] + {l : List α} {k k' : α} (k_beq : k == k') + (distinct : l.Pairwise (fun a b => (a == b) = false)) (mem : k ∈ l) : + m.contains k = false → getKey? (insertManyIfNewUnit m l) k' = some k := + Raw₀.Const.getKey?_insertManyIfNewUnit_list_of_mem_of_contains_eq_false ⟨m.1, _⟩ m.2 k_beq distinct mem + +theorem getKey?_insertManyIfNewUnit_list_of_contains_of_contains [EquivBEq α] [LawfulHashable α] + {l : List α} {k : α} + (distinct : l.Pairwise (fun a b => (a == b) = false)) + (h' : l.contains k) : + m.contains k → getKey? (insertManyIfNewUnit m l) k = getKey? m k := + Raw₀.Const.getKey?_insertManyIfNewUnit_list_of_contains_of_contains ⟨m.1, _⟩ m.2 distinct h' + +theorem getKey_insertManyIfNewUnit_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] + {l : List α} {k : α} (h₁ : l.contains k = false) {h'} : + getKey (insertManyIfNewUnit m l) k h' = + getKey m k (contains_of_contains_insertManyIfNewUnit_list m h₁ h') := + Raw₀.Const.getKey_insertManyIfNewUnit_list_of_contains_eq_false ⟨m.1, _⟩ m.2 h₁ + +theorem getKey_insertManyIfNewUnit_list_of_mem_of_contains_eq_false [EquivBEq α] [LawfulHashable α] + {l : List α} + {k k' : α} (k_beq : k == k') + (distinct : l.Pairwise (fun a b => (a == b) = false)) + (mem : k ∈ l) {h'} : + m.contains k = false → + getKey (insertManyIfNewUnit m l) k' h' = k := + Raw₀.Const.getKey_insertManyIfNewUnit_list_of_mem_of_contains_eq_false ⟨m.1, _⟩ m.2 k_beq distinct mem + +theorem getKey_insertManyIfNewUnit_list_mem_of_contains_of_contains [EquivBEq α] [LawfulHashable α] + {l : List α} {k : α} + (distinct : l.Pairwise (fun a b => (a == b) = false)) + (h₁ : l.contains k) {mem' : m.contains k} {h'} : + getKey (insertManyIfNewUnit m l) k h' = getKey m k mem' := + Raw₀.Const.getKey_insertManyIfNewUnit_list_mem_of_contains_of_contains ⟨m.1, _⟩ m.2 distinct h₁ + +theorem getKey!_insertManyIfNewUnit_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] + [Inhabited α] {l : List α} {k : α} + (h' : l.contains k = false) : + getKey! (insertManyIfNewUnit m l) k = getKey! m k := + Raw₀.Const.getKey!_insertManyIfNewUnit_list_of_contains_eq_false ⟨m.1, _⟩ m.2 h' + +theorem getKey!_insertManyIfNewUnit_list_of_mem_of_contains_eq_false [EquivBEq α] [LawfulHashable α] + [Inhabited α] {l : List α} {k k' : α} (k_beq : k == k') + (distinct : l.Pairwise (fun a b => (a == b) = false)) + (mem : k ∈ l) : contains m k = false → + getKey! (insertManyIfNewUnit m l) k' = k := + Raw₀.Const.getKey!_insertManyIfNewUnit_list_of_mem_of_contains_eq_false ⟨m.1, _⟩ m.2 k_beq distinct mem + +theorem getKey!_insertManyIfNewUnit_list_mem_of_contains_of_contains [EquivBEq α] [LawfulHashable α] + [Inhabited α] {l : List α} {k : α} + (distinct : l.Pairwise (fun a b => (a == b) = false)) + (h' : l.contains k) : m.contains k → + getKey! (insertManyIfNewUnit m l) k = getKey! m k := + Raw₀.Const.getKey!_insertManyIfNewUnit_list_mem_of_contains_of_contains ⟨m.1, _⟩ m.2 distinct h' + +theorem getKeyD_insertManyIfNewUnit_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] + {l : List α} {k fallback : α} + (h' : l.contains k = false) : + getKeyD (insertManyIfNewUnit m l) k fallback = getKeyD m k fallback := + Raw₀.Const.getKeyD_insertManyIfNewUnit_list_of_contains_eq_false ⟨m.1, _⟩ m.2 h' + +theorem getKeyD_insertManyIfNewUnit_list_of_mem_of_contains_eq_false [EquivBEq α] [LawfulHashable α] + {l : List α} {k k' fallback : α} (k_beq : k == k') + (distinct : l.Pairwise (fun a b => (a == b) = false)) + (mem : k ∈ l ) : m.contains k = false → + getKeyD (insertManyIfNewUnit m l) k' fallback = k := + Raw₀.Const.getKeyD_insertManyIfNewUnit_list_of_mem_of_contains_eq_false ⟨m.1, _⟩ m.2 k_beq distinct mem + +theorem getKeyD_insertManyIfNewUnit_list_of_contains_of_contains [EquivBEq α] [LawfulHashable α] + {l : List α} {k fallback : α} + (distinct : l.Pairwise (fun a b => (a == b) = false)) + (h' : l.contains k) : m.contains k → + getKeyD (insertManyIfNewUnit m l) k fallback = getKeyD m k fallback := + Raw₀.Const.getKeyD_insertManyIfNewUnit_list_of_contains_of_contains ⟨m.1, _⟩ m.2 distinct h' + +theorem size_insertManyIfNewUnit_list [EquivBEq α] [LawfulHashable α] + {l : List α} + (distinct : l.Pairwise (fun a b => (a == b) = false)) : + (∀ (a : α), m.contains a → l.contains a = false) → + (insertManyIfNewUnit m l).size = m.size + l.length := + Raw₀.Const.size_insertManyIfNewUnit_list ⟨m.1, _⟩ m.2 distinct + +theorem size_insertManyIfNewUnit_list_le [EquivBEq α] [LawfulHashable α] + {l : List α} : + (insertManyIfNewUnit m l).size ≤ m.size + l.length := + Raw₀.Const.size_insertManyIfNewUnit_list_le ⟨m.1, _⟩ m.2 + +theorem isEmpty_insertManyIfNewUnit_list [EquivBEq α] [LawfulHashable α] + {l : List α} : + (insertManyIfNewUnit m l).isEmpty = (m.isEmpty && l.isEmpty) := + Raw₀.Const.isEmpty_insertManyIfNewUnit_list ⟨m.1, _⟩ m.2 + +theorem get?_insertManyIfNewUnit_list [EquivBEq α] [LawfulHashable α] + {l : List α} {k : α} : + get? (insertManyIfNewUnit m l) k = + if m.contains k ∨ l.contains k then some () else none := + Raw₀.Const.get?_insertManyIfNewUnit_list ⟨m.1, _⟩ m.2 + +theorem get_insertManyIfNewUnit_list + {l : List α} {k : α} {h} : + get (insertManyIfNewUnit m l) k h = () := + Raw₀.Const.get_insertManyIfNewUnit_list ⟨m.1, _⟩ + +theorem get!_insertManyIfNewUnit_list + {l : List α} {k : α} : + get! (insertManyIfNewUnit m l) k = () := + Raw₀.Const.get!_insertManyIfNewUnit_list ⟨m.1, _⟩ + +theorem getD_insertManyIfNewUnit_list + {l : List α} {k : α} {fallback : Unit} : + getD (insertManyIfNewUnit m l) k fallback = () := by + simp + +end Const +end Std.DHashMap + +namespace Std.DHashMap +variable [BEq α] [Hashable α] +@[simp] +theorem ofList_nil [EquivBEq α] [LawfulHashable α] : + ofList ([] : List ((a : α) × (β a))) = empty := by + simp only [ofList, insertMany_nil] + rfl + +@[simp] +theorem ofList_singleton {k : α} {v : β k} [EquivBEq α] [LawfulHashable α] : + ofList [⟨k, v⟩] = empty.insert k v := by + simp only [ofList, insertMany_list_singleton] + rfl + +theorem ofList_cons [EquivBEq α] [LawfulHashable α] {hd : (a : α) × (β a)} {tl : List ((a : α) × (β a))} : + ofList (hd::tl) = ((∅ : DHashMap α β).insert hd.1 hd.2).insertMany tl := by + simp only [ofList] + rw [insertMany_cons] + +theorem contains_ofList [EquivBEq α] [LawfulHashable α] + {l : List ((a : α) × β a)} {k : α} : + (ofList l).contains k = (l.map Sigma.fst).contains k := by + simp only [ofList, contains_insertMany_list, contains_emptyc, Bool.false_or] + +theorem get?_ofList_of_contains_eq_false [LawfulBEq α] + {l : List ((a : α) × β a)} {k : α} + (h : (l.map Sigma.fst).contains k = false) : + (ofList l).get? k = none := by + simp only [ofList] + rw [get?_insertMany_list_of_contains_eq_false h] + apply get?_emptyc + +theorem get?_ofList_of_mem [LawfulBEq α] + {l : List ((a : α) × β a)} {k k' : α} (k_beq : k == k') {v : β k} + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) + (mem : ⟨k, v⟩ ∈ l) : + (ofList l).get? k' = some (cast (by congr; apply LawfulBEq.eq_of_beq k_beq) v) := by + simp only [ofList] + rw [get?_insertMany_list_of_mem k_beq distinct mem] + +theorem get_ofList_of_mem [LawfulBEq α] + {l : List ((a : α) × β a)} {k k' : α} (k_beq : k == k') {v : β k} + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) + (mem : ⟨k, v⟩ ∈ l) + {h} : + (ofList l).get k' h = cast (by congr; apply LawfulBEq.eq_of_beq k_beq) v := by + simp only [ofList] + rw [get_insertMany_list_of_mem k_beq distinct mem] + +theorem get!_ofList_of_contains_eq_false [LawfulBEq α] + {l : List ((a : α) × β a)} {k : α} [Inhabited (β k)] + (h : (l.map Sigma.fst).contains k = false) : + (ofList l).get! k = default := by + simp only [ofList] + rw [get!_insertMany_list_of_contains_eq_false h] + apply get!_emptyc + +theorem get!_ofList_of_mem [LawfulBEq α] + {l : List ((a : α) × β a)} {k k' : α} (k_beq : k == k') {v : β k} [Inhabited (β k')] + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) + (mem : ⟨k, v⟩ ∈ l) : + (ofList l).get! k' = cast (by congr; apply LawfulBEq.eq_of_beq k_beq) v := by + simp only [ofList] + rw [get!_insertMany_list_of_mem k_beq distinct mem] + +theorem getD_ofList_of_contains_eq_false [LawfulBEq α] + {l : List ((a : α) × β a)} {k : α} {fallback : β k} + (contains_eq_false : (l.map Sigma.fst).contains k = false) : + (ofList l).getD k fallback = fallback := by + simp only [ofList] + rw [getD_insertMany_list_of_contains_eq_false contains_eq_false] + apply getD_emptyc + +theorem getD_ofList_of_mem [LawfulBEq α] + {l : List ((a : α) × β a)} {k k' : α} (k_beq : k == k') {v : β k} {fallback : β k'} + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) + (mem : ⟨k, v⟩ ∈ l) : + (ofList l).getD k' fallback = cast (by congr; apply LawfulBEq.eq_of_beq k_beq) v := by + simp only [ofList] + rw [getD_insertMany_list_of_mem k_beq distinct mem] + +theorem getKey?_ofList_of_contains_eq_false [EquivBEq α] [LawfulHashable α] + {l : List ((a : α) × β a)} {k : α} + (h : (l.map Sigma.fst).contains k = false) : + (ofList l).getKey? k = none := by + simp only [ofList] + rw [getKey?_insertMany_list_of_contains_eq_false h] + apply getKey?_emptyc + +theorem getKey?_ofList_of_mem [EquivBEq α] [LawfulHashable α] + {l : List ((a : α) × β a)} + {k k' : α} (k_beq : k == k') + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) + (mem : k ∈ l.map Sigma.fst) : + (ofList l).getKey? k' = some k := by + simp only [ofList] + rw [getKey?_insertMany_list_of_mem k_beq distinct mem] + +theorem getKey_ofList_of_mem [EquivBEq α] [LawfulHashable α] + {l : List ((a : α) × β a)} + {k k' : α} (k_beq : k == k') + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) + (mem : k ∈ l.map Sigma.fst) + {h'} : + (ofList l).getKey k' h' = k := by + simp only [ofList] + rw [getKey_insertMany_list_of_mem k_beq distinct mem] + +theorem getKey!_ofList_of_contains_eq_false [EquivBEq α] [LawfulHashable α] [Inhabited α] + {l : List ((a : α) × β a)} {k : α} + (h : (l.map Sigma.fst).contains k = false) : + (ofList l).getKey! k = default := by + simp only [ofList] + rw [getKey!_insertMany_list_of_contains_eq_false h] + apply getKey!_emptyc + +theorem getKey!_ofList_of_mem [EquivBEq α] [LawfulHashable α] [Inhabited α] + {l : List ((a : α) × β a)} + {k k' : α} (k_beq : k == k') + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) + (mem : k ∈ l.map Sigma.fst) : + (ofList l).getKey! k' = k := by + simp only [ofList] + rw [getKey!_insertMany_list_of_mem k_beq distinct mem] + +theorem getKeyD_ofList_of_contains_eq_false [EquivBEq α] [LawfulHashable α] + {l : List ((a : α) × β a)} {k fallback : α} + (h : (l.map Sigma.fst).contains k = false) : + (ofList l).getKeyD k fallback = fallback := by + simp only [ofList] + rw [getKeyD_insertMany_list_of_contains_eq_false h] + apply getKeyD_emptyc + +theorem getKeyD_ofList_of_mem [EquivBEq α] [LawfulHashable α] + {l : List ((a : α) × β a)} + {k k' fallback : α} (k_beq : k == k') + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) + (mem : k ∈ l.map Sigma.fst) : + (ofList l).getKeyD k' fallback = k := by + simp only [ofList] + rw [getKeyD_insertMany_list_of_mem k_beq distinct mem] + +theorem size_ofList [EquivBEq α] [LawfulHashable α] + {l : List ((a : α) × β a)} (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) : + (ofList l).size = l.length := by + simp only [ofList] + rw [size_insertMany_list distinct] + · simp only [size_emptyc, Nat.zero_add] + · simp only [contains_emptyc, Bool.false_eq_true, false_implies, implies_true] + +theorem size_ofList_le [EquivBEq α] [LawfulHashable α] + {l : List ((a : α) × β a)} : + (ofList l).size ≤ l.length := by + simp only [ofList] + rw [← Nat.zero_add l.length, ← size_emptyc] + apply (size_insertMany_list_le) + +theorem isEmpty_ofList [EquivBEq α] [LawfulHashable α] + {l : List ((a : α) × β a)} : + (ofList l).isEmpty = l.isEmpty := by + simp only [ofList] + rw [isEmpty_insertMany_list] + simp + +namespace Const +variable {β : Type v} + +@[simp] +theorem ofList_nil [EquivBEq α] [LawfulHashable α] : + ofList ([] : List (α × β)) = empty := by + simp only [ofList, insertMany_nil] + rfl + +@[simp] +theorem ofList_singleton {k : α} {v : β} [EquivBEq α] [LawfulHashable α] : + ofList [⟨k, v⟩] = empty.insert k v := by + simp only [ofList, insertMany_list_singleton] + rfl + +theorem ofList_cons [EquivBEq α] [LawfulHashable α] {hd : α × β} {tl : List (α × β)} : + ofList (hd::tl) = insertMany ((∅ : DHashMap α (fun _ => β)).insert hd.1 hd.2) tl := by + simp only [ofList] + rw [insertMany_cons] + +theorem contains_ofList [EquivBEq α] [LawfulHashable α] + {l : List (α × β)} {k : α} : + (ofList l).contains k = (l.map Prod.fst).contains k := by + simp only [ofList, contains_insertMany_list, contains_emptyc, Bool.false_or] + +theorem get?_ofList_of_contains_eq_false [LawfulBEq α] + {l : List (α × β)} {k : α} + (h : (l.map Prod.fst).contains k = false) : + get? (ofList l) k = none := by + simp only [ofList] + rw [get?_insertMany_list_of_contains_eq_false _ h] + apply get?_emptyc + +theorem get?_ofList_of_mem [LawfulBEq α] + {l : List (α × β)} {k k' : α} (k_beq : k == k') {v : β} + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) + (mem : ⟨k, v⟩ ∈ l) : + get? (ofList l) k' = some v := by + simp only [ofList] + rw [get?_insertMany_list_of_mem _ k_beq mem distinct] + +theorem get_ofList_of_mem [LawfulBEq α] + {l : List (α × β)} {k k' : α} (k_beq : k == k') {v : β} + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) + (mem : ⟨k, v⟩ ∈ l) + {h} : + get (ofList l) k' h = v := by + simp only [ofList] + rw [get_insertMany_list_of_mem _ k_beq mem distinct] + +theorem get!_ofList_of_contains_eq_false [LawfulBEq α] + {l : List (α × β)} {k : α} [Inhabited β] + (h : (l.map Prod.fst).contains k = false) : + get! (ofList l) k = default := by + simp only [ofList] + rw [get!_insertMany_list_of_contains_eq_false _ h] + apply get!_emptyc + +theorem get!_ofList_of_mem [LawfulBEq α] + {l : List (α × β)} {k k' : α} (k_beq : k == k') {v : β} [Inhabited β] + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) + (mem : ⟨k, v⟩ ∈ l) : + get! (ofList l) k' = v := by + simp only [ofList] + rw [get!_insertMany_list_of_mem _ k_beq mem distinct] + +theorem getD_ofList_of_contains_eq_false [LawfulBEq α] + {l : List (α × β)} {k : α} {fallback : β} + (contains_eq_false : (l.map Prod.fst).contains k = false) : + getD (ofList l) k fallback = fallback := by + simp only [ofList] + rw [getD_insertMany_list_of_contains_eq_false _ contains_eq_false] + apply getD_emptyc + +theorem getD_ofList_of_mem [LawfulBEq α] + {l : List (α × β)} {k k' : α} (k_beq : k == k') {v : β} {fallback : β} + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) + (mem : ⟨k, v⟩ ∈ l) : + getD (ofList l) k' fallback = v := by + simp only [ofList] + rw [getD_insertMany_list_of_mem _ k_beq mem distinct] + +theorem getKey?_ofList_of_contains_eq_false [EquivBEq α] [LawfulHashable α] + {l : List (α × β)} {k : α} + (h : (l.map Prod.fst).contains k = false) : + (ofList l).getKey? k = none := by + simp only [ofList] + rw [getKey?_insertMany_list_of_contains_eq_false _ h] + apply getKey?_emptyc + +theorem getKey?_ofList_of_mem [EquivBEq α] [LawfulHashable α] + {l : List (α × β)} + {k k' : α} (k_beq : k == k') + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) + (mem : k ∈ l.map Prod.fst) : + (ofList l).getKey? k' = some k := by + simp only [ofList] + rw [getKey?_insertMany_list_of_mem _ k_beq distinct mem] + +theorem getKey_ofList_of_mem [EquivBEq α] [LawfulHashable α] + {l : List (α × β)} + {k k' : α} (k_beq : k == k') + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) + (mem : k ∈ l.map Prod.fst) + {h'} : + (ofList l).getKey k' h' = k := by + simp only [ofList] + rw [getKey_insertMany_list_of_mem _ k_beq distinct mem] + +theorem getKey!_ofList_of_contains_eq_false [EquivBEq α] [LawfulHashable α] [Inhabited α] + {l : List (α × β)} {k : α} + (h : (l.map Prod.fst).contains k = false) : + (ofList l).getKey! k = default := by + simp only [ofList] + rw [getKey!_insertMany_list_of_contains_eq_false _ h] + apply getKey!_emptyc + +theorem getKey!_ofList_of_mem [EquivBEq α] [LawfulHashable α] [Inhabited α] + {l : List (α × β)} + {k k' : α} (k_beq : k == k') + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) + (mem : k ∈ l.map Prod.fst) : + (ofList l).getKey! k' = k := by + simp only [ofList] + rw [getKey!_insertMany_list_of_mem _ k_beq distinct mem] + +theorem getKeyD_ofList_of_contains_eq_false [EquivBEq α] [LawfulHashable α] + {l : List (α × β)} {k fallback : α} + (h : (l.map Prod.fst).contains k = false) : + (ofList l).getKeyD k fallback = fallback := by + simp only [ofList] + rw [getKeyD_insertMany_list_of_contains_eq_false _ h] + apply getKeyD_emptyc + +theorem getKeyD_ofList_of_mem [EquivBEq α] [LawfulHashable α] + {l : List (α × β)} + {k k' fallback : α} (k_beq : k == k') + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) + (mem : k ∈ l.map Prod.fst) : + (ofList l).getKeyD k' fallback = k := by + simp only [ofList] + rw [getKeyD_insertMany_list_of_mem _ k_beq distinct mem] + +theorem size_ofList [EquivBEq α] [LawfulHashable α] + {l : List (α × β)} (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) : + (ofList l).size = l.length := by + simp only [ofList] + rw [size_insertMany_list _ distinct] + · simp only [size_emptyc, Nat.zero_add] + · simp only [contains_emptyc, Bool.false_eq_true, false_implies, implies_true] + +theorem size_ofList_le [EquivBEq α] [LawfulHashable α] + {l : List (α × β)} : + (ofList l).size ≤ l.length := by + simp only [ofList] + rw [← Nat.zero_add l.length, ← size_emptyc] + apply (size_insertMany_list_le _) + +theorem isEmpty_ofList [EquivBEq α] [LawfulHashable α] + {l : List (α × β)} : + (ofList l).isEmpty = l.isEmpty := by + simp [ofList, isEmpty_insertMany_list _] + +@[simp] +theorem unitOfList_nil [EquivBEq α] [LawfulHashable α] : + unitOfList ([] : List α) = empty := by + simp only [unitOfList, insertManyIfNewUnit_nil _] + rfl + +@[simp] +theorem unitOfList_singleton [EquivBEq α] [LawfulHashable α] {k : α} : + unitOfList [k] = empty.insertIfNew k () := by + simp only [unitOfList, insertManyIfNewUnit_list_singleton _] + rfl + +theorem unitOfList_cons [EquivBEq α] [LawfulHashable α] {hd : α} {tl : List α} : + unitOfList (hd::tl) = insertManyIfNewUnit ((∅ : DHashMap α (fun _ => Unit)).insertIfNew hd ()) tl := by + simp only [unitOfList] + rw [insertManyIfNewUnit_cons _] + +theorem contains_unitOfList [EquivBEq α] [LawfulHashable α] + {l : List α} {k : α} : + (unitOfList l).contains k = l.contains k := by + simp [unitOfList, contains_insertManyIfNewUnit_list _] + +theorem getKey?_unitOfList_of_contains_eq_false [EquivBEq α] [LawfulHashable α] + {l : List α} {k : α} (h' : l.contains k = false) : + getKey? (unitOfList l) k = none := by + simp [unitOfList, getKey?_insertManyIfNewUnit_list_of_contains_eq_false _ h'] + +theorem getKey?_unitOfList_of_mem [EquivBEq α] [LawfulHashable α] + {l : List α} {k k' : α} (k_beq : k == k') + (distinct : l.Pairwise (fun a b => (a == b) = false)) (mem : k ∈ l) : + getKey? (unitOfList l) k' = some k := by + simp [unitOfList, getKey?_insertManyIfNewUnit_list_of_mem_of_contains_eq_false _ k_beq + distinct mem contains_emptyc] + +theorem getKey_unitOfList_of_mem [EquivBEq α] [LawfulHashable α] + {l : List α} + {k k' : α} (k_beq : k == k') + (distinct : l.Pairwise (fun a b => (a == b) = false)) + (mem : k ∈ l) {h'} : + getKey (unitOfList l) k' h' = k := by + simp [unitOfList, getKey_insertManyIfNewUnit_list_of_mem_of_contains_eq_false _ k_beq + distinct mem contains_emptyc] + +theorem getKey!_unitOfList_of_contains_eq_false [EquivBEq α] [LawfulHashable α] + [Inhabited α] {l : List α} {k : α} + (h' : l.contains k = false) : + getKey! (unitOfList l) k = default := by + simp [unitOfList, getKey!_insertManyIfNewUnit_list_of_contains_eq_false _ h'] + +theorem getKey!_unitOfList_of_mem [EquivBEq α] [LawfulHashable α] + [Inhabited α] {l : List α} {k k' : α} (k_beq : k == k') + (distinct : l.Pairwise (fun a b => (a == b) = false)) + (mem : k ∈ l) : + getKey! (unitOfList l) k' = k := by + simp [unitOfList, getKey!_insertManyIfNewUnit_list_of_mem_of_contains_eq_false _ k_beq + distinct mem contains_emptyc] + +theorem getKeyD_unitOfList_of_contains_eq_false [EquivBEq α] [LawfulHashable α] + {l : List α} {k fallback : α} + (h' : l.contains k = false) : + getKeyD (unitOfList l) k fallback = fallback := by + simp [unitOfList, getKeyD_insertManyIfNewUnit_list_of_contains_eq_false _ h'] + +theorem getKeyD_unitOfList_of_mem [EquivBEq α] [LawfulHashable α] + {l : List α} {k k' fallback : α} (k_beq : k == k') + (distinct : l.Pairwise (fun a b => (a == b) = false)) + {mem : k ∈ l } : + getKeyD (unitOfList l) k' fallback = k := by + simp [unitOfList, getKeyD_insertManyIfNewUnit_list_of_mem_of_contains_eq_false _ k_beq + distinct mem contains_emptyc] + +theorem size_unitOfList [EquivBEq α] [LawfulHashable α] + {l : List α} + (distinct : l.Pairwise (fun a b => (a == b) = false)) : + (unitOfList l).size = l.length := by + simp only [unitOfList] + rw [size_insertManyIfNewUnit_list _ distinct] + · simp + · simp + +theorem size_unitOfList_le [EquivBEq α] [LawfulHashable α] + {l : List α} : + (unitOfList l).size ≤ l.length := by + simp only [unitOfList] + apply Nat.le_trans (size_insertManyIfNewUnit_list_le _) + simp + + +theorem isEmpty_unitOfList [EquivBEq α] [LawfulHashable α] + {l : List α} : + (unitOfList l).isEmpty = l.isEmpty := by + simp only [unitOfList] + rw [isEmpty_insertManyIfNewUnit_list _] + simp + +theorem get?_unitOfList [EquivBEq α] [LawfulHashable α] + {l : List α} {k : α} : + get? (unitOfList l) k = + if l.contains k then some () else none := by + simp only [unitOfList] + rw [get?_insertManyIfNewUnit_list _] + simp + +theorem get_unitOfList + {l : List α} {k : α} {h} : + get (unitOfList l) k h = () := by + simp + +theorem get!_unitOfList + {l : List α} {k : α} : + get! (unitOfList l) k = () := by + simp + +theorem getD_unitOfList + {l : List α} {k : α} {fallback : Unit} : + getD (unitOfList l) k fallback = () := by + simp + +end Const end Std.DHashMap From 340ac3a78b754e5facf411fc30c035a1d9f7741f Mon Sep 17 00:00:00 2001 From: jt0202 Date: Sun, 5 Jan 2025 17:58:57 +0100 Subject: [PATCH 58/83] Small style fixes --- .../DHashMap/Internal/List/Associative.lean | 22 +++++-------------- src/Std/Data/DHashMap/Internal/Model.lean | 1 + src/Std/Data/DHashMap/Internal/RawLemmas.lean | 6 +++++ 3 files changed, 12 insertions(+), 17 deletions(-) diff --git a/src/Std/Data/DHashMap/Internal/List/Associative.lean b/src/Std/Data/DHashMap/Internal/List/Associative.lean index 18ed12a7f4e9..2a77d4aacfd8 100644 --- a/src/Std/Data/DHashMap/Internal/List/Associative.lean +++ b/src/Std/Data/DHashMap/Internal/List/Associative.lean @@ -225,21 +225,6 @@ def containsKey [BEq α] (a : α) : List ((a : α) × β a) → Bool @[simp] theorem containsKey_cons [BEq α] {l : List ((a : α) × β a)} {k a : α} {v : β k} : containsKey a (⟨k, v⟩ :: l) = (k == a || containsKey a l) := rfl -theorem containsKey_eq_false_iff [BEq α][PartialEquivBEq α] {l : List ((a : α) × β a)} {a : α} : - containsKey a l = false ↔ ∀ (b : ((a : α) × β a)), b ∈ l → (a == b.fst) = false := by - induction l with - | nil => simp - | cons hd tl ih => - simp only [containsKey, Bool.or_eq_false_iff, ih, List.mem_cons, forall_eq_or_imp, - and_congr_left_iff, Bool.coe_false_iff_false, Bool.not_eq_eq_eq_not, Bool.not_not] - intro _ - rw [Bool.eq_iff_iff] - constructor - · intro h - apply PartialEquivBEq.symm h - · intro h - apply PartialEquivBEq.symm h - theorem containsKey_cons_eq_false [BEq α] {l : List ((a : α) × β a)} {k a : α} {v : β k} : (containsKey a (⟨k, v⟩ :: l) = false) ↔ ((k == a) = false) ∧ (containsKey a l = false) := by simp [containsKey_cons, not_or] @@ -901,7 +886,7 @@ theorem DistinctKeys.nil [BEq α] : DistinctKeys ([] : List ((a : α) × β a)) theorem DistinctKeys.def [BEq α] {l : List ((a : α) × β a)} : DistinctKeys l ↔ l.Pairwise (fun a b => (a.1 == b.1) = false) := ⟨fun h => by simpa [keys_eq_map, List.pairwise_map] using h.distinct, - fun h => ⟨by simpa [keys_eq_map, List.pairwise_map] using h⟩⟩ + fun h => ⟨by simpa [keys_eq_map, List.pairwise_map] using h⟩⟩ open List @@ -938,6 +923,10 @@ theorem containsKey_eq_false_iff_forall_mem_keys [BEq α] [PartialEquivBEq α] (containsKey a l) = false ↔ ∀ a' ∈ keys l, (a == a') = false := by simp only [Bool.eq_false_iff, ne_eq, containsKey_iff_exists, not_exists, not_and] +theorem containsKey_eq_false_iff [BEq α][PartialEquivBEq α] {l : List ((a : α) × β a)} {a : α} : + containsKey a l = false ↔ ∀ (b : ((a : α) × β a)), b ∈ l → (a == b.fst) = false := by + simp [containsKey_eq_false_iff_forall_mem_keys, keys_eq_map] + @[simp] theorem distinctKeys_cons_iff [BEq α] [PartialEquivBEq α] {l : List ((a : α) × β a)} {k : α} {v : β k} : DistinctKeys (⟨k, v⟩ :: l) ↔ DistinctKeys l ∧ (containsKey k l) = false := by @@ -1004,7 +993,6 @@ theorem getEntry?_of_mem [BEq α] [PartialEquivBEq α] · rw [getEntry?_cons_of_false, ih hl.tail hkv] exact BEq.neq_of_neq_of_beq (containsKey_eq_false_iff.1 hl.containsKey_eq_false _ hkv) hk - /-- Internal implementation detail of the hash map -/ def insertEntry [BEq α] (k : α) (v : β k) (l : List ((a : α) × β a)) : List ((a : α) × β a) := bif containsKey k l then replaceEntry k v l else ⟨k, v⟩ :: l diff --git a/src/Std/Data/DHashMap/Internal/Model.lean b/src/Std/Data/DHashMap/Internal/Model.lean index 30db8950f7f6..a98dc78a218e 100644 --- a/src/Std/Data/DHashMap/Internal/Model.lean +++ b/src/Std/Data/DHashMap/Internal/Model.lean @@ -359,6 +359,7 @@ def Const.insertListIfNewUnitₘ [BEq α] [Hashable α](m : Raw₀ α (fun _ => match l with | .nil => m | .cons hd tl => insertListIfNewUnitₘ (m.insertIfNew hd ()) tl + end /-! # Equivalence between model functions and real implementations -/ diff --git a/src/Std/Data/DHashMap/Internal/RawLemmas.lean b/src/Std/Data/DHashMap/Internal/RawLemmas.lean index 809577ae19b2..073c1b8c1ccc 100644 --- a/src/Std/Data/DHashMap/Internal/RawLemmas.lean +++ b/src/Std/Data/DHashMap/Internal/RawLemmas.lean @@ -859,6 +859,7 @@ theorem insertMany_cons {l : List ((a : α) × β a)} {k : α} {v : β k} : | nil => simp [insertListₘ] | cons hd tl => simp [insertListₘ] +@[simp] theorem contains_insertMany_list [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l : List ((a : α) × β a)} {k : α} : (m.insertMany l).1.contains k = (m.contains k || (l.map Sigma.fst).contains k) := by @@ -994,6 +995,7 @@ theorem size_insertMany_list_le [EquivBEq α] [LawfulHashable α] (h : m.1.WF) (m.insertMany l).1.1.size ≤ m.1.size + l.length := by simp_to_model using length_insertList_le +@[simp] theorem isEmpty_insertMany_list [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l : List ((a : α) × β a)} : (m.insertMany l).1.1.isEmpty = (m.1.isEmpty && l.isEmpty) := by @@ -1020,6 +1022,7 @@ theorem insertMany_cons {l : List (α × β)} {k : α} {v : β} : | nil => simp [insertListₘ] | cons hd tl => simp [insertListₘ] +@[simp] theorem contains_insertMany_list [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l : List (α × β)} {k : α} : (Const.insertMany m l).1.contains k = (m.contains k || (l.map Prod.fst).contains k) := by @@ -1101,6 +1104,7 @@ theorem size_insertMany_list_le [EquivBEq α] [LawfulHashable α] (h : m.1.WF) (insertMany m l).1.1.size ≤ m.1.size + l.length := by simp_to_model using length_insertListConst_le +@[simp] theorem isEmpty_insertMany_list [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l : List (α × β)} : (insertMany m l).1.1.isEmpty = (m.1.isEmpty && l.isEmpty) := by @@ -1174,6 +1178,7 @@ theorem insertManyIfNewUnit_cons {l : List α} {k : α} : | nil => simp [insertListIfNewUnitₘ] | cons hd tl => simp [insertListIfNewUnitₘ] +@[simp] theorem contains_insertManyIfNewUnit_list [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l : List α} {k : α} : (insertManyIfNewUnit m l).1.contains k = (m.contains k || l.contains k) := by @@ -1276,6 +1281,7 @@ theorem size_insertManyIfNewUnit_list_le [EquivBEq α] [LawfulHashable α] (h : (insertManyIfNewUnit m l).1.1.size ≤ m.1.size + l.length := by simp_to_model using length_insertListIfNewUnit_le +@[simp] theorem isEmpty_insertManyIfNewUnit_list [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l : List α} : (insertManyIfNewUnit m l).1.1.isEmpty = (m.1.isEmpty && l.isEmpty) := by From 7624c21deb62a57a33407c39677a6b7c8a52ced4 Mon Sep 17 00:00:00 2001 From: jt0202 Date: Sun, 5 Jan 2025 20:19:34 +0100 Subject: [PATCH 59/83] Added predecessors for ofList into Raw0 --- src/Std/Data/DHashMap/Internal/RawLemmas.lean | 422 +++++++++++++++++- src/Std/Data/DHashMap/Lemmas.lean | 8 +- 2 files changed, 418 insertions(+), 12 deletions(-) diff --git a/src/Std/Data/DHashMap/Internal/RawLemmas.lean b/src/Std/Data/DHashMap/Internal/RawLemmas.lean index 073c1b8c1ccc..413e144f8eb9 100644 --- a/src/Std/Data/DHashMap/Internal/RawLemmas.lean +++ b/src/Std/Data/DHashMap/Internal/RawLemmas.lean @@ -1117,8 +1117,9 @@ theorem get?_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable simp_to_model using getValue?_insertListConst_of_contains_eq_false theorem get?_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] (h : m.1.WF) - {l : List (α × β)} {k k' : α} (k_beq : k == k') {v : β} (mem : ⟨k, v⟩ ∈ l) - (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) : + {l : List (α × β)} {k k' : α} (k_beq : k == k') {v : β} + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) (mem : ⟨k, v⟩ ∈ l) + : get? (insertMany m l).1 k' = v := by simp_to_model using getValue?_insertListConst_of_mem @@ -1130,8 +1131,8 @@ theorem get_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable simp_to_model using getValue_insertListConst_of_contains_eq_false theorem get_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] (h : m.1.WF) - {l : List (α × β)} {k k' : α} (k_beq : k == k') {v : β} (mem : ⟨k, v⟩ ∈ l) - (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) {h'} : + {l : List (α × β)} {k k' : α} (k_beq : k == k') {v : β} + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) (mem : ⟨k, v⟩ ∈ l) {h'} : get (insertMany m l).1 k' h' = v := by simp_to_model using getValue_insertListConst_of_mem @@ -1142,8 +1143,8 @@ theorem get!_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable simp_to_model using getValue!_insertListConst_of_contains_eq_false theorem get!_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] [Inhabited β] (h : m.1.WF) - {l : List (α × β)} {k k' : α} (k_beq : k == k') {v : β} (mem : ⟨k, v⟩ ∈ l) - (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) : + {l : List (α × β)} {k k' : α} (k_beq : k == k') {v : β} + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) (mem : ⟨k, v⟩ ∈ l): get! (insertMany m l).1 k' = v := by simp_to_model using getValue!_insertListConst_of_mem @@ -1154,8 +1155,8 @@ theorem getD_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable simp_to_model using getValueD_insertListConst_of_contains_eq_false theorem getD_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] (h : m.1.WF) - {l : List (α × β)} {k k' : α} (k_beq : k == k') {v fallback : β} (mem : ⟨k, v⟩ ∈ l) - (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) : + {l : List (α × β)} {k k' : α} (k_beq : k == k') {v fallback : β} + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) (mem : ⟨k, v⟩ ∈ l) : getD (insertMany m l).1 k' fallback = v := by simp_to_model using getValueD_insertListConst_of_mem @@ -1311,4 +1312,409 @@ theorem getD_insertManyIfNewUnit_list end Const end Raw₀ +namespace Raw₀ +variable [BEq α] [Hashable α] +@[simp] +theorem insertMany_empty_list_nil [EquivBEq α] [LawfulHashable α] : + (insertMany empty ([] : List ((a : α) × (β a)))).1 = empty := by + simp + +@[simp] +theorem insertMany_empty_list_singleton {k : α} {v : β k} [EquivBEq α] [LawfulHashable α] : + (insertMany empty [⟨k, v⟩]).1 = empty.insert k v := by + simp + +theorem insertMany_empty_list_cons [EquivBEq α] [LawfulHashable α] {k : α} {v : β k} {tl : List ((a : α) × (β a))} : + (insertMany empty (⟨k, v⟩ :: tl)).1 = ((empty.insert k v).insertMany tl).1 := by + rw [insertMany_cons] + +theorem contains_insertMany_empty_list [EquivBEq α] [LawfulHashable α] + {l : List ((a : α) × β a)} {k : α} : + (insertMany empty l).1.contains k = (l.map Sigma.fst).contains k := by + simp [contains_insertMany_list _ Raw.WF.empty₀] + +theorem get?_insertMany_empty_list_of_contains_eq_false [LawfulBEq α] + {l : List ((a : α) × β a)} {k : α} + (h : (l.map Sigma.fst).contains k = false) : + (insertMany empty l).1.get? k = none := by + simp [get?_insertMany_list_of_contains_eq_false _ Raw.WF.empty₀ h] + +theorem get?_insertMany_empty_list_of_mem [LawfulBEq α] + {l : List ((a : α) × β a)} {k k' : α} (k_beq : k == k') {v : β k} + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) + (mem : ⟨k, v⟩ ∈ l) : + (insertMany empty l).1.get? k' = some (cast (by congr; apply LawfulBEq.eq_of_beq k_beq) v) := by + rw [get?_insertMany_list_of_mem _ Raw.WF.empty₀ k_beq distinct mem] + +theorem get_insertMany_empty_list_of_mem [LawfulBEq α] + {l : List ((a : α) × β a)} {k k' : α} (k_beq : k == k') {v : β k} + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) + (mem : ⟨k, v⟩ ∈ l) + {h} : + (insertMany empty l).1.get k' h = cast (by congr; apply LawfulBEq.eq_of_beq k_beq) v := by + rw [get_insertMany_list_of_mem _ Raw.WF.empty₀ k_beq distinct mem] + +theorem get!_insertMany_empty_list_of_contains_eq_false [LawfulBEq α] + {l : List ((a : α) × β a)} {k : α} [Inhabited (β k)] + (h : (l.map Sigma.fst).contains k = false) : + (insertMany empty l).1.get! k = default := by + simp [get!_insertMany_list_of_contains_eq_false _ Raw.WF.empty₀ h] + apply get!_empty + +theorem get!_insertMany_empty_list_of_mem [LawfulBEq α] + {l : List ((a : α) × β a)} {k k' : α} (k_beq : k == k') {v : β k} [Inhabited (β k')] + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) + (mem : ⟨k, v⟩ ∈ l) : + (insertMany empty l).1.get! k' = cast (by congr; apply LawfulBEq.eq_of_beq k_beq) v := by + rw [get!_insertMany_list_of_mem _ Raw.WF.empty₀ k_beq distinct mem] + +theorem getD_insertMany_empty_list_of_contains_eq_false [LawfulBEq α] + {l : List ((a : α) × β a)} {k : α} {fallback : β k} + (contains_eq_false : (l.map Sigma.fst).contains k = false) : + (insertMany empty l).1.getD k fallback = fallback := by + rw [getD_insertMany_list_of_contains_eq_false _ Raw.WF.empty₀ contains_eq_false] + apply getD_empty + +theorem getD_insertMany_empty_list_of_mem [LawfulBEq α] + {l : List ((a : α) × β a)} {k k' : α} (k_beq : k == k') {v : β k} {fallback : β k'} + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) + (mem : ⟨k, v⟩ ∈ l) : + (insertMany empty l).1.getD k' fallback = cast (by congr; apply LawfulBEq.eq_of_beq k_beq) v := by + rw [getD_insertMany_list_of_mem _ Raw.WF.empty₀ k_beq distinct mem] + +theorem getKey?_insertMany_empty_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] + {l : List ((a : α) × β a)} {k : α} + (h : (l.map Sigma.fst).contains k = false) : + (insertMany empty l).1.getKey? k = none := by + rw [getKey?_insertMany_list_of_contains_eq_false _ Raw.WF.empty₀ h] + apply getKey?_empty + +theorem getKey?_insertMany_empty_list_of_mem [EquivBEq α] [LawfulHashable α] + {l : List ((a : α) × β a)} + {k k' : α} (k_beq : k == k') + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) + (mem : k ∈ l.map Sigma.fst) : + (insertMany empty l).1.getKey? k' = some k := by + rw [getKey?_insertMany_list_of_mem _ Raw.WF.empty₀ k_beq distinct mem] + +theorem getKey_insertMany_empty_list_of_mem [EquivBEq α] [LawfulHashable α] + {l : List ((a : α) × β a)} + {k k' : α} (k_beq : k == k') + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) + (mem : k ∈ l.map Sigma.fst) + {h'} : + (insertMany empty l).1.getKey k' h' = k := by + rw [getKey_insertMany_list_of_mem _ Raw.WF.empty₀ k_beq distinct mem] + +theorem getKey!_insertMany_empty_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] [Inhabited α] + {l : List ((a : α) × β a)} {k : α} + (h : (l.map Sigma.fst).contains k = false) : + (insertMany empty l).1.getKey! k = default := by + rw [getKey!_insertMany_list_of_contains_eq_false _ Raw.WF.empty₀ h] + apply getKey!_empty + +theorem getKey!_insertMany_empty_list_of_mem [EquivBEq α] [LawfulHashable α] [Inhabited α] + {l : List ((a : α) × β a)} + {k k' : α} (k_beq : k == k') + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) + (mem : k ∈ l.map Sigma.fst) : + (insertMany empty l).1.getKey! k' = k := by + rw [getKey!_insertMany_list_of_mem _ Raw.WF.empty₀ k_beq distinct mem] + +theorem getKeyD_insertMany_empty_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] + {l : List ((a : α) × β a)} {k fallback : α} + (h : (l.map Sigma.fst).contains k = false) : + (insertMany empty l).1.getKeyD k fallback = fallback := by + rw [getKeyD_insertMany_list_of_contains_eq_false _ Raw.WF.empty₀ h] + apply getKeyD_empty + +theorem getKeyD_insertMany_empty_list_of_mem [EquivBEq α] [LawfulHashable α] + {l : List ((a : α) × β a)} + {k k' fallback : α} (k_beq : k == k') + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) + (mem : k ∈ l.map Sigma.fst) : + (insertMany empty l).1.getKeyD k' fallback = k := by + rw [getKeyD_insertMany_list_of_mem _ Raw.WF.empty₀ k_beq distinct mem] + +theorem size_insertMany_empty_list [EquivBEq α] [LawfulHashable α] + {l : List ((a : α) × β a)} (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) : + (insertMany empty l).1.1.size = l.length := by + rw [size_insertMany_list _ Raw.WF.empty₀ distinct] + · simp only [size_empty, Nat.zero_add] + · simp only [contains_empty, Bool.false_eq_true, false_implies, implies_true] + +theorem size_insertMany_empty_list_le [EquivBEq α] [LawfulHashable α] + {l : List ((a : α) × β a)} : + (insertMany empty l).1.1.size ≤ l.length := by + rw [← Nat.zero_add l.length] + apply (size_insertMany_list_le _ Raw.WF.empty₀) + +theorem isEmpty_insertMany_empty_list [EquivBEq α] [LawfulHashable α] + {l : List ((a : α) × β a)} : + (insertMany empty l).1.1.isEmpty = l.isEmpty := by + simp [isEmpty_insertMany_list _ Raw.WF.empty₀] + +namespace Const +variable {β : Type v} + +@[simp] +theorem insertMany_empty_list_nil [EquivBEq α] [LawfulHashable α] : + (insertMany empty ([] : List (α × β))).1 = empty := by + simp only [insertMany_nil] + +@[simp] +theorem insertMany_empty_list_singleton {k : α} {v : β} [EquivBEq α] [LawfulHashable α] : + (insertMany empty [⟨k, v⟩]).1 = empty.insert k v := by + simp only [insertMany_list_singleton] + +theorem insertMany_empty_list_cons [EquivBEq α] [LawfulHashable α] {k : α} {v : β} {tl : List (α × β)} : + (insertMany empty (⟨k, v⟩ :: tl)) = (insertMany (empty.insert k v) tl).1 := by + rw [insertMany_cons] + +theorem contains_insertMany_empty_list [EquivBEq α] [LawfulHashable α] + {l : List (α × β)} {k : α} : + (insertMany (empty : Raw₀ α (fun _ => β)) l).1.contains k = (l.map Prod.fst).contains k := by + simp [contains_insertMany_list _ Raw.WF.empty₀] + +theorem get?_insertMany_empty_list_of_contains_eq_false [LawfulBEq α] + {l : List (α × β)} {k : α} + (h : (l.map Prod.fst).contains k = false) : + get? (insertMany (empty : Raw₀ α (fun _ => β)) l).1 k = none := by + rw [get?_insertMany_list_of_contains_eq_false _ Raw.WF.empty₀ h] + apply get?_empty + +theorem get?_insertMany_empty_list_of_mem [LawfulBEq α] + {l : List (α × β)} {k k' : α} (k_beq : k == k') {v : β} + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) + (mem : ⟨k, v⟩ ∈ l) : + get? (insertMany (empty : Raw₀ α (fun _ => β)) l) k' = some v := by + rw [get?_insertMany_list_of_mem _ Raw.WF.empty₀ k_beq distinct mem] + +theorem get_insertMany_empty_list_of_mem [LawfulBEq α] + {l : List (α × β)} {k k' : α} (k_beq : k == k') {v : β} + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) + (mem : ⟨k, v⟩ ∈ l) + {h} : + get (insertMany (empty : Raw₀ α (fun _ => β)) l) k' h = v := by + + rw [get_insertMany_list_of_mem _ Raw.WF.empty₀ k_beq distinct mem] + +theorem get!_insertMany_empty_list_of_contains_eq_false [LawfulBEq α] + {l : List (α × β)} {k : α} [Inhabited β] + (h : (l.map Prod.fst).contains k = false) : + get! (insertMany (empty : Raw₀ α (fun _ => β)) l) k = (default : β) := by + + rw [get!_insertMany_list_of_contains_eq_false _ Raw.WF.empty₀ h] + apply get!_empty + +theorem get!_insertMany_empty_list_of_mem [LawfulBEq α] + {l : List (α × β)} {k k' : α} (k_beq : k == k') {v : β} [Inhabited β] + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) + (mem : ⟨k, v⟩ ∈ l) : + get! (insertMany (empty : Raw₀ α (fun _ => β)) l) k' = v := by + + rw [get!_insertMany_list_of_mem _ Raw.WF.empty₀ k_beq distinct mem] + +theorem getD_insertMany_empty_list_of_contains_eq_false [LawfulBEq α] + {l : List (α × β)} {k : α} {fallback : β} + (contains_eq_false : (l.map Prod.fst).contains k = false) : + getD (insertMany (empty : Raw₀ α (fun _ => β)) l) k fallback = fallback := by + + rw [getD_insertMany_list_of_contains_eq_false _ Raw.WF.empty₀ contains_eq_false] + apply getD_empty + +theorem getD_insertMany_empty_list_of_mem [LawfulBEq α] + {l : List (α × β)} {k k' : α} (k_beq : k == k') {v : β} {fallback : β} + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) + (mem : ⟨k, v⟩ ∈ l) : + getD (insertMany (empty : Raw₀ α (fun _ => β)) l) k' fallback = v := by + rw [getD_insertMany_list_of_mem _ Raw.WF.empty₀ k_beq distinct mem] + +theorem getKey?_insertMany_empty_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] + {l : List (α × β)} {k : α} + (h : (l.map Prod.fst).contains k = false) : + (insertMany (empty : Raw₀ α (fun _ => β)) l).1.getKey? k = none := by + rw [getKey?_insertMany_list_of_contains_eq_false _ Raw.WF.empty₀ h] + apply getKey?_empty + +theorem getKey?_insertMany_empty_list_of_mem [EquivBEq α] [LawfulHashable α] + {l : List (α × β)} + {k k' : α} (k_beq : k == k') + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) + (mem : k ∈ l.map Prod.fst) : + (insertMany (empty : Raw₀ α (fun _ => β)) l).1.getKey? k' = some k := by + rw [getKey?_insertMany_list_of_mem _ Raw.WF.empty₀ k_beq distinct mem] + +theorem getKey_insertMany_empty_list_of_mem [EquivBEq α] [LawfulHashable α] + {l : List (α × β)} + {k k' : α} (k_beq : k == k') + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) + (mem : k ∈ l.map Prod.fst) + {h'} : + (insertMany (empty : Raw₀ α (fun _ => β)) l).1.getKey k' h' = k := by + + rw [getKey_insertMany_list_of_mem _ Raw.WF.empty₀ k_beq distinct mem] + +theorem getKey!_insertMany_empty_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] + [Inhabited α] {l : List (α × β)} {k : α} + (h : (l.map Prod.fst).contains k = false) : + (insertMany (empty : Raw₀ α (fun _ => β)) l).1.getKey! k = default := by + rw [getKey!_insertMany_list_of_contains_eq_false _ Raw.WF.empty₀ h] + apply getKey!_empty + +theorem getKey!_insertMany_empty_list_of_mem [EquivBEq α] [LawfulHashable α] [Inhabited α] + {l : List (α × β)} + {k k' : α} (k_beq : k == k') + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) + (mem : k ∈ l.map Prod.fst) : + (insertMany (empty : Raw₀ α (fun _ => β)) l).1.getKey! k' = k := by + rw [getKey!_insertMany_list_of_mem _ Raw.WF.empty₀ k_beq distinct mem] + +theorem getKeyD_insertMany_empty_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] + {l : List (α × β)} {k fallback : α} + (h : (l.map Prod.fst).contains k = false) : + (insertMany (empty : Raw₀ α (fun _ => β)) l).1.getKeyD k fallback = fallback := by + rw [getKeyD_insertMany_list_of_contains_eq_false _ Raw.WF.empty₀ h] + apply getKeyD_empty + +theorem getKeyD_insertMany_empty_list_of_mem [EquivBEq α] [LawfulHashable α] + {l : List (α × β)} + {k k' fallback : α} (k_beq : k == k') + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) + (mem : k ∈ l.map Prod.fst) : + (insertMany (empty : Raw₀ α (fun _ => β)) l).1.getKeyD k' fallback = k := by + rw [getKeyD_insertMany_list_of_mem _ Raw.WF.empty₀ k_beq distinct mem] + +theorem size_insertMany_empty_list [EquivBEq α] [LawfulHashable α] + {l : List (α × β)} (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) : + (insertMany (empty : Raw₀ α (fun _ => β)) l).1.1.size = l.length := by + rw [size_insertMany_list _ Raw.WF.empty₀ distinct] + · simp only [size_empty, Nat.zero_add] + · simp only [contains_empty, Bool.false_eq_true, false_implies, implies_true] + +theorem size_insertMany_empty_list_le [EquivBEq α] [LawfulHashable α] + {l : List (α × β)} : + (insertMany (empty : Raw₀ α (fun _ => β)) l).1.1.size ≤ l.length := by + rw [← Nat.zero_add l.length] + apply (size_insertMany_list_le _ Raw.WF.empty₀) + +theorem isEmpty_insertMany_empty_list [EquivBEq α] [LawfulHashable α] + {l : List (α × β)} : + (insertMany (empty : Raw₀ α (fun _ => β)) l).1.1.isEmpty = l.isEmpty := by + simp [isEmpty_insertMany_list _ Raw.WF.empty₀] + +@[simp] +theorem insertManyIfNewUnit_empty_list_nil [EquivBEq α] [LawfulHashable α] : + insertManyIfNewUnit (empty : Raw₀ α (fun _ => Unit)) ([] : List α) = + (empty : Raw₀ α (fun _ => Unit)) := by + simp + +@[simp] +theorem insertManyIfNewUnit_empty_list_singleton [EquivBEq α] [LawfulHashable α] {k : α} : + (insertManyIfNewUnit (empty : Raw₀ α (fun _ => Unit)) [k]).1 = empty.insertIfNew k () := by + simp + +theorem insertManyIfNewUnit_empty_list_cons [EquivBEq α] [LawfulHashable α] {hd : α} {tl : List α} : + insertManyIfNewUnit (empty : Raw₀ α (fun _ => Unit)) (hd :: tl) = + (insertManyIfNewUnit (empty.insertIfNew hd ()) tl).1 := by + rw [insertManyIfNewUnit_cons] + +theorem contains_insertManyIfNewUnit_empty_list [EquivBEq α] [LawfulHashable α] + {l : List α} {k : α} : + (insertManyIfNewUnit (empty : Raw₀ α (fun _ => Unit)) l).1.contains k = l.contains k := by + simp [contains_insertManyIfNewUnit_list _ Raw.WF.empty₀] + +theorem getKey?_insertManyIfNewUnit_empty_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] + {l : List α} {k : α} (h' : l.contains k = false) : + getKey? (insertManyIfNewUnit (empty : Raw₀ α (fun _ => Unit)) l).1 k = none := by + simp [getKey?_insertManyIfNewUnit_list_of_contains_eq_false _ Raw.WF.empty₀ h'] + +theorem getKey?_insertManyIfNewUnit_empty_list_of_mem [EquivBEq α] [LawfulHashable α] + {l : List α} {k k' : α} (k_beq : k == k') + (distinct : l.Pairwise (fun a b => (a == b) = false)) (mem : k ∈ l) : + getKey? (insertManyIfNewUnit (empty : Raw₀ α (fun _ => Unit)) l).1 k' = some k := by + simp [getKey?_insertManyIfNewUnit_list_of_mem_of_contains_eq_false _ Raw.WF.empty₀ k_beq + distinct mem contains_empty] + +theorem getKey_insertManyIfNewUnit_empty_list_of_mem [EquivBEq α] [LawfulHashable α] + {l : List α} + {k k' : α} (k_beq : k == k') + (distinct : l.Pairwise (fun a b => (a == b) = false)) + (mem : k ∈ l) {h'} : + getKey (insertManyIfNewUnit (empty : Raw₀ α (fun _ => Unit)) l).1 k' h' = k := by + simp [getKey_insertManyIfNewUnit_list_of_mem_of_contains_eq_false _ Raw.WF.empty₀ k_beq + distinct mem contains_empty] + +theorem getKey!_insertManyIfNewUnit_empty_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] + [Inhabited α] {l : List α} {k : α} + (h' : l.contains k = false) : + getKey! (insertManyIfNewUnit (empty : Raw₀ α (fun _ => Unit)) l).1 k = default := by + simp [getKey!_insertManyIfNewUnit_list_of_contains_eq_false _ Raw.WF.empty₀ h'] + apply getKey!_empty + +theorem getKey!_insertManyIfNewUnit_empty_list_of_mem [EquivBEq α] [LawfulHashable α] + [Inhabited α] {l : List α} {k k' : α} (k_beq : k == k') + (distinct : l.Pairwise (fun a b => (a == b) = false)) + (mem : k ∈ l) : + getKey! (insertManyIfNewUnit (empty : Raw₀ α (fun _ => Unit)) l).1 k' = k := by + simp [getKey!_insertManyIfNewUnit_list_of_mem_of_contains_eq_false _ Raw.WF.empty₀ k_beq + distinct mem contains_empty] + +theorem getKeyD_insertManyIfNewUnit_empty_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] + {l : List α} {k fallback : α} + (h' : l.contains k = false) : + getKeyD (insertManyIfNewUnit (empty : Raw₀ α (fun _ => Unit)) l).1 k fallback = fallback := by + simp [getKeyD_insertManyIfNewUnit_list_of_contains_eq_false _ Raw.WF.empty₀ h'] + apply getKeyD_empty + +theorem getKeyD_insertManyIfNewUnit_empty_list_of_mem [EquivBEq α] [LawfulHashable α] + {l : List α} {k k' fallback : α} (k_beq : k == k') + (distinct : l.Pairwise (fun a b => (a == b) = false)) + {mem : k ∈ l } : + getKeyD (insertManyIfNewUnit (empty : Raw₀ α (fun _ => Unit)) l).1 k' fallback = k := by + simp [getKeyD_insertManyIfNewUnit_list_of_mem_of_contains_eq_false _ Raw.WF.empty₀ k_beq + distinct mem contains_empty] + +theorem size_insertManyIfNewUnit_empty_list [EquivBEq α] [LawfulHashable α] + {l : List α} + (distinct : l.Pairwise (fun a b => (a == b) = false)) : + (insertManyIfNewUnit (empty : Raw₀ α (fun _ => Unit)) l).1.1.size = l.length := by + simp [size_insertManyIfNewUnit_list _ Raw.WF.empty₀ distinct] + +theorem size_insertManyIfNewUnit_empty_list_le [EquivBEq α] [LawfulHashable α] + {l : List α} : + (insertManyIfNewUnit (empty : Raw₀ α (fun _ => Unit)) l).1.1.size ≤ l.length := by + apply Nat.le_trans (size_insertManyIfNewUnit_list_le _ Raw.WF.empty₀) + simp + +theorem isEmpty_insertManyIfNewUnit_empty_list [EquivBEq α] [LawfulHashable α] + {l : List α} : + (insertManyIfNewUnit (empty : Raw₀ α (fun _ => Unit)) l).1.1.isEmpty = l.isEmpty := by + rw [isEmpty_insertManyIfNewUnit_list _ Raw.WF.empty₀] + simp + +theorem get?_insertManyIfNewUnit_empty_list [EquivBEq α] [LawfulHashable α] + {l : List α} {k : α} : + get? (insertManyIfNewUnit (empty : Raw₀ α (fun _ => Unit)) l) k = + if l.contains k then some () else none := by + rw [get?_insertManyIfNewUnit_list _ Raw.WF.empty₀] + simp + +theorem get_insertManyIfNewUnit_empty_list + {l : List α} {k : α} {h} : + get (insertManyIfNewUnit (empty : Raw₀ α (fun _ => Unit)) l) k h = () := by + simp + +theorem get!_insertManyIfNewUnit_empty_list + {l : List α} {k : α} : + get! (insertManyIfNewUnit (empty : Raw₀ α (fun _ => Unit)) l) k = () := by + simp + +theorem getD_insertManyIfNewUnit_empty_list + {l : List α} {k : α} {fallback : Unit} : + getD (insertManyIfNewUnit (empty : Raw₀ α (fun _ => Unit)) l) k fallback = () := by + simp + +end Const +end Raw₀ end Std.DHashMap.Internal diff --git a/src/Std/Data/DHashMap/Lemmas.lean b/src/Std/Data/DHashMap/Lemmas.lean index 080ecea6833e..2a5165d0deee 100644 --- a/src/Std/Data/DHashMap/Lemmas.lean +++ b/src/Std/Data/DHashMap/Lemmas.lean @@ -1240,7 +1240,7 @@ theorem get?_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] {l : List (α × β)} {k k' : α} (k_beq : k == k') {v : β} (mem : ⟨k, v⟩ ∈ l) (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) : get? (insertMany m l) k' = v := - Raw₀.Const.get?_insertMany_list_of_mem ⟨m.1, _⟩ m.2 k_beq mem distinct + Raw₀.Const.get?_insertMany_list_of_mem ⟨m.1, _⟩ m.2 k_beq distinct mem theorem get_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] {l : List (α × β)} {k : α} @@ -1253,7 +1253,7 @@ theorem get_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] {l : List (α × β)} {k k' : α} (k_beq : k == k') {v : β} (mem : ⟨k, v⟩ ∈ l) (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) {h'} : get (insertMany m l) k' h' = v := - Raw₀.Const.get_insertMany_list_of_mem ⟨m.1, _⟩ m.2 k_beq mem distinct + Raw₀.Const.get_insertMany_list_of_mem ⟨m.1, _⟩ m.2 k_beq distinct mem theorem get!_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] [Inhabited β] {l : List (α × β)} {k : α} @@ -1265,7 +1265,7 @@ theorem get!_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] [Inhabited {l : List (α × β)} {k k' : α} (k_beq : k == k') {v : β} (mem : ⟨k, v⟩ ∈ l) (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) : get! (insertMany m l) k' = v := - Raw₀.Const.get!_insertMany_list_of_mem ⟨m.1, _⟩ m.2 k_beq mem distinct + Raw₀.Const.get!_insertMany_list_of_mem ⟨m.1, _⟩ m.2 k_beq distinct mem theorem getD_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] {l : List (α × β)} {k : α} {fallback : β} @@ -1277,7 +1277,7 @@ theorem getD_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] {l : List (α × β)} {k k' : α} (k_beq : k == k') {v fallback : β} (mem : ⟨k, v⟩ ∈ l) (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) : getD (insertMany m l) k' fallback = v := - Raw₀.Const.getD_insertMany_list_of_mem ⟨m.1, _⟩ m.2 k_beq mem distinct + Raw₀.Const.getD_insertMany_list_of_mem ⟨m.1, _⟩ m.2 k_beq distinct mem variable (m : DHashMap α (fun _ => Unit)) From 37fd60f0ec33dadf1c741d02a3cde14f0bf3b790 Mon Sep 17 00:00:00 2001 From: jt0202 Date: Sun, 5 Jan 2025 21:37:16 +0100 Subject: [PATCH 60/83] Rework of DHashMap lemmas for ofList --- src/Std/Data/DHashMap/Internal/RawLemmas.lean | 2 +- src/Std/Data/DHashMap/Lemmas.lean | 441 ++++++++---------- 2 files changed, 189 insertions(+), 254 deletions(-) diff --git a/src/Std/Data/DHashMap/Internal/RawLemmas.lean b/src/Std/Data/DHashMap/Internal/RawLemmas.lean index 413e144f8eb9..c942d16c70d6 100644 --- a/src/Std/Data/DHashMap/Internal/RawLemmas.lean +++ b/src/Std/Data/DHashMap/Internal/RawLemmas.lean @@ -1670,7 +1670,7 @@ theorem getKeyD_insertManyIfNewUnit_empty_list_of_contains_eq_false [EquivBEq α theorem getKeyD_insertManyIfNewUnit_empty_list_of_mem [EquivBEq α] [LawfulHashable α] {l : List α} {k k' fallback : α} (k_beq : k == k') (distinct : l.Pairwise (fun a b => (a == b) = false)) - {mem : k ∈ l } : + (mem : k ∈ l) : getKeyD (insertManyIfNewUnit (empty : Raw₀ α (fun _ => Unit)) l).1 k' fallback = k := by simp [getKeyD_insertManyIfNewUnit_list_of_mem_of_contains_eq_false _ Raw.WF.empty₀ k_beq distinct mem contains_empty] diff --git a/src/Std/Data/DHashMap/Lemmas.lean b/src/Std/Data/DHashMap/Lemmas.lean index 2a5165d0deee..fc72c5ba33a5 100644 --- a/src/Std/Data/DHashMap/Lemmas.lean +++ b/src/Std/Data/DHashMap/Lemmas.lean @@ -966,23 +966,19 @@ theorem distinct_keys [EquivBEq α] [LawfulHashable α] : @[simp] theorem insertMany_nil : - m.insertMany [] = m := by - apply Subtype.eq - congr + m.insertMany [] = m := + Subtype.eq (congrArg Subtype.val (Raw₀.insertMany_nil ⟨m.1, m.2.size_buckets_pos⟩) :) @[simp] theorem insertMany_list_singleton {k : α} {v : β k} : - m.insertMany [⟨k, v⟩] = m.insert k v := by - apply Subtype.eq - congr + m.insertMany [⟨k, v⟩] = m.insert k v := + Subtype.eq (congrArg Subtype.val (Raw₀.insertMany_list_singleton ⟨m.1, m.2.size_buckets_pos⟩) :) theorem insertMany_cons {l : List ((a : α) × β a)} {k : α} {v : β k} : - (m.insertMany (⟨k, v⟩ :: l)) = ((m.insert k v).insertMany l) := by - apply Subtype.eq - simp [insertMany] - rw [Raw₀.insertMany_cons ⟨m.1, _⟩] - rfl + (m.insertMany (⟨k, v⟩ :: l)) = ((m.insert k v).insertMany l) := + Subtype.eq (congrArg Subtype.val (Raw₀.insertMany_cons ⟨m.1, m.2.size_buckets_pos⟩) :) +@[simp] theorem contains_insertMany_list [EquivBEq α] [LawfulHashable α] {l : List ((a : α) × β a)} {k : α} : (m.insertMany l).contains k = (m.contains k || (l.map Sigma.fst).contains k) := @@ -1118,6 +1114,7 @@ theorem size_insertMany_list_le [EquivBEq α] [LawfulHashable α] (m.insertMany l).size ≤ m.size + l.length := Raw₀.size_insertMany_list_le ⟨m.1, _⟩ m.2 +@[simp] theorem isEmpty_insertMany_list [EquivBEq α] [LawfulHashable α] {l : List ((a : α) × β a)} : (m.insertMany l).isEmpty = (m.isEmpty && l.isEmpty) := @@ -1128,22 +1125,19 @@ variable {β : Type v} (m : DHashMap α (fun _ => β)) @[simp] theorem insertMany_nil : - insertMany m [] = m := by - apply Subtype.eq - congr + insertMany m [] = m := + Subtype.eq (congrArg Subtype.val (Raw₀.Const.insertMany_nil ⟨m.1, m.2.size_buckets_pos⟩) :) @[simp] theorem insertMany_list_singleton {k : α} {v : β} : - insertMany m [⟨k, v⟩] = m.insert k v := by - apply Subtype.eq - congr + insertMany m [⟨k, v⟩] = m.insert k v := + Subtype.eq (congrArg Subtype.val (Raw₀.Const.insertMany_list_singleton ⟨m.1, m.2.size_buckets_pos⟩) :) theorem insertMany_cons {l : List (α × β)} {k : α} {v : β} : - (insertMany m (⟨k, v⟩ :: l)) = (insertMany (m.insert k v) l) := by - apply Subtype.eq - simp [insertMany] - rw [Raw₀.Const.insertMany_cons ⟨m.1, _⟩] - rfl + (insertMany m (⟨k, v⟩ :: l)) = (insertMany (m.insert k v) l) := + Subtype.eq (congrArg Subtype.val (Raw₀.Const.insertMany_cons ⟨m.1, m.2.size_buckets_pos⟩) :) + +@[simp] theorem contains_insertMany_list [EquivBEq α] [LawfulHashable α] {l : List (α × β)} {k : α} : (Const.insertMany m l).contains k = (m.contains k || (l.map Prod.fst).contains k) := @@ -1225,6 +1219,7 @@ theorem size_insertMany_list_le [EquivBEq α] [LawfulHashable α] (insertMany m l).size ≤ m.size + l.length := Raw₀.Const.size_insertMany_list_le ⟨m.1, _⟩ m.2 +@[simp] theorem isEmpty_insertMany_list [EquivBEq α] [LawfulHashable α] {l : List (α × β)} : (insertMany m l).isEmpty = (m.isEmpty && l.isEmpty) := @@ -1237,10 +1232,11 @@ theorem get?_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable Raw₀.Const.get?_insertMany_list_of_contains_eq_false ⟨m.1, _⟩ m.2 h' theorem get?_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] - {l : List (α × β)} {k k' : α} (k_beq : k == k') {v : β} (mem : ⟨k, v⟩ ∈ l) - (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) : + {l : List (α × β)} {k k' : α} (k_beq : k == k') {v : β} + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) (mem : ⟨k, v⟩ ∈ l) + : get? (insertMany m l) k' = v := - Raw₀.Const.get?_insertMany_list_of_mem ⟨m.1, _⟩ m.2 k_beq distinct mem + Raw₀.Const.get?_insertMany_list_of_mem ⟨m.1, _⟩ m.2 k_beq distinct mem theorem get_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] {l : List (α × β)} {k : α} @@ -1250,10 +1246,10 @@ theorem get_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable Raw₀.Const.get_insertMany_list_of_contains_eq_false ⟨m.1, _⟩ m.2 h₁ theorem get_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] - {l : List (α × β)} {k k' : α} (k_beq : k == k') {v : β} (mem : ⟨k, v⟩ ∈ l) - (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) {h'} : + {l : List (α × β)} {k k' : α} (k_beq : k == k') {v : β} + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) (mem : ⟨k, v⟩ ∈ l) {h'} : get (insertMany m l) k' h' = v := - Raw₀.Const.get_insertMany_list_of_mem ⟨m.1, _⟩ m.2 k_beq distinct mem + Raw₀.Const.get_insertMany_list_of_mem ⟨m.1, _⟩ m.2 k_beq distinct mem theorem get!_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] [Inhabited β] {l : List (α × β)} {k : α} @@ -1262,10 +1258,10 @@ theorem get!_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable Raw₀.Const.get!_insertMany_list_of_contains_eq_false ⟨m.1, _⟩ m.2 h' theorem get!_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] [Inhabited β] - {l : List (α × β)} {k k' : α} (k_beq : k == k') {v : β} (mem : ⟨k, v⟩ ∈ l) - (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) : + {l : List (α × β)} {k k' : α} (k_beq : k == k') {v : β} + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) (mem : ⟨k, v⟩ ∈ l): get! (insertMany m l) k' = v := - Raw₀.Const.get!_insertMany_list_of_mem ⟨m.1, _⟩ m.2 k_beq distinct mem + Raw₀.Const.get!_insertMany_list_of_mem ⟨m.1, _⟩ m.2 k_beq distinct mem theorem getD_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] {l : List (α × β)} {k : α} {fallback : β} @@ -1274,32 +1270,32 @@ theorem getD_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable Raw₀.Const.getD_insertMany_list_of_contains_eq_false ⟨m.1, _⟩ m.2 h' theorem getD_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] - {l : List (α × β)} {k k' : α} (k_beq : k == k') {v fallback : β} (mem : ⟨k, v⟩ ∈ l) - (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) : + {l : List (α × β)} {k k' : α} (k_beq : k == k') {v fallback : β} + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) (mem : ⟨k, v⟩ ∈ l) : getD (insertMany m l) k' fallback = v := - Raw₀.Const.getD_insertMany_list_of_mem ⟨m.1, _⟩ m.2 k_beq distinct mem + Raw₀.Const.getD_insertMany_list_of_mem ⟨m.1, _⟩ m.2 k_beq distinct mem variable (m : DHashMap α (fun _ => Unit)) @[simp] theorem insertManyIfNewUnit_nil : - insertManyIfNewUnit m [] = m := by - apply Subtype.eq - congr + insertManyIfNewUnit m [] = m := + Subtype.eq (congrArg Subtype.val (Raw₀.Const.insertManyIfNewUnit_nil ⟨m.1, m.2.size_buckets_pos⟩) :) + @[simp] theorem insertManyIfNewUnit_list_singleton [EquivBEq α] [LawfulHashable α] {k : α} : - insertManyIfNewUnit m [k] = m.insertIfNew k () := by - apply Subtype.eq - congr + insertManyIfNewUnit m [k] = m.insertIfNew k () := + Subtype.eq (congrArg Subtype.val + (Raw₀.Const.insertManyIfNewUnit_list_singleton ⟨m.1, m.2.size_buckets_pos⟩) :) + theorem insertManyIfNewUnit_cons {l : List α} {k : α} : - (insertManyIfNewUnit m (k :: l)) = (insertManyIfNewUnit (m.insertIfNew k ()) l) := by - apply Subtype.eq - simp [insertManyIfNewUnit] - rw [Raw₀.Const.insertManyIfNewUnit_cons ⟨m.1, _⟩] - rfl + (insertManyIfNewUnit m (k :: l)) = (insertManyIfNewUnit (m.insertIfNew k ()) l) := + Subtype.eq (congrArg Subtype.val (Raw₀.Const.insertManyIfNewUnit_cons ⟨m.1, m.2.size_buckets_pos⟩) :) + +@[simp] theorem contains_insertManyIfNewUnit_list [EquivBEq α] [LawfulHashable α] {l : List α} {k : α} : (insertManyIfNewUnit m l).contains k = (m.contains k || l.contains k) := @@ -1319,7 +1315,8 @@ theorem getKey?_insertManyIfNewUnit_list_of_mem_of_contains_eq_false [EquivBEq {l : List α} {k k' : α} (k_beq : k == k') (distinct : l.Pairwise (fun a b => (a == b) = false)) (mem : k ∈ l) : m.contains k = false → getKey? (insertManyIfNewUnit m l) k' = some k := - Raw₀.Const.getKey?_insertManyIfNewUnit_list_of_mem_of_contains_eq_false ⟨m.1, _⟩ m.2 k_beq distinct mem + Raw₀.Const.getKey?_insertManyIfNewUnit_list_of_mem_of_contains_eq_false ⟨m.1, _⟩ + m.2 k_beq distinct mem theorem getKey?_insertManyIfNewUnit_list_of_contains_of_contains [EquivBEq α] [LawfulHashable α] {l : List α} {k : α} @@ -1341,14 +1338,15 @@ theorem getKey_insertManyIfNewUnit_list_of_mem_of_contains_eq_false [EquivBEq α (mem : k ∈ l) {h'} : m.contains k = false → getKey (insertManyIfNewUnit m l) k' h' = k := - Raw₀.Const.getKey_insertManyIfNewUnit_list_of_mem_of_contains_eq_false ⟨m.1, _⟩ m.2 k_beq distinct mem + Raw₀.Const.getKey_insertManyIfNewUnit_list_of_mem_of_contains_eq_false ⟨m.1, _⟩ m.2 k_beq + distinct mem theorem getKey_insertManyIfNewUnit_list_mem_of_contains_of_contains [EquivBEq α] [LawfulHashable α] {l : List α} {k : α} (distinct : l.Pairwise (fun a b => (a == b) = false)) (h₁ : l.contains k) {mem' : m.contains k} {h'} : getKey (insertManyIfNewUnit m l) k h' = getKey m k mem' := - Raw₀.Const.getKey_insertManyIfNewUnit_list_mem_of_contains_of_contains ⟨m.1, _⟩ m.2 distinct h₁ + Raw₀.Const.getKey_insertManyIfNewUnit_list_mem_of_contains_of_contains ⟨m.1, _⟩ m.2 distinct h₁ theorem getKey!_insertManyIfNewUnit_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] [Inhabited α] {l : List α} {k : α} @@ -1361,7 +1359,8 @@ theorem getKey!_insertManyIfNewUnit_list_of_mem_of_contains_eq_false [EquivBEq (distinct : l.Pairwise (fun a b => (a == b) = false)) (mem : k ∈ l) : contains m k = false → getKey! (insertManyIfNewUnit m l) k' = k := - Raw₀.Const.getKey!_insertManyIfNewUnit_list_of_mem_of_contains_eq_false ⟨m.1, _⟩ m.2 k_beq distinct mem + Raw₀.Const.getKey!_insertManyIfNewUnit_list_of_mem_of_contains_eq_false ⟨m.1, _⟩ m.2 k_beq + distinct mem theorem getKey!_insertManyIfNewUnit_list_mem_of_contains_of_contains [EquivBEq α] [LawfulHashable α] [Inhabited α] {l : List α} {k : α} @@ -1381,7 +1380,8 @@ theorem getKeyD_insertManyIfNewUnit_list_of_mem_of_contains_eq_false [EquivBEq (distinct : l.Pairwise (fun a b => (a == b) = false)) (mem : k ∈ l ) : m.contains k = false → getKeyD (insertManyIfNewUnit m l) k' fallback = k := - Raw₀.Const.getKeyD_insertManyIfNewUnit_list_of_mem_of_contains_eq_false ⟨m.1, _⟩ m.2 k_beq distinct mem + Raw₀.Const.getKeyD_insertManyIfNewUnit_list_of_mem_of_contains_eq_false ⟨m.1, _⟩ m.2 k_beq + distinct mem theorem getKeyD_insertManyIfNewUnit_list_of_contains_of_contains [EquivBEq α] [LawfulHashable α] {l : List α} {k fallback : α} @@ -1402,6 +1402,7 @@ theorem size_insertManyIfNewUnit_list_le [EquivBEq α] [LawfulHashable α] (insertManyIfNewUnit m l).size ≤ m.size + l.length := Raw₀.Const.size_insertManyIfNewUnit_list_le ⟨m.1, _⟩ m.2 +@[simp] theorem isEmpty_insertManyIfNewUnit_list [EquivBEq α] [LawfulHashable α] {l : List α} : (insertManyIfNewUnit m l).isEmpty = (m.isEmpty && l.isEmpty) := @@ -1429,105 +1430,89 @@ theorem getD_insertManyIfNewUnit_list simp end Const -end Std.DHashMap +end DHashMap -namespace Std.DHashMap -variable [BEq α] [Hashable α] +namespace DHashMap @[simp] theorem ofList_nil [EquivBEq α] [LawfulHashable α] : - ofList ([] : List ((a : α) × (β a))) = empty := by - simp only [ofList, insertMany_nil] - rfl + ofList ([] : List ((a : α) × (β a))) = ∅ := + Subtype.eq (congrArg Subtype.val (Raw₀.insertMany_empty_list_nil (α := α)) :) @[simp] theorem ofList_singleton {k : α} {v : β k} [EquivBEq α] [LawfulHashable α] : - ofList [⟨k, v⟩] = empty.insert k v := by - simp only [ofList, insertMany_list_singleton] - rfl + ofList [⟨k, v⟩] = empty.insert k v := + Subtype.eq (congrArg Subtype.val (Raw₀.insertMany_empty_list_singleton (α := α)) :) -theorem ofList_cons [EquivBEq α] [LawfulHashable α] {hd : (a : α) × (β a)} {tl : List ((a : α) × (β a))} : - ofList (hd::tl) = ((∅ : DHashMap α β).insert hd.1 hd.2).insertMany tl := by - simp only [ofList] - rw [insertMany_cons] +theorem ofList_cons [EquivBEq α] [LawfulHashable α] {k : α} {v : β k} {tl : List ((a : α) × (β a))} : + ofList (⟨k, v⟩ :: tl) = (empty.insert k v).insertMany tl := + Subtype.eq (congrArg Subtype.val (Raw₀.insertMany_empty_list_cons (α := α)) :) +@[simp] theorem contains_ofList [EquivBEq α] [LawfulHashable α] {l : List ((a : α) × β a)} {k : α} : - (ofList l).contains k = (l.map Sigma.fst).contains k := by - simp only [ofList, contains_insertMany_list, contains_emptyc, Bool.false_or] + (ofList l).contains k = (l.map Sigma.fst).contains k := + Raw₀.contains_insertMany_empty_list theorem get?_ofList_of_contains_eq_false [LawfulBEq α] {l : List ((a : α) × β a)} {k : α} (h : (l.map Sigma.fst).contains k = false) : - (ofList l).get? k = none := by - simp only [ofList] - rw [get?_insertMany_list_of_contains_eq_false h] - apply get?_emptyc + (ofList l).get? k = none := + Raw₀.get?_insertMany_empty_list_of_contains_eq_false h theorem get?_ofList_of_mem [LawfulBEq α] {l : List ((a : α) × β a)} {k k' : α} (k_beq : k == k') {v : β k} (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) (mem : ⟨k, v⟩ ∈ l) : - (ofList l).get? k' = some (cast (by congr; apply LawfulBEq.eq_of_beq k_beq) v) := by - simp only [ofList] - rw [get?_insertMany_list_of_mem k_beq distinct mem] + (ofList l).get? k' = some (cast (by congr; apply LawfulBEq.eq_of_beq k_beq) v) := + Raw₀.get?_insertMany_empty_list_of_mem k_beq distinct mem theorem get_ofList_of_mem [LawfulBEq α] {l : List ((a : α) × β a)} {k k' : α} (k_beq : k == k') {v : β k} (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) (mem : ⟨k, v⟩ ∈ l) {h} : - (ofList l).get k' h = cast (by congr; apply LawfulBEq.eq_of_beq k_beq) v := by - simp only [ofList] - rw [get_insertMany_list_of_mem k_beq distinct mem] + (ofList l).get k' h = cast (by congr; apply LawfulBEq.eq_of_beq k_beq) v := + Raw₀.get_insertMany_empty_list_of_mem k_beq distinct mem theorem get!_ofList_of_contains_eq_false [LawfulBEq α] {l : List ((a : α) × β a)} {k : α} [Inhabited (β k)] (h : (l.map Sigma.fst).contains k = false) : - (ofList l).get! k = default := by - simp only [ofList] - rw [get!_insertMany_list_of_contains_eq_false h] - apply get!_emptyc + (ofList l).get! k = default := + Raw₀.get!_insertMany_empty_list_of_contains_eq_false h theorem get!_ofList_of_mem [LawfulBEq α] {l : List ((a : α) × β a)} {k k' : α} (k_beq : k == k') {v : β k} [Inhabited (β k')] (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) (mem : ⟨k, v⟩ ∈ l) : - (ofList l).get! k' = cast (by congr; apply LawfulBEq.eq_of_beq k_beq) v := by - simp only [ofList] - rw [get!_insertMany_list_of_mem k_beq distinct mem] + (ofList l).get! k' = cast (by congr; apply LawfulBEq.eq_of_beq k_beq) v := + Raw₀.get!_insertMany_empty_list_of_mem k_beq distinct mem theorem getD_ofList_of_contains_eq_false [LawfulBEq α] {l : List ((a : α) × β a)} {k : α} {fallback : β k} (contains_eq_false : (l.map Sigma.fst).contains k = false) : - (ofList l).getD k fallback = fallback := by - simp only [ofList] - rw [getD_insertMany_list_of_contains_eq_false contains_eq_false] - apply getD_emptyc + (ofList l).getD k fallback = fallback := + Raw₀.getD_insertMany_empty_list_of_contains_eq_false contains_eq_false theorem getD_ofList_of_mem [LawfulBEq α] {l : List ((a : α) × β a)} {k k' : α} (k_beq : k == k') {v : β k} {fallback : β k'} (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) (mem : ⟨k, v⟩ ∈ l) : - (ofList l).getD k' fallback = cast (by congr; apply LawfulBEq.eq_of_beq k_beq) v := by - simp only [ofList] - rw [getD_insertMany_list_of_mem k_beq distinct mem] + (ofList l).getD k' fallback = cast (by congr; apply LawfulBEq.eq_of_beq k_beq) v := + Raw₀.getD_insertMany_empty_list_of_mem k_beq distinct mem theorem getKey?_ofList_of_contains_eq_false [EquivBEq α] [LawfulHashable α] {l : List ((a : α) × β a)} {k : α} (h : (l.map Sigma.fst).contains k = false) : - (ofList l).getKey? k = none := by - simp only [ofList] - rw [getKey?_insertMany_list_of_contains_eq_false h] - apply getKey?_emptyc + (ofList l).getKey? k = none := + Raw₀.getKey?_insertMany_empty_list_of_contains_eq_false h theorem getKey?_ofList_of_mem [EquivBEq α] [LawfulHashable α] {l : List ((a : α) × β a)} {k k' : α} (k_beq : k == k') (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) (mem : k ∈ l.map Sigma.fst) : - (ofList l).getKey? k' = some k := by - simp only [ofList] - rw [getKey?_insertMany_list_of_mem k_beq distinct mem] + (ofList l).getKey? k' = some k := + Raw₀.getKey?_insertMany_empty_list_of_mem k_beq distinct mem theorem getKey_ofList_of_mem [EquivBEq α] [LawfulHashable α] {l : List ((a : α) × β a)} @@ -1535,164 +1520,135 @@ theorem getKey_ofList_of_mem [EquivBEq α] [LawfulHashable α] (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) (mem : k ∈ l.map Sigma.fst) {h'} : - (ofList l).getKey k' h' = k := by - simp only [ofList] - rw [getKey_insertMany_list_of_mem k_beq distinct mem] + (ofList l).getKey k' h' = k := + Raw₀.getKey_insertMany_empty_list_of_mem k_beq distinct mem theorem getKey!_ofList_of_contains_eq_false [EquivBEq α] [LawfulHashable α] [Inhabited α] {l : List ((a : α) × β a)} {k : α} (h : (l.map Sigma.fst).contains k = false) : - (ofList l).getKey! k = default := by - simp only [ofList] - rw [getKey!_insertMany_list_of_contains_eq_false h] - apply getKey!_emptyc + (ofList l).getKey! k = default := + Raw₀.getKey!_insertMany_empty_list_of_contains_eq_false h theorem getKey!_ofList_of_mem [EquivBEq α] [LawfulHashable α] [Inhabited α] {l : List ((a : α) × β a)} {k k' : α} (k_beq : k == k') (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) (mem : k ∈ l.map Sigma.fst) : - (ofList l).getKey! k' = k := by - simp only [ofList] - rw [getKey!_insertMany_list_of_mem k_beq distinct mem] + (ofList l).getKey! k' = k := + Raw₀.getKey!_insertMany_empty_list_of_mem k_beq distinct mem theorem getKeyD_ofList_of_contains_eq_false [EquivBEq α] [LawfulHashable α] {l : List ((a : α) × β a)} {k fallback : α} (h : (l.map Sigma.fst).contains k = false) : - (ofList l).getKeyD k fallback = fallback := by - simp only [ofList] - rw [getKeyD_insertMany_list_of_contains_eq_false h] - apply getKeyD_emptyc + (ofList l).getKeyD k fallback = fallback := + Raw₀.getKeyD_insertMany_empty_list_of_contains_eq_false h theorem getKeyD_ofList_of_mem [EquivBEq α] [LawfulHashable α] {l : List ((a : α) × β a)} {k k' fallback : α} (k_beq : k == k') (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) (mem : k ∈ l.map Sigma.fst) : - (ofList l).getKeyD k' fallback = k := by - simp only [ofList] - rw [getKeyD_insertMany_list_of_mem k_beq distinct mem] + (ofList l).getKeyD k' fallback = k := + Raw₀.getKeyD_insertMany_empty_list_of_mem k_beq distinct mem theorem size_ofList [EquivBEq α] [LawfulHashable α] {l : List ((a : α) × β a)} (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) : - (ofList l).size = l.length := by - simp only [ofList] - rw [size_insertMany_list distinct] - · simp only [size_emptyc, Nat.zero_add] - · simp only [contains_emptyc, Bool.false_eq_true, false_implies, implies_true] + (ofList l).size = l.length := + Raw₀.size_insertMany_empty_list distinct theorem size_ofList_le [EquivBEq α] [LawfulHashable α] {l : List ((a : α) × β a)} : - (ofList l).size ≤ l.length := by - simp only [ofList] - rw [← Nat.zero_add l.length, ← size_emptyc] - apply (size_insertMany_list_le) + (ofList l).size ≤ l.length := + Raw₀.size_insertMany_empty_list_le -theorem isEmpty_ofList [EquivBEq α] [LawfulHashable α] +@[simp] +theorem isEmpty_ofList[EquivBEq α] [LawfulHashable α] {l : List ((a : α) × β a)} : - (ofList l).isEmpty = l.isEmpty := by - simp only [ofList] - rw [isEmpty_insertMany_list] - simp + (ofList l).isEmpty = l.isEmpty := + Raw₀.isEmpty_insertMany_empty_list namespace Const variable {β : Type v} - @[simp] theorem ofList_nil [EquivBEq α] [LawfulHashable α] : - ofList ([] : List (α × β)) = empty := by - simp only [ofList, insertMany_nil] - rfl + ofList ([] : List (α × β)) = ∅ := + Subtype.eq (congrArg Subtype.val (Raw₀.Const.insertMany_empty_list_nil (α:= α)) :) @[simp] theorem ofList_singleton {k : α} {v : β} [EquivBEq α] [LawfulHashable α] : - ofList [⟨k, v⟩] = empty.insert k v := by - simp only [ofList, insertMany_list_singleton] - rfl + ofList [⟨k, v⟩] = empty.insert k v := + Subtype.eq (congrArg Subtype.val (Raw₀.Const.insertMany_empty_list_singleton (α:= α)) :) -theorem ofList_cons [EquivBEq α] [LawfulHashable α] {hd : α × β} {tl : List (α × β)} : - ofList (hd::tl) = insertMany ((∅ : DHashMap α (fun _ => β)).insert hd.1 hd.2) tl := by - simp only [ofList] - rw [insertMany_cons] +theorem ofList_cons [EquivBEq α] [LawfulHashable α] {k : α} {v : β} {tl : List (α × β)} : + (insertMany empty (⟨k, v⟩ :: tl)) = (insertMany (empty.insert k v) tl) := + Subtype.eq (congrArg Subtype.val (Raw₀.Const.insertMany_empty_list_cons (α:= α)) :) +@[simp] theorem contains_ofList [EquivBEq α] [LawfulHashable α] {l : List (α × β)} {k : α} : - (ofList l).contains k = (l.map Prod.fst).contains k := by - simp only [ofList, contains_insertMany_list, contains_emptyc, Bool.false_or] + (ofList l).contains k = (l.map Prod.fst).contains k := + Raw₀.Const.contains_insertMany_empty_list theorem get?_ofList_of_contains_eq_false [LawfulBEq α] {l : List (α × β)} {k : α} (h : (l.map Prod.fst).contains k = false) : - get? (ofList l) k = none := by - simp only [ofList] - rw [get?_insertMany_list_of_contains_eq_false _ h] - apply get?_emptyc + get? (ofList l) k = none := + Raw₀.Const.get?_insertMany_empty_list_of_contains_eq_false h theorem get?_ofList_of_mem [LawfulBEq α] {l : List (α × β)} {k k' : α} (k_beq : k == k') {v : β} (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) (mem : ⟨k, v⟩ ∈ l) : - get? (ofList l) k' = some v := by - simp only [ofList] - rw [get?_insertMany_list_of_mem _ k_beq mem distinct] + get? (ofList l) k' = some v := + Raw₀.Const.get?_insertMany_empty_list_of_mem k_beq distinct mem theorem get_ofList_of_mem [LawfulBEq α] {l : List (α × β)} {k k' : α} (k_beq : k == k') {v : β} (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) (mem : ⟨k, v⟩ ∈ l) {h} : - get (ofList l) k' h = v := by - simp only [ofList] - rw [get_insertMany_list_of_mem _ k_beq mem distinct] + get (ofList l) k' h = v := + Raw₀.Const.get_insertMany_empty_list_of_mem k_beq distinct mem theorem get!_ofList_of_contains_eq_false [LawfulBEq α] {l : List (α × β)} {k : α} [Inhabited β] (h : (l.map Prod.fst).contains k = false) : - get! (ofList l) k = default := by - simp only [ofList] - rw [get!_insertMany_list_of_contains_eq_false _ h] - apply get!_emptyc + get! (ofList l) k = (default : β) := + Raw₀.Const.get!_insertMany_empty_list_of_contains_eq_false h theorem get!_ofList_of_mem [LawfulBEq α] {l : List (α × β)} {k k' : α} (k_beq : k == k') {v : β} [Inhabited β] (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) (mem : ⟨k, v⟩ ∈ l) : - get! (ofList l) k' = v := by - simp only [ofList] - rw [get!_insertMany_list_of_mem _ k_beq mem distinct] + get! (ofList l) k' = v := + Raw₀.Const.get!_insertMany_empty_list_of_mem k_beq distinct mem theorem getD_ofList_of_contains_eq_false [LawfulBEq α] {l : List (α × β)} {k : α} {fallback : β} (contains_eq_false : (l.map Prod.fst).contains k = false) : - getD (ofList l) k fallback = fallback := by - simp only [ofList] - rw [getD_insertMany_list_of_contains_eq_false _ contains_eq_false] - apply getD_emptyc + getD (ofList l) k fallback = fallback := + Raw₀.Const.getD_insertMany_empty_list_of_contains_eq_false contains_eq_false theorem getD_ofList_of_mem [LawfulBEq α] {l : List (α × β)} {k k' : α} (k_beq : k == k') {v : β} {fallback : β} (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) (mem : ⟨k, v⟩ ∈ l) : - getD (ofList l) k' fallback = v := by - simp only [ofList] - rw [getD_insertMany_list_of_mem _ k_beq mem distinct] + getD (ofList l) k' fallback = v := + Raw₀.Const.getD_insertMany_empty_list_of_mem k_beq distinct mem theorem getKey?_ofList_of_contains_eq_false [EquivBEq α] [LawfulHashable α] {l : List (α × β)} {k : α} (h : (l.map Prod.fst).contains k = false) : - (ofList l).getKey? k = none := by - simp only [ofList] - rw [getKey?_insertMany_list_of_contains_eq_false _ h] - apply getKey?_emptyc + (ofList l).getKey? k = none := + Raw₀.Const.getKey?_insertMany_empty_list_of_contains_eq_false h theorem getKey?_ofList_of_mem [EquivBEq α] [LawfulHashable α] {l : List (α × β)} {k k' : α} (k_beq : k == k') (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) (mem : k ∈ l.map Prod.fst) : - (ofList l).getKey? k' = some k := by - simp only [ofList] - rw [getKey?_insertMany_list_of_mem _ k_beq distinct mem] + (ofList l).getKey? k' = some k := + Raw₀.Const.getKey?_insertMany_empty_list_of_mem k_beq distinct mem theorem getKey_ofList_of_mem [EquivBEq α] [LawfulHashable α] {l : List (α × β)} @@ -1700,177 +1656,156 @@ theorem getKey_ofList_of_mem [EquivBEq α] [LawfulHashable α] (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) (mem : k ∈ l.map Prod.fst) {h'} : - (ofList l).getKey k' h' = k := by - simp only [ofList] - rw [getKey_insertMany_list_of_mem _ k_beq distinct mem] + (ofList l).getKey k' h' = k := + Raw₀.Const.getKey_insertMany_empty_list_of_mem k_beq distinct mem -theorem getKey!_ofList_of_contains_eq_false [EquivBEq α] [LawfulHashable α] [Inhabited α] - {l : List (α × β)} {k : α} +theorem getKey!_ofList_of_contains_eq_false [EquivBEq α] [LawfulHashable α] + [Inhabited α] {l : List (α × β)} {k : α} (h : (l.map Prod.fst).contains k = false) : - (ofList l).getKey! k = default := by - simp only [ofList] - rw [getKey!_insertMany_list_of_contains_eq_false _ h] - apply getKey!_emptyc + (ofList l).getKey! k = default := + Raw₀.Const.getKey!_insertMany_empty_list_of_contains_eq_false h theorem getKey!_ofList_of_mem [EquivBEq α] [LawfulHashable α] [Inhabited α] {l : List (α × β)} {k k' : α} (k_beq : k == k') (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) (mem : k ∈ l.map Prod.fst) : - (ofList l).getKey! k' = k := by - simp only [ofList] - rw [getKey!_insertMany_list_of_mem _ k_beq distinct mem] + (ofList l).getKey! k' = k := + Raw₀.Const.getKey!_insertMany_empty_list_of_mem k_beq distinct mem theorem getKeyD_ofList_of_contains_eq_false [EquivBEq α] [LawfulHashable α] {l : List (α × β)} {k fallback : α} (h : (l.map Prod.fst).contains k = false) : - (ofList l).getKeyD k fallback = fallback := by - simp only [ofList] - rw [getKeyD_insertMany_list_of_contains_eq_false _ h] - apply getKeyD_emptyc + (ofList l).getKeyD k fallback = fallback := + Raw₀.Const.getKeyD_insertMany_empty_list_of_contains_eq_false h theorem getKeyD_ofList_of_mem [EquivBEq α] [LawfulHashable α] {l : List (α × β)} {k k' fallback : α} (k_beq : k == k') (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) (mem : k ∈ l.map Prod.fst) : - (ofList l).getKeyD k' fallback = k := by - simp only [ofList] - rw [getKeyD_insertMany_list_of_mem _ k_beq distinct mem] + (ofList l).getKeyD k' fallback = k := + Raw₀.Const.getKeyD_insertMany_empty_list_of_mem k_beq distinct mem theorem size_ofList [EquivBEq α] [LawfulHashable α] {l : List (α × β)} (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) : - (ofList l).size = l.length := by - simp only [ofList] - rw [size_insertMany_list _ distinct] - · simp only [size_emptyc, Nat.zero_add] - · simp only [contains_emptyc, Bool.false_eq_true, false_implies, implies_true] + (ofList l).size = l.length := + Raw₀.Const.size_insertMany_empty_list distinct theorem size_ofList_le [EquivBEq α] [LawfulHashable α] {l : List (α × β)} : - (ofList l).size ≤ l.length := by - simp only [ofList] - rw [← Nat.zero_add l.length, ← size_emptyc] - apply (size_insertMany_list_le _) + (ofList l).size ≤ l.length := + Raw₀.Const.size_insertMany_empty_list_le +@[simp] theorem isEmpty_ofList [EquivBEq α] [LawfulHashable α] {l : List (α × β)} : - (ofList l).isEmpty = l.isEmpty := by - simp [ofList, isEmpty_insertMany_list _] + (ofList l).isEmpty = l.isEmpty := + Raw₀.Const.isEmpty_insertMany_empty_list @[simp] theorem unitOfList_nil [EquivBEq α] [LawfulHashable α] : - unitOfList ([] : List α) = empty := by - simp only [unitOfList, insertManyIfNewUnit_nil _] - rfl + unitOfList ([] : List α) = ∅ := + Subtype.eq (congrArg Subtype.val (Raw₀.Const.insertManyIfNewUnit_empty_list_nil (α := α)) :) @[simp] theorem unitOfList_singleton [EquivBEq α] [LawfulHashable α] {k : α} : - unitOfList [k] = empty.insertIfNew k () := by - simp only [unitOfList, insertManyIfNewUnit_list_singleton _] - rfl + unitOfList [k] = empty.insertIfNew k () := + Subtype.eq (congrArg Subtype.val (Raw₀.Const.insertManyIfNewUnit_empty_list_singleton (α := α)) :) theorem unitOfList_cons [EquivBEq α] [LawfulHashable α] {hd : α} {tl : List α} : - unitOfList (hd::tl) = insertManyIfNewUnit ((∅ : DHashMap α (fun _ => Unit)).insertIfNew hd ()) tl := by - simp only [unitOfList] - rw [insertManyIfNewUnit_cons _] + unitOfList (hd :: tl) = + (insertManyIfNewUnit (empty.insertIfNew hd ()) tl) := + Subtype.eq (congrArg Subtype.val (Raw₀.Const.insertManyIfNewUnit_empty_list_cons (α := α)) :) +@[simp] theorem contains_unitOfList [EquivBEq α] [LawfulHashable α] {l : List α} {k : α} : - (unitOfList l).contains k = l.contains k := by - simp [unitOfList, contains_insertManyIfNewUnit_list _] + (unitOfList l).contains k = l.contains k := + Raw₀.Const.contains_insertManyIfNewUnit_empty_list theorem getKey?_unitOfList_of_contains_eq_false [EquivBEq α] [LawfulHashable α] {l : List α} {k : α} (h' : l.contains k = false) : - getKey? (unitOfList l) k = none := by - simp [unitOfList, getKey?_insertManyIfNewUnit_list_of_contains_eq_false _ h'] + getKey? (unitOfList l) k = none := + Raw₀.Const.getKey?_insertManyIfNewUnit_empty_list_of_contains_eq_false h' theorem getKey?_unitOfList_of_mem [EquivBEq α] [LawfulHashable α] {l : List α} {k k' : α} (k_beq : k == k') (distinct : l.Pairwise (fun a b => (a == b) = false)) (mem : k ∈ l) : - getKey? (unitOfList l) k' = some k := by - simp [unitOfList, getKey?_insertManyIfNewUnit_list_of_mem_of_contains_eq_false _ k_beq - distinct mem contains_emptyc] + getKey? (unitOfList l) k' = some k := + Raw₀.Const.getKey?_insertManyIfNewUnit_empty_list_of_mem k_beq distinct mem theorem getKey_unitOfList_of_mem [EquivBEq α] [LawfulHashable α] {l : List α} {k k' : α} (k_beq : k == k') (distinct : l.Pairwise (fun a b => (a == b) = false)) (mem : k ∈ l) {h'} : - getKey (unitOfList l) k' h' = k := by - simp [unitOfList, getKey_insertManyIfNewUnit_list_of_mem_of_contains_eq_false _ k_beq - distinct mem contains_emptyc] + getKey (unitOfList l) k' h' = k := + Raw₀.Const.getKey_insertManyIfNewUnit_empty_list_of_mem k_beq distinct mem theorem getKey!_unitOfList_of_contains_eq_false [EquivBEq α] [LawfulHashable α] [Inhabited α] {l : List α} {k : α} (h' : l.contains k = false) : - getKey! (unitOfList l) k = default := by - simp [unitOfList, getKey!_insertManyIfNewUnit_list_of_contains_eq_false _ h'] + getKey! (unitOfList l) k = default := + Raw₀.Const.getKey!_insertManyIfNewUnit_empty_list_of_contains_eq_false h' theorem getKey!_unitOfList_of_mem [EquivBEq α] [LawfulHashable α] [Inhabited α] {l : List α} {k k' : α} (k_beq : k == k') (distinct : l.Pairwise (fun a b => (a == b) = false)) (mem : k ∈ l) : - getKey! (unitOfList l) k' = k := by - simp [unitOfList, getKey!_insertManyIfNewUnit_list_of_mem_of_contains_eq_false _ k_beq - distinct mem contains_emptyc] + getKey! (unitOfList l) k' = k := + Raw₀.Const.getKey!_insertManyIfNewUnit_empty_list_of_mem k_beq distinct mem theorem getKeyD_unitOfList_of_contains_eq_false [EquivBEq α] [LawfulHashable α] {l : List α} {k fallback : α} (h' : l.contains k = false) : - getKeyD (unitOfList l) k fallback = fallback := by - simp [unitOfList, getKeyD_insertManyIfNewUnit_list_of_contains_eq_false _ h'] + getKeyD (unitOfList l) k fallback = fallback := + Raw₀.Const.getKeyD_insertManyIfNewUnit_empty_list_of_contains_eq_false h' theorem getKeyD_unitOfList_of_mem [EquivBEq α] [LawfulHashable α] {l : List α} {k k' fallback : α} (k_beq : k == k') (distinct : l.Pairwise (fun a b => (a == b) = false)) - {mem : k ∈ l } : - getKeyD (unitOfList l) k' fallback = k := by - simp [unitOfList, getKeyD_insertManyIfNewUnit_list_of_mem_of_contains_eq_false _ k_beq - distinct mem contains_emptyc] + (mem : k ∈ l) : + getKeyD (unitOfList l) k' fallback = k := + Raw₀.Const.getKeyD_insertManyIfNewUnit_empty_list_of_mem k_beq distinct mem theorem size_unitOfList [EquivBEq α] [LawfulHashable α] {l : List α} (distinct : l.Pairwise (fun a b => (a == b) = false)) : - (unitOfList l).size = l.length := by - simp only [unitOfList] - rw [size_insertManyIfNewUnit_list _ distinct] - · simp - · simp + (unitOfList l).size = l.length := + Raw₀.Const.size_insertManyIfNewUnit_empty_list distinct theorem size_unitOfList_le [EquivBEq α] [LawfulHashable α] {l : List α} : - (unitOfList l).size ≤ l.length := by - simp only [unitOfList] - apply Nat.le_trans (size_insertManyIfNewUnit_list_le _) - simp - + (unitOfList l).size ≤ l.length := + Raw₀.Const.size_insertManyIfNewUnit_empty_list_le +@[simp] theorem isEmpty_unitOfList [EquivBEq α] [LawfulHashable α] {l : List α} : - (unitOfList l).isEmpty = l.isEmpty := by - simp only [unitOfList] - rw [isEmpty_insertManyIfNewUnit_list _] - simp + (unitOfList l).isEmpty = l.isEmpty := + Raw₀.Const.isEmpty_insertManyIfNewUnit_empty_list +@[simp] theorem get?_unitOfList [EquivBEq α] [LawfulHashable α] {l : List α} {k : α} : get? (unitOfList l) k = - if l.contains k then some () else none := by - simp only [unitOfList] - rw [get?_insertManyIfNewUnit_list _] - simp + if l.contains k then some () else none := + Raw₀.Const.get?_insertManyIfNewUnit_empty_list +@[simp] theorem get_unitOfList {l : List α} {k : α} {h} : - get (unitOfList l) k h = () := by - simp + get (unitOfList l) k h = () := + Raw₀.Const.get_insertManyIfNewUnit_empty_list +@[simp] theorem get!_unitOfList {l : List α} {k : α} : - get! (unitOfList l) k = () := by - simp + get! (unitOfList l) k = () := + Raw₀.Const.get!_insertManyIfNewUnit_empty_list +@[simp] theorem getD_unitOfList {l : List α} {k : α} {fallback : Unit} : getD (unitOfList l) k fallback = () := by From 5f9dd86a86753408b7e903fdf81e701e4170d790 Mon Sep 17 00:00:00 2001 From: jt0202 Date: Sun, 5 Jan 2025 22:35:12 +0100 Subject: [PATCH 61/83] Style fixes --- src/Std/Data/DHashMap/Lemmas.lean | 214 +++++++++++++++--------------- 1 file changed, 105 insertions(+), 109 deletions(-) diff --git a/src/Std/Data/DHashMap/Lemmas.lean b/src/Std/Data/DHashMap/Lemmas.lean index fc72c5ba33a5..70cdaeb52349 100644 --- a/src/Std/Data/DHashMap/Lemmas.lean +++ b/src/Std/Data/DHashMap/Lemmas.lean @@ -975,7 +975,7 @@ theorem insertMany_list_singleton {k : α} {v : β k} : Subtype.eq (congrArg Subtype.val (Raw₀.insertMany_list_singleton ⟨m.1, m.2.size_buckets_pos⟩) :) theorem insertMany_cons {l : List ((a : α) × β a)} {k : α} {v : β k} : - (m.insertMany (⟨k, v⟩ :: l)) = ((m.insert k v).insertMany l) := + m.insertMany (⟨k, v⟩ :: l) = (m.insert k v).insertMany l := Subtype.eq (congrArg Subtype.val (Raw₀.insertMany_cons ⟨m.1, m.2.size_buckets_pos⟩) :) @[simp] @@ -987,20 +987,20 @@ theorem contains_insertMany_list [EquivBEq α] [LawfulHashable α] theorem contains_of_contains_insertMany_list [EquivBEq α] [LawfulHashable α] {l : List ((a : α) × β a)} {k : α} : (m.insertMany l).contains k → (l.map Sigma.fst).contains k = false → m.contains k := - Raw₀.contains_of_contains_insertMany_list ⟨m.1, _⟩ m.2 + Raw₀.contains_of_contains_insertMany_list ⟨m.1, _⟩ m.2 theorem get?_insertMany_list_of_contains_eq_false [LawfulBEq α] {l : List ((a : α) × β a)} {k : α} (h' : (l.map Sigma.fst).contains k = false) : (m.insertMany l).get? k = m.get? k := - Raw₀.get?_insertMany_list_of_contains_eq_false ⟨m.1, _⟩ m.2 h' + Raw₀.get?_insertMany_list_of_contains_eq_false ⟨m.1, _⟩ m.2 h' theorem get?_insertMany_list_of_mem [LawfulBEq α] {l : List ((a : α) × β a)} {k k' : α} (k_beq : k == k') {v : β k} (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) (mem : ⟨k, v⟩ ∈ l) : (m.insertMany l).get? k' = some (cast (by congr; apply LawfulBEq.eq_of_beq k_beq) v) := - Raw₀.get?_insertMany_list_of_mem ⟨m.1, _⟩ m.2 k_beq distinct mem + Raw₀.get?_insertMany_list_of_mem ⟨m.1, _⟩ m.2 k_beq distinct mem theorem get_insertMany_list_of_contains_eq_false [LawfulBEq α] {l : List ((a : α) × β a)} {k : α} @@ -1008,7 +1008,7 @@ theorem get_insertMany_list_of_contains_eq_false [LawfulBEq α] {h'} : (m.insertMany l).get k h' = m.get k (contains_of_contains_insertMany_list h' contains) := - Raw₀.get_insertMany_list_of_contains_eq_false ⟨m.1, _⟩ m.2 contains + Raw₀.get_insertMany_list_of_contains_eq_false ⟨m.1, _⟩ m.2 contains theorem get_insertMany_list_of_mem [LawfulBEq α] {l : List ((a : α) × β a)} {k k' : α} (k_beq : k == k') {v : β k} @@ -1016,39 +1016,39 @@ theorem get_insertMany_list_of_mem [LawfulBEq α] (mem : ⟨k, v⟩ ∈ l) {h'} : (m.insertMany l).get k' h' = cast (by congr; apply LawfulBEq.eq_of_beq k_beq) v := - Raw₀.get_insertMany_list_of_mem ⟨m.1, _⟩ m.2 k_beq distinct mem + Raw₀.get_insertMany_list_of_mem ⟨m.1, _⟩ m.2 k_beq distinct mem theorem get!_insertMany_list_of_contains_eq_false [LawfulBEq α] {l : List ((a : α) × β a)} {k : α} [Inhabited (β k)] (h' : (l.map Sigma.fst).contains k = false) : (m.insertMany l).get! k = m.get! k := - Raw₀.get!_insertMany_list_of_contains_eq_false ⟨m.1, _⟩ m.2 h' + Raw₀.get!_insertMany_list_of_contains_eq_false ⟨m.1, _⟩ m.2 h' theorem get!_insertMany_list_of_mem [LawfulBEq α] {l : List ((a : α) × β a)} {k k' : α} (k_beq : k == k') {v : β k} [Inhabited (β k')] (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) (mem : ⟨k, v⟩ ∈ l) : (m.insertMany l).get! k' = cast (by congr; apply LawfulBEq.eq_of_beq k_beq) v := - Raw₀.get!_insertMany_list_of_mem ⟨m.1, _⟩ m.2 k_beq distinct mem + Raw₀.get!_insertMany_list_of_mem ⟨m.1, _⟩ m.2 k_beq distinct mem theorem getD_insertMany_list_of_contains_eq_false [LawfulBEq α] {l : List ((a : α) × β a)} {k : α} {fallback : β k} (not_mem : (l.map Sigma.fst).contains k = false) : (m.insertMany l).getD k fallback = m.getD k fallback := - Raw₀.getD_insertMany_list_of_contains_eq_false ⟨m.1, _⟩ m.2 not_mem + Raw₀.getD_insertMany_list_of_contains_eq_false ⟨m.1, _⟩ m.2 not_mem theorem getD_insertMany_list_of_mem [LawfulBEq α] {l : List ((a : α) × β a)} {k k' : α} (k_beq : k == k') {v : β k} {fallback : β k'} (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) (mem : ⟨k, v⟩ ∈ l) : (m.insertMany l).getD k' fallback = cast (by congr; apply LawfulBEq.eq_of_beq k_beq) v := - Raw₀.getD_insertMany_list_of_mem ⟨m.1, _⟩ m.2 k_beq distinct mem + Raw₀.getD_insertMany_list_of_mem ⟨m.1, _⟩ m.2 k_beq distinct mem theorem getKey?_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] {l : List ((a : α) × β a)} {k : α} (h' : (l.map Sigma.fst).contains k = false) : (m.insertMany l).getKey? k = m.getKey? k := - Raw₀.getKey?_insertMany_list_of_contains_eq_false ⟨m.1, _⟩ m.2 h' + Raw₀.getKey?_insertMany_list_of_contains_eq_false ⟨m.1, _⟩ m.2 h' theorem getKey?_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] {l : List ((a : α) × β a)} @@ -1056,7 +1056,7 @@ theorem getKey?_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) (mem : k ∈ l.map Sigma.fst) : (m.insertMany l).getKey? k' = some k := - Raw₀.getKey?_insertMany_list_of_mem ⟨m.1, _⟩ m.2 k_beq distinct mem + Raw₀.getKey?_insertMany_list_of_mem ⟨m.1, _⟩ m.2 k_beq distinct mem theorem getKey_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] {l : List ((a : α) × β a)} {k : α} @@ -1064,7 +1064,7 @@ theorem getKey_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashabl {h'} : (m.insertMany l).getKey k h' = m.getKey k (contains_of_contains_insertMany_list h' h₁) := - Raw₀.getKey_insertMany_list_of_contains_eq_false ⟨m.1, _⟩ m.2 h₁ + Raw₀.getKey_insertMany_list_of_contains_eq_false ⟨m.1, _⟩ m.2 h₁ theorem getKey_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] {l : List ((a : α) × β a)} @@ -1073,13 +1073,13 @@ theorem getKey_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] (mem : k ∈ l.map Sigma.fst) {h'} : (m.insertMany l).getKey k' h' = k := - Raw₀.getKey_insertMany_list_of_mem ⟨m.1, _⟩ m.2 k_beq distinct mem + Raw₀.getKey_insertMany_list_of_mem ⟨m.1, _⟩ m.2 k_beq distinct mem theorem getKey!_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] [Inhabited α] {l : List ((a : α) × β a)} {k : α} (h' : (l.map Sigma.fst).contains k = false) : (m.insertMany l).getKey! k = m.getKey! k := - Raw₀.getKey!_insertMany_list_of_contains_eq_false ⟨m.1, _⟩ m.2 h' + Raw₀.getKey!_insertMany_list_of_contains_eq_false ⟨m.1, _⟩ m.2 h' theorem getKey!_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] [Inhabited α] {l : List ((a : α) × β a)} @@ -1087,13 +1087,13 @@ theorem getKey!_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] [Inhabi (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) (mem : k ∈ l.map Sigma.fst) : (m.insertMany l).getKey! k' = k := - Raw₀.getKey!_insertMany_list_of_mem ⟨m.1, _⟩ m.2 k_beq distinct mem + Raw₀.getKey!_insertMany_list_of_mem ⟨m.1, _⟩ m.2 k_beq distinct mem theorem getKeyD_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] {l : List ((a : α) × β a)} {k fallback : α} (h' : (l.map Sigma.fst).contains k = false) : (m.insertMany l).getKeyD k fallback = m.getKeyD k fallback := - Raw₀.getKeyD_insertMany_list_of_contains_eq_false ⟨m.1, _⟩ m.2 h' + Raw₀.getKeyD_insertMany_list_of_contains_eq_false ⟨m.1, _⟩ m.2 h' theorem getKeyD_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] {l : List ((a : α) × β a)} @@ -1101,24 +1101,24 @@ theorem getKeyD_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) (mem : k ∈ l.map Sigma.fst) : (m.insertMany l).getKeyD k' fallback = k := - Raw₀.getKeyD_insertMany_list_of_mem ⟨m.1, _⟩ m.2 k_beq distinct mem + Raw₀.getKeyD_insertMany_list_of_mem ⟨m.1, _⟩ m.2 k_beq distinct mem theorem size_insertMany_list [EquivBEq α] [LawfulHashable α] {l : List ((a : α) × β a)} (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) : (∀ (a : α), m.contains a → (l.map Sigma.fst).contains a = false) → (m.insertMany l).size = m.size + l.length := - Raw₀.size_insertMany_list ⟨m.1, _⟩ m.2 distinct + Raw₀.size_insertMany_list ⟨m.1, _⟩ m.2 distinct theorem size_insertMany_list_le [EquivBEq α] [LawfulHashable α] {l : List ((a : α) × β a)} : (m.insertMany l).size ≤ m.size + l.length := - Raw₀.size_insertMany_list_le ⟨m.1, _⟩ m.2 + Raw₀.size_insertMany_list_le ⟨m.1, _⟩ m.2 @[simp] theorem isEmpty_insertMany_list [EquivBEq α] [LawfulHashable α] {l : List ((a : α) × β a)} : (m.insertMany l).isEmpty = (m.isEmpty && l.isEmpty) := - Raw₀.isEmpty_insertMany_list ⟨m.1, _⟩ m.2 + Raw₀.isEmpty_insertMany_list ⟨m.1, _⟩ m.2 namespace Const variable {β : Type v} (m : DHashMap α (fun _ => β)) @@ -1134,33 +1134,33 @@ theorem insertMany_list_singleton {k : α} {v : β} : Subtype.eq (congrArg Subtype.val (Raw₀.Const.insertMany_list_singleton ⟨m.1, m.2.size_buckets_pos⟩) :) theorem insertMany_cons {l : List (α × β)} {k : α} {v : β} : - (insertMany m (⟨k, v⟩ :: l)) = (insertMany (m.insert k v) l) := + insertMany m (⟨k, v⟩ :: l) = insertMany (m.insert k v) l := Subtype.eq (congrArg Subtype.val (Raw₀.Const.insertMany_cons ⟨m.1, m.2.size_buckets_pos⟩) :) @[simp] theorem contains_insertMany_list [EquivBEq α] [LawfulHashable α] {l : List (α × β)} {k : α} : (Const.insertMany m l).contains k = (m.contains k || (l.map Prod.fst).contains k) := - Raw₀.Const.contains_insertMany_list ⟨m.1, _⟩ m.2 + Raw₀.Const.contains_insertMany_list ⟨m.1, _⟩ m.2 theorem contains_of_contains_insertMany_list [EquivBEq α] [LawfulHashable α] {l : List ( α × β )} {k : α} : (insertMany m l).contains k → (l.map Prod.fst).contains k = false → m.contains k := - Raw₀.Const.contains_of_contains_insertMany_list ⟨m.1, _⟩ m.2 + Raw₀.Const.contains_of_contains_insertMany_list ⟨m.1, _⟩ m.2 theorem getKey?_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] {l : List (α × β)} {k : α} (h' : (l.map Prod.fst).contains k = false) : (insertMany m l).getKey? k = m.getKey? k := - Raw₀.Const.getKey?_insertMany_list_of_contains_eq_false ⟨m.1, _⟩ m.2 h' + Raw₀.Const.getKey?_insertMany_list_of_contains_eq_false ⟨m.1, _⟩ m.2 h' theorem getKey?_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] - {l : List (α × β)} + {l : List (α × β)} {k k' : α} (k_beq : k == k') (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) (mem : k ∈ l.map Prod.fst) : (insertMany m l).getKey? k' = some k := - Raw₀.Const.getKey?_insertMany_list_of_mem ⟨m.1, _⟩ m.2 k_beq distinct mem + Raw₀.Const.getKey?_insertMany_list_of_mem ⟨m.1, _⟩ m.2 k_beq distinct mem theorem getKey_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] {l : List (α × β)} {k : α} @@ -1168,7 +1168,7 @@ theorem getKey_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashabl {h'} : (insertMany m l).getKey k h' = m.getKey k (contains_of_contains_insertMany_list _ h' h₁) := - Raw₀.Const.getKey_insertMany_list_of_contains_eq_false ⟨m.1, _⟩ m.2 h₁ + Raw₀.Const.getKey_insertMany_list_of_contains_eq_false ⟨m.1, _⟩ m.2 h₁ theorem getKey_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] {l : List (α × β)} @@ -1177,13 +1177,13 @@ theorem getKey_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] (mem : k ∈ l.map Prod.fst) {h'} : (insertMany m l).getKey k' h' = k := - Raw₀.Const.getKey_insertMany_list_of_mem ⟨m.1, _⟩ m.2 k_beq distinct mem + Raw₀.Const.getKey_insertMany_list_of_mem ⟨m.1, _⟩ m.2 k_beq distinct mem theorem getKey!_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] [Inhabited α] {l : List (α × β)} {k : α} (h' : (l.map Prod.fst).contains k = false) : (insertMany m l).getKey! k = m.getKey! k := - Raw₀.Const.getKey!_insertMany_list_of_contains_eq_false ⟨m.1, _⟩ m.2 h' + Raw₀.Const.getKey!_insertMany_list_of_contains_eq_false ⟨m.1, _⟩ m.2 h' theorem getKey!_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] [Inhabited α] {l : List (α × β)} @@ -1191,13 +1191,13 @@ theorem getKey!_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] [Inhabi (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) (mem : k ∈ l.map Prod.fst) : (insertMany m l).getKey! k' = k := - Raw₀.Const.getKey!_insertMany_list_of_mem ⟨m.1, _⟩ m.2 k_beq distinct mem + Raw₀.Const.getKey!_insertMany_list_of_mem ⟨m.1, _⟩ m.2 k_beq distinct mem theorem getKeyD_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] {l : List (α × β)} {k fallback : α} (h' : (l.map Prod.fst).contains k = false) : (insertMany m l).getKeyD k fallback = m.getKeyD k fallback := - Raw₀.Const.getKeyD_insertMany_list_of_contains_eq_false ⟨m.1, _⟩ m.2 h' + Raw₀.Const.getKeyD_insertMany_list_of_contains_eq_false ⟨m.1, _⟩ m.2 h' theorem getKeyD_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] {l : List (α × β)} @@ -1205,131 +1205,127 @@ theorem getKeyD_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) (mem : k ∈ l.map Prod.fst) : (insertMany m l).getKeyD k' fallback = k := - Raw₀.Const.getKeyD_insertMany_list_of_mem ⟨m.1, _⟩ m.2 k_beq distinct mem + Raw₀.Const.getKeyD_insertMany_list_of_mem ⟨m.1, _⟩ m.2 k_beq distinct mem theorem size_insertMany_list [EquivBEq α] [LawfulHashable α] {l : List (α × β)} (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) : (∀ (a : α), m.contains a → (l.map Prod.fst).contains a = false) → (insertMany m l).size = m.size + l.length := - Raw₀.Const.size_insertMany_list ⟨m.1, _⟩ m.2 distinct + Raw₀.Const.size_insertMany_list ⟨m.1, _⟩ m.2 distinct theorem size_insertMany_list_le [EquivBEq α] [LawfulHashable α] {l : List (α × β)} : (insertMany m l).size ≤ m.size + l.length := - Raw₀.Const.size_insertMany_list_le ⟨m.1, _⟩ m.2 + Raw₀.Const.size_insertMany_list_le ⟨m.1, _⟩ m.2 @[simp] theorem isEmpty_insertMany_list [EquivBEq α] [LawfulHashable α] {l : List (α × β)} : (insertMany m l).isEmpty = (m.isEmpty && l.isEmpty) := - Raw₀.Const.isEmpty_insertMany_list ⟨m.1, _⟩ m.2 + Raw₀.Const.isEmpty_insertMany_list ⟨m.1, _⟩ m.2 theorem get?_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] {l : List (α × β)} {k : α} (h' : (l.map Prod.fst).contains k = false) : get? (insertMany m l) k = get? m k := - Raw₀.Const.get?_insertMany_list_of_contains_eq_false ⟨m.1, _⟩ m.2 h' + Raw₀.Const.get?_insertMany_list_of_contains_eq_false ⟨m.1, _⟩ m.2 h' theorem get?_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] {l : List (α × β)} {k k' : α} (k_beq : k == k') {v : β} (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) (mem : ⟨k, v⟩ ∈ l) : get? (insertMany m l) k' = v := - Raw₀.Const.get?_insertMany_list_of_mem ⟨m.1, _⟩ m.2 k_beq distinct mem + Raw₀.Const.get?_insertMany_list_of_mem ⟨m.1, _⟩ m.2 k_beq distinct mem theorem get_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] {l : List (α × β)} {k : α} (h₁ : (l.map Prod.fst).contains k = false) {h'} : get (insertMany m l) k h' = get m k (contains_of_contains_insertMany_list _ h' h₁) := - Raw₀.Const.get_insertMany_list_of_contains_eq_false ⟨m.1, _⟩ m.2 h₁ + Raw₀.Const.get_insertMany_list_of_contains_eq_false ⟨m.1, _⟩ m.2 h₁ theorem get_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] {l : List (α × β)} {k k' : α} (k_beq : k == k') {v : β} (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) (mem : ⟨k, v⟩ ∈ l) {h'} : get (insertMany m l) k' h' = v := - Raw₀.Const.get_insertMany_list_of_mem ⟨m.1, _⟩ m.2 k_beq distinct mem + Raw₀.Const.get_insertMany_list_of_mem ⟨m.1, _⟩ m.2 k_beq distinct mem theorem get!_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] - [Inhabited β] {l : List (α × β)} {k : α} + [Inhabited β] {l : List (α × β)} {k : α} (h' : (l.map Prod.fst).contains k = false) : get! (insertMany m l) k = get! m k := - Raw₀.Const.get!_insertMany_list_of_contains_eq_false ⟨m.1, _⟩ m.2 h' + Raw₀.Const.get!_insertMany_list_of_contains_eq_false ⟨m.1, _⟩ m.2 h' theorem get!_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] [Inhabited β] {l : List (α × β)} {k k' : α} (k_beq : k == k') {v : β} (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) (mem : ⟨k, v⟩ ∈ l): get! (insertMany m l) k' = v := - Raw₀.Const.get!_insertMany_list_of_mem ⟨m.1, _⟩ m.2 k_beq distinct mem + Raw₀.Const.get!_insertMany_list_of_mem ⟨m.1, _⟩ m.2 k_beq distinct mem theorem getD_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] {l : List (α × β)} {k : α} {fallback : β} (h' : (l.map Prod.fst).contains k = false) : getD (insertMany m l) k fallback = getD m k fallback := - Raw₀.Const.getD_insertMany_list_of_contains_eq_false ⟨m.1, _⟩ m.2 h' + Raw₀.Const.getD_insertMany_list_of_contains_eq_false ⟨m.1, _⟩ m.2 h' theorem getD_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] {l : List (α × β)} {k k' : α} (k_beq : k == k') {v fallback : β} (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) (mem : ⟨k, v⟩ ∈ l) : getD (insertMany m l) k' fallback = v := - Raw₀.Const.getD_insertMany_list_of_mem ⟨m.1, _⟩ m.2 k_beq distinct mem + Raw₀.Const.getD_insertMany_list_of_mem ⟨m.1, _⟩ m.2 k_beq distinct mem variable (m : DHashMap α (fun _ => Unit)) - @[simp] theorem insertManyIfNewUnit_nil : insertManyIfNewUnit m [] = m := Subtype.eq (congrArg Subtype.val (Raw₀.Const.insertManyIfNewUnit_nil ⟨m.1, m.2.size_buckets_pos⟩) :) - @[simp] theorem insertManyIfNewUnit_list_singleton [EquivBEq α] [LawfulHashable α] {k : α} : insertManyIfNewUnit m [k] = m.insertIfNew k () := Subtype.eq (congrArg Subtype.val (Raw₀.Const.insertManyIfNewUnit_list_singleton ⟨m.1, m.2.size_buckets_pos⟩) :) - theorem insertManyIfNewUnit_cons {l : List α} {k : α} : - (insertManyIfNewUnit m (k :: l)) = (insertManyIfNewUnit (m.insertIfNew k ()) l) := + insertManyIfNewUnit m (k :: l) = insertManyIfNewUnit (m.insertIfNew k ()) l := Subtype.eq (congrArg Subtype.val (Raw₀.Const.insertManyIfNewUnit_cons ⟨m.1, m.2.size_buckets_pos⟩) :) - @[simp] theorem contains_insertManyIfNewUnit_list [EquivBEq α] [LawfulHashable α] {l : List α} {k : α} : (insertManyIfNewUnit m l).contains k = (m.contains k || l.contains k) := - Raw₀.Const.contains_insertManyIfNewUnit_list ⟨m.1, _⟩ m.2 + Raw₀.Const.contains_insertManyIfNewUnit_list ⟨m.1, _⟩ m.2 theorem contains_of_contains_insertManyIfNewUnit_list [EquivBEq α] [LawfulHashable α] {l : List α} {k : α} (h₂ : l.contains k = false) : (insertManyIfNewUnit m l).contains k → m.contains k := - Raw₀.Const.contains_of_contains_insertManyIfNewUnit_list ⟨m.1, _⟩ m.2 h₂ + Raw₀.Const.contains_of_contains_insertManyIfNewUnit_list ⟨m.1, _⟩ m.2 h₂ theorem getKey?_insertManyIfNewUnit_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] {l : List α} {k : α} (h' : l.contains k = false) : getKey? (insertManyIfNewUnit m l) k = getKey? m k := - Raw₀.Const.getKey?_insertManyIfNewUnit_list_of_contains_eq_false ⟨m.1, _⟩ m.2 h' + Raw₀.Const.getKey?_insertManyIfNewUnit_list_of_contains_eq_false ⟨m.1, _⟩ m.2 h' theorem getKey?_insertManyIfNewUnit_list_of_mem_of_contains_eq_false [EquivBEq α] [LawfulHashable α] {l : List α} {k k' : α} (k_beq : k == k') (distinct : l.Pairwise (fun a b => (a == b) = false)) (mem : k ∈ l) : m.contains k = false → getKey? (insertManyIfNewUnit m l) k' = some k := Raw₀.Const.getKey?_insertManyIfNewUnit_list_of_mem_of_contains_eq_false ⟨m.1, _⟩ - m.2 k_beq distinct mem + m.2 k_beq distinct mem theorem getKey?_insertManyIfNewUnit_list_of_contains_of_contains [EquivBEq α] [LawfulHashable α] {l : List α} {k : α} (distinct : l.Pairwise (fun a b => (a == b) = false)) (h' : l.contains k) : m.contains k → getKey? (insertManyIfNewUnit m l) k = getKey? m k := - Raw₀.Const.getKey?_insertManyIfNewUnit_list_of_contains_of_contains ⟨m.1, _⟩ m.2 distinct h' + Raw₀.Const.getKey?_insertManyIfNewUnit_list_of_contains_of_contains ⟨m.1, _⟩ m.2 distinct h' theorem getKey_insertManyIfNewUnit_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] {l : List α} {k : α} (h₁ : l.contains k = false) {h'} : getKey (insertManyIfNewUnit m l) k h' = getKey m k (contains_of_contains_insertManyIfNewUnit_list m h₁ h') := - Raw₀.Const.getKey_insertManyIfNewUnit_list_of_contains_eq_false ⟨m.1, _⟩ m.2 h₁ + Raw₀.Const.getKey_insertManyIfNewUnit_list_of_contains_eq_false ⟨m.1, _⟩ m.2 h₁ theorem getKey_insertManyIfNewUnit_list_of_mem_of_contains_eq_false [EquivBEq α] [LawfulHashable α] {l : List α} @@ -1338,90 +1334,90 @@ theorem getKey_insertManyIfNewUnit_list_of_mem_of_contains_eq_false [EquivBEq α (mem : k ∈ l) {h'} : m.contains k = false → getKey (insertManyIfNewUnit m l) k' h' = k := - Raw₀.Const.getKey_insertManyIfNewUnit_list_of_mem_of_contains_eq_false ⟨m.1, _⟩ m.2 k_beq - distinct mem + Raw₀.Const.getKey_insertManyIfNewUnit_list_of_mem_of_contains_eq_false ⟨m.1, _⟩ m.2 k_beq + distinct mem theorem getKey_insertManyIfNewUnit_list_mem_of_contains_of_contains [EquivBEq α] [LawfulHashable α] {l : List α} {k : α} (distinct : l.Pairwise (fun a b => (a == b) = false)) (h₁ : l.contains k) {mem' : m.contains k} {h'} : getKey (insertManyIfNewUnit m l) k h' = getKey m k mem' := - Raw₀.Const.getKey_insertManyIfNewUnit_list_mem_of_contains_of_contains ⟨m.1, _⟩ m.2 distinct h₁ + Raw₀.Const.getKey_insertManyIfNewUnit_list_mem_of_contains_of_contains ⟨m.1, _⟩ m.2 distinct h₁ theorem getKey!_insertManyIfNewUnit_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] - [Inhabited α] {l : List α} {k : α} + [Inhabited α] {l : List α} {k : α} (h' : l.contains k = false) : getKey! (insertManyIfNewUnit m l) k = getKey! m k := - Raw₀.Const.getKey!_insertManyIfNewUnit_list_of_contains_eq_false ⟨m.1, _⟩ m.2 h' + Raw₀.Const.getKey!_insertManyIfNewUnit_list_of_contains_eq_false ⟨m.1, _⟩ m.2 h' theorem getKey!_insertManyIfNewUnit_list_of_mem_of_contains_eq_false [EquivBEq α] [LawfulHashable α] - [Inhabited α] {l : List α} {k k' : α} (k_beq : k == k') + [Inhabited α] {l : List α} {k k' : α} (k_beq : k == k') (distinct : l.Pairwise (fun a b => (a == b) = false)) (mem : k ∈ l) : contains m k = false → getKey! (insertManyIfNewUnit m l) k' = k := - Raw₀.Const.getKey!_insertManyIfNewUnit_list_of_mem_of_contains_eq_false ⟨m.1, _⟩ m.2 k_beq - distinct mem + Raw₀.Const.getKey!_insertManyIfNewUnit_list_of_mem_of_contains_eq_false ⟨m.1, _⟩ m.2 k_beq + distinct mem theorem getKey!_insertManyIfNewUnit_list_mem_of_contains_of_contains [EquivBEq α] [LawfulHashable α] - [Inhabited α] {l : List α} {k : α} + [Inhabited α] {l : List α} {k : α} (distinct : l.Pairwise (fun a b => (a == b) = false)) (h' : l.contains k) : m.contains k → - getKey! (insertManyIfNewUnit m l) k = getKey! m k := - Raw₀.Const.getKey!_insertManyIfNewUnit_list_mem_of_contains_of_contains ⟨m.1, _⟩ m.2 distinct h' + getKey! (insertManyIfNewUnit m l) k = getKey! m k := + Raw₀.Const.getKey!_insertManyIfNewUnit_list_mem_of_contains_of_contains ⟨m.1, _⟩ m.2 distinct h' theorem getKeyD_insertManyIfNewUnit_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] {l : List α} {k fallback : α} (h' : l.contains k = false) : getKeyD (insertManyIfNewUnit m l) k fallback = getKeyD m k fallback := - Raw₀.Const.getKeyD_insertManyIfNewUnit_list_of_contains_eq_false ⟨m.1, _⟩ m.2 h' + Raw₀.Const.getKeyD_insertManyIfNewUnit_list_of_contains_eq_false ⟨m.1, _⟩ m.2 h' theorem getKeyD_insertManyIfNewUnit_list_of_mem_of_contains_eq_false [EquivBEq α] [LawfulHashable α] {l : List α} {k k' fallback : α} (k_beq : k == k') (distinct : l.Pairwise (fun a b => (a == b) = false)) (mem : k ∈ l ) : m.contains k = false → getKeyD (insertManyIfNewUnit m l) k' fallback = k := - Raw₀.Const.getKeyD_insertManyIfNewUnit_list_of_mem_of_contains_eq_false ⟨m.1, _⟩ m.2 k_beq - distinct mem + Raw₀.Const.getKeyD_insertManyIfNewUnit_list_of_mem_of_contains_eq_false ⟨m.1, _⟩ m.2 k_beq + distinct mem theorem getKeyD_insertManyIfNewUnit_list_of_contains_of_contains [EquivBEq α] [LawfulHashable α] {l : List α} {k fallback : α} (distinct : l.Pairwise (fun a b => (a == b) = false)) (h' : l.contains k) : m.contains k → getKeyD (insertManyIfNewUnit m l) k fallback = getKeyD m k fallback := - Raw₀.Const.getKeyD_insertManyIfNewUnit_list_of_contains_of_contains ⟨m.1, _⟩ m.2 distinct h' + Raw₀.Const.getKeyD_insertManyIfNewUnit_list_of_contains_of_contains ⟨m.1, _⟩ m.2 distinct h' theorem size_insertManyIfNewUnit_list [EquivBEq α] [LawfulHashable α] {l : List α} (distinct : l.Pairwise (fun a b => (a == b) = false)) : (∀ (a : α), m.contains a → l.contains a = false) → (insertManyIfNewUnit m l).size = m.size + l.length := - Raw₀.Const.size_insertManyIfNewUnit_list ⟨m.1, _⟩ m.2 distinct + Raw₀.Const.size_insertManyIfNewUnit_list ⟨m.1, _⟩ m.2 distinct theorem size_insertManyIfNewUnit_list_le [EquivBEq α] [LawfulHashable α] {l : List α} : (insertManyIfNewUnit m l).size ≤ m.size + l.length := - Raw₀.Const.size_insertManyIfNewUnit_list_le ⟨m.1, _⟩ m.2 + Raw₀.Const.size_insertManyIfNewUnit_list_le ⟨m.1, _⟩ m.2 @[simp] theorem isEmpty_insertManyIfNewUnit_list [EquivBEq α] [LawfulHashable α] {l : List α} : (insertManyIfNewUnit m l).isEmpty = (m.isEmpty && l.isEmpty) := - Raw₀.Const.isEmpty_insertManyIfNewUnit_list ⟨m.1, _⟩ m.2 + Raw₀.Const.isEmpty_insertManyIfNewUnit_list ⟨m.1, _⟩ m.2 theorem get?_insertManyIfNewUnit_list [EquivBEq α] [LawfulHashable α] {l : List α} {k : α} : get? (insertManyIfNewUnit m l) k = if m.contains k ∨ l.contains k then some () else none := - Raw₀.Const.get?_insertManyIfNewUnit_list ⟨m.1, _⟩ m.2 + Raw₀.Const.get?_insertManyIfNewUnit_list ⟨m.1, _⟩ m.2 theorem get_insertManyIfNewUnit_list {l : List α} {k : α} {h} : - get (insertManyIfNewUnit m l) k h = () := + get (insertManyIfNewUnit m l) k h = () := Raw₀.Const.get_insertManyIfNewUnit_list ⟨m.1, _⟩ theorem get!_insertManyIfNewUnit_list {l : List α} {k : α} : - get! (insertManyIfNewUnit m l) k = () := + get! (insertManyIfNewUnit m l) k = () := Raw₀.Const.get!_insertManyIfNewUnit_list ⟨m.1, _⟩ theorem getD_insertManyIfNewUnit_list @@ -1457,14 +1453,14 @@ theorem get?_ofList_of_contains_eq_false [LawfulBEq α] {l : List ((a : α) × β a)} {k : α} (h : (l.map Sigma.fst).contains k = false) : (ofList l).get? k = none := - Raw₀.get?_insertMany_empty_list_of_contains_eq_false h + Raw₀.get?_insertMany_empty_list_of_contains_eq_false h theorem get?_ofList_of_mem [LawfulBEq α] {l : List ((a : α) × β a)} {k k' : α} (k_beq : k == k') {v : β k} (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) (mem : ⟨k, v⟩ ∈ l) : (ofList l).get? k' = some (cast (by congr; apply LawfulBEq.eq_of_beq k_beq) v) := - Raw₀.get?_insertMany_empty_list_of_mem k_beq distinct mem + Raw₀.get?_insertMany_empty_list_of_mem k_beq distinct mem theorem get_ofList_of_mem [LawfulBEq α] {l : List ((a : α) × β a)} {k k' : α} (k_beq : k == k') {v : β k} @@ -1472,26 +1468,26 @@ theorem get_ofList_of_mem [LawfulBEq α] (mem : ⟨k, v⟩ ∈ l) {h} : (ofList l).get k' h = cast (by congr; apply LawfulBEq.eq_of_beq k_beq) v := - Raw₀.get_insertMany_empty_list_of_mem k_beq distinct mem + Raw₀.get_insertMany_empty_list_of_mem k_beq distinct mem theorem get!_ofList_of_contains_eq_false [LawfulBEq α] {l : List ((a : α) × β a)} {k : α} [Inhabited (β k)] (h : (l.map Sigma.fst).contains k = false) : (ofList l).get! k = default := - Raw₀.get!_insertMany_empty_list_of_contains_eq_false h + Raw₀.get!_insertMany_empty_list_of_contains_eq_false h theorem get!_ofList_of_mem [LawfulBEq α] {l : List ((a : α) × β a)} {k k' : α} (k_beq : k == k') {v : β k} [Inhabited (β k')] (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) (mem : ⟨k, v⟩ ∈ l) : (ofList l).get! k' = cast (by congr; apply LawfulBEq.eq_of_beq k_beq) v := - Raw₀.get!_insertMany_empty_list_of_mem k_beq distinct mem + Raw₀.get!_insertMany_empty_list_of_mem k_beq distinct mem theorem getD_ofList_of_contains_eq_false [LawfulBEq α] {l : List ((a : α) × β a)} {k : α} {fallback : β k} (contains_eq_false : (l.map Sigma.fst).contains k = false) : (ofList l).getD k fallback = fallback := - Raw₀.getD_insertMany_empty_list_of_contains_eq_false contains_eq_false + Raw₀.getD_insertMany_empty_list_of_contains_eq_false contains_eq_false theorem getD_ofList_of_mem [LawfulBEq α] {l : List ((a : α) × β a)} {k k' : α} (k_beq : k == k') {v : β k} {fallback : β k'} @@ -1504,7 +1500,7 @@ theorem getKey?_ofList_of_contains_eq_false [EquivBEq α] [LawfulHashable α] {l : List ((a : α) × β a)} {k : α} (h : (l.map Sigma.fst).contains k = false) : (ofList l).getKey? k = none := - Raw₀.getKey?_insertMany_empty_list_of_contains_eq_false h + Raw₀.getKey?_insertMany_empty_list_of_contains_eq_false h theorem getKey?_ofList_of_mem [EquivBEq α] [LawfulHashable α] {l : List ((a : α) × β a)} @@ -1512,7 +1508,7 @@ theorem getKey?_ofList_of_mem [EquivBEq α] [LawfulHashable α] (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) (mem : k ∈ l.map Sigma.fst) : (ofList l).getKey? k' = some k := - Raw₀.getKey?_insertMany_empty_list_of_mem k_beq distinct mem + Raw₀.getKey?_insertMany_empty_list_of_mem k_beq distinct mem theorem getKey_ofList_of_mem [EquivBEq α] [LawfulHashable α] {l : List ((a : α) × β a)} @@ -1521,13 +1517,13 @@ theorem getKey_ofList_of_mem [EquivBEq α] [LawfulHashable α] (mem : k ∈ l.map Sigma.fst) {h'} : (ofList l).getKey k' h' = k := - Raw₀.getKey_insertMany_empty_list_of_mem k_beq distinct mem + Raw₀.getKey_insertMany_empty_list_of_mem k_beq distinct mem theorem getKey!_ofList_of_contains_eq_false [EquivBEq α] [LawfulHashable α] [Inhabited α] {l : List ((a : α) × β a)} {k : α} (h : (l.map Sigma.fst).contains k = false) : (ofList l).getKey! k = default := - Raw₀.getKey!_insertMany_empty_list_of_contains_eq_false h + Raw₀.getKey!_insertMany_empty_list_of_contains_eq_false h theorem getKey!_ofList_of_mem [EquivBEq α] [LawfulHashable α] [Inhabited α] {l : List ((a : α) × β a)} @@ -1535,7 +1531,7 @@ theorem getKey!_ofList_of_mem [EquivBEq α] [LawfulHashable α] [Inhabited α] (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) (mem : k ∈ l.map Sigma.fst) : (ofList l).getKey! k' = k := - Raw₀.getKey!_insertMany_empty_list_of_mem k_beq distinct mem + Raw₀.getKey!_insertMany_empty_list_of_mem k_beq distinct mem theorem getKeyD_ofList_of_contains_eq_false [EquivBEq α] [LawfulHashable α] {l : List ((a : α) × β a)} {k fallback : α} @@ -1549,7 +1545,7 @@ theorem getKeyD_ofList_of_mem [EquivBEq α] [LawfulHashable α] (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) (mem : k ∈ l.map Sigma.fst) : (ofList l).getKeyD k' fallback = k := - Raw₀.getKeyD_insertMany_empty_list_of_mem k_beq distinct mem + Raw₀.getKeyD_insertMany_empty_list_of_mem k_beq distinct mem theorem size_ofList [EquivBEq α] [LawfulHashable α] {l : List ((a : α) × β a)} (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) : @@ -1580,7 +1576,7 @@ theorem ofList_singleton {k : α} {v : β} [EquivBEq α] [LawfulHashable α] : Subtype.eq (congrArg Subtype.val (Raw₀.Const.insertMany_empty_list_singleton (α:= α)) :) theorem ofList_cons [EquivBEq α] [LawfulHashable α] {k : α} {v : β} {tl : List (α × β)} : - (insertMany empty (⟨k, v⟩ :: tl)) = (insertMany (empty.insert k v) tl) := + insertMany empty (⟨k, v⟩ :: tl) = insertMany (empty.insert k v) tl := Subtype.eq (congrArg Subtype.val (Raw₀.Const.insertMany_empty_list_cons (α:= α)) :) @[simp] @@ -1593,14 +1589,14 @@ theorem get?_ofList_of_contains_eq_false [LawfulBEq α] {l : List (α × β)} {k : α} (h : (l.map Prod.fst).contains k = false) : get? (ofList l) k = none := - Raw₀.Const.get?_insertMany_empty_list_of_contains_eq_false h + Raw₀.Const.get?_insertMany_empty_list_of_contains_eq_false h theorem get?_ofList_of_mem [LawfulBEq α] {l : List (α × β)} {k k' : α} (k_beq : k == k') {v : β} (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) (mem : ⟨k, v⟩ ∈ l) : get? (ofList l) k' = some v := - Raw₀.Const.get?_insertMany_empty_list_of_mem k_beq distinct mem + Raw₀.Const.get?_insertMany_empty_list_of_mem k_beq distinct mem theorem get_ofList_of_mem [LawfulBEq α] {l : List (α × β)} {k k' : α} (k_beq : k == k') {v : β} @@ -1608,20 +1604,20 @@ theorem get_ofList_of_mem [LawfulBEq α] (mem : ⟨k, v⟩ ∈ l) {h} : get (ofList l) k' h = v := - Raw₀.Const.get_insertMany_empty_list_of_mem k_beq distinct mem + Raw₀.Const.get_insertMany_empty_list_of_mem k_beq distinct mem theorem get!_ofList_of_contains_eq_false [LawfulBEq α] {l : List (α × β)} {k : α} [Inhabited β] (h : (l.map Prod.fst).contains k = false) : get! (ofList l) k = (default : β) := - Raw₀.Const.get!_insertMany_empty_list_of_contains_eq_false h + Raw₀.Const.get!_insertMany_empty_list_of_contains_eq_false h theorem get!_ofList_of_mem [LawfulBEq α] {l : List (α × β)} {k k' : α} (k_beq : k == k') {v : β} [Inhabited β] (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) (mem : ⟨k, v⟩ ∈ l) : get! (ofList l) k' = v := - Raw₀.Const.get!_insertMany_empty_list_of_mem k_beq distinct mem + Raw₀.Const.get!_insertMany_empty_list_of_mem k_beq distinct mem theorem getD_ofList_of_contains_eq_false [LawfulBEq α] {l : List (α × β)} {k : α} {fallback : β} @@ -1634,7 +1630,7 @@ theorem getD_ofList_of_mem [LawfulBEq α] (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) (mem : ⟨k, v⟩ ∈ l) : getD (ofList l) k' fallback = v := - Raw₀.Const.getD_insertMany_empty_list_of_mem k_beq distinct mem + Raw₀.Const.getD_insertMany_empty_list_of_mem k_beq distinct mem theorem getKey?_ofList_of_contains_eq_false [EquivBEq α] [LawfulHashable α] {l : List (α × β)} {k : α} @@ -1648,7 +1644,7 @@ theorem getKey?_ofList_of_mem [EquivBEq α] [LawfulHashable α] (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) (mem : k ∈ l.map Prod.fst) : (ofList l).getKey? k' = some k := - Raw₀.Const.getKey?_insertMany_empty_list_of_mem k_beq distinct mem + Raw₀.Const.getKey?_insertMany_empty_list_of_mem k_beq distinct mem theorem getKey_ofList_of_mem [EquivBEq α] [LawfulHashable α] {l : List (α × β)} @@ -1657,13 +1653,13 @@ theorem getKey_ofList_of_mem [EquivBEq α] [LawfulHashable α] (mem : k ∈ l.map Prod.fst) {h'} : (ofList l).getKey k' h' = k := - Raw₀.Const.getKey_insertMany_empty_list_of_mem k_beq distinct mem + Raw₀.Const.getKey_insertMany_empty_list_of_mem k_beq distinct mem theorem getKey!_ofList_of_contains_eq_false [EquivBEq α] [LawfulHashable α] [Inhabited α] {l : List (α × β)} {k : α} (h : (l.map Prod.fst).contains k = false) : (ofList l).getKey! k = default := - Raw₀.Const.getKey!_insertMany_empty_list_of_contains_eq_false h + Raw₀.Const.getKey!_insertMany_empty_list_of_contains_eq_false h theorem getKey!_ofList_of_mem [EquivBEq α] [LawfulHashable α] [Inhabited α] {l : List (α × β)} @@ -1671,7 +1667,7 @@ theorem getKey!_ofList_of_mem [EquivBEq α] [LawfulHashable α] [Inhabited α] (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) (mem : k ∈ l.map Prod.fst) : (ofList l).getKey! k' = k := - Raw₀.Const.getKey!_insertMany_empty_list_of_mem k_beq distinct mem + Raw₀.Const.getKey!_insertMany_empty_list_of_mem k_beq distinct mem theorem getKeyD_ofList_of_contains_eq_false [EquivBEq α] [LawfulHashable α] {l : List (α × β)} {k fallback : α} @@ -1685,7 +1681,7 @@ theorem getKeyD_ofList_of_mem [EquivBEq α] [LawfulHashable α] (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) (mem : k ∈ l.map Prod.fst) : (ofList l).getKeyD k' fallback = k := - Raw₀.Const.getKeyD_insertMany_empty_list_of_mem k_beq distinct mem + Raw₀.Const.getKeyD_insertMany_empty_list_of_mem k_beq distinct mem theorem size_ofList [EquivBEq α] [LawfulHashable α] {l : List (α × β)} (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) : @@ -1694,7 +1690,7 @@ theorem size_ofList [EquivBEq α] [LawfulHashable α] theorem size_ofList_le [EquivBEq α] [LawfulHashable α] {l : List (α × β)} : - (ofList l).size ≤ l.length := + (ofList l).size ≤ l.length := Raw₀.Const.size_insertMany_empty_list_le @[simp] @@ -1715,7 +1711,7 @@ theorem unitOfList_singleton [EquivBEq α] [LawfulHashable α] {k : α} : theorem unitOfList_cons [EquivBEq α] [LawfulHashable α] {hd : α} {tl : List α} : unitOfList (hd :: tl) = - (insertManyIfNewUnit (empty.insertIfNew hd ()) tl) := + insertManyIfNewUnit (empty.insertIfNew hd ()) tl := Subtype.eq (congrArg Subtype.val (Raw₀.Const.insertManyIfNewUnit_empty_list_cons (α := α)) :) @[simp] @@ -1727,13 +1723,13 @@ theorem contains_unitOfList [EquivBEq α] [LawfulHashable α] theorem getKey?_unitOfList_of_contains_eq_false [EquivBEq α] [LawfulHashable α] {l : List α} {k : α} (h' : l.contains k = false) : getKey? (unitOfList l) k = none := - Raw₀.Const.getKey?_insertManyIfNewUnit_empty_list_of_contains_eq_false h' + Raw₀.Const.getKey?_insertManyIfNewUnit_empty_list_of_contains_eq_false h' theorem getKey?_unitOfList_of_mem [EquivBEq α] [LawfulHashable α] {l : List α} {k k' : α} (k_beq : k == k') (distinct : l.Pairwise (fun a b => (a == b) = false)) (mem : k ∈ l) : getKey? (unitOfList l) k' = some k := - Raw₀.Const.getKey?_insertManyIfNewUnit_empty_list_of_mem k_beq distinct mem + Raw₀.Const.getKey?_insertManyIfNewUnit_empty_list_of_mem k_beq distinct mem theorem getKey_unitOfList_of_mem [EquivBEq α] [LawfulHashable α] {l : List α} @@ -1741,7 +1737,7 @@ theorem getKey_unitOfList_of_mem [EquivBEq α] [LawfulHashable α] (distinct : l.Pairwise (fun a b => (a == b) = false)) (mem : k ∈ l) {h'} : getKey (unitOfList l) k' h' = k := - Raw₀.Const.getKey_insertManyIfNewUnit_empty_list_of_mem k_beq distinct mem + Raw₀.Const.getKey_insertManyIfNewUnit_empty_list_of_mem k_beq distinct mem theorem getKey!_unitOfList_of_contains_eq_false [EquivBEq α] [LawfulHashable α] [Inhabited α] {l : List α} {k : α} @@ -1754,7 +1750,7 @@ theorem getKey!_unitOfList_of_mem [EquivBEq α] [LawfulHashable α] (distinct : l.Pairwise (fun a b => (a == b) = false)) (mem : k ∈ l) : getKey! (unitOfList l) k' = k := - Raw₀.Const.getKey!_insertManyIfNewUnit_empty_list_of_mem k_beq distinct mem + Raw₀.Const.getKey!_insertManyIfNewUnit_empty_list_of_mem k_beq distinct mem theorem getKeyD_unitOfList_of_contains_eq_false [EquivBEq α] [LawfulHashable α] {l : List α} {k fallback : α} @@ -1796,13 +1792,13 @@ theorem get?_unitOfList [EquivBEq α] [LawfulHashable α] @[simp] theorem get_unitOfList {l : List α} {k : α} {h} : - get (unitOfList l) k h = () := + get (unitOfList l) k h = () := Raw₀.Const.get_insertManyIfNewUnit_empty_list @[simp] theorem get!_unitOfList {l : List α} {k : α} : - get! (unitOfList l) k = () := + get! (unitOfList l) k = () := Raw₀.Const.get!_insertManyIfNewUnit_empty_list @[simp] From d07c76e1aee8d1befe490c62a79b119a9502dbd3 Mon Sep 17 00:00:00 2001 From: jt0202 Date: Mon, 6 Jan 2025 11:22:41 +0100 Subject: [PATCH 62/83] Rewrite of DHashMap.RawLemmas --- src/Std/Data/DHashMap/Internal/Raw.lean | 16 ++ src/Std/Data/DHashMap/RawLemmas.lean | 248 ++++++++++-------------- 2 files changed, 119 insertions(+), 145 deletions(-) diff --git a/src/Std/Data/DHashMap/Internal/Raw.lean b/src/Std/Data/DHashMap/Internal/Raw.lean index 5c9b4c7606cb..5fd8539388d1 100644 --- a/src/Std/Data/DHashMap/Internal/Raw.lean +++ b/src/Std/Data/DHashMap/Internal/Raw.lean @@ -207,6 +207,11 @@ theorem insertMany_val [BEq α][Hashable α] {m : Raw₀ α β} {ρ : Type w} [F m.val.insertMany l = m.insertMany l := by simp [Raw.insertMany, m.2] +theorem ofList_eq [BEq α] [Hashable α] {l: List ((a : α) × β a)} : + Raw.ofList l = Raw₀.insertMany Raw₀.empty l := by + simp only [Raw.ofList, Raw.insertMany, (Raw.WF.empty).size_buckets_pos ∅, ↓reduceDIte] + congr + section variable {β : Type v} @@ -219,6 +224,11 @@ theorem Const.insertMany_val [BEq α][Hashable α] {m : Raw₀ α (fun _ => β)} Raw.Const.insertMany m.val l = Raw₀.Const.insertMany m l := by simp [Raw.Const.insertMany, m.2] +theorem Const.ofList_eq [BEq α] [Hashable α] {l: List (α × β)} : + Raw.Const.ofList l = Raw₀.Const.insertMany Raw₀.empty l := by + simp only [Raw.Const.ofList, Raw.Const.insertMany, (Raw.WF.empty).size_buckets_pos ∅, ↓reduceDIte] + congr + theorem Const.insertManyIfNewUnit_eq {ρ : Type w} [ForIn Id ρ α] [BEq α] [Hashable α] {m : Raw α (fun _ => Unit)} {l : ρ} (h : m.WF): Raw.Const.insertManyIfNewUnit m l = Raw₀.Const.insertManyIfNewUnit ⟨m, h.size_buckets_pos⟩ l := by @@ -229,6 +239,12 @@ theorem Const.insertManyIfNewUnit_val {ρ : Type w} [ForIn Id ρ α] [BEq α] [H Raw.Const.insertManyIfNewUnit m.val l = Raw₀.Const.insertManyIfNewUnit m l := by simp [Raw.Const.insertManyIfNewUnit, m.2] +theorem Const.unitOfList_eq [BEq α] [Hashable α] {l: List α} : + Raw.Const.unitOfList l = Raw₀.Const.insertManyIfNewUnit Raw₀.empty l := by + simp only [Raw.Const.unitOfList, Raw.Const.insertManyIfNewUnit, (Raw.WF.empty).size_buckets_pos ∅, + ↓reduceDIte] + congr + theorem Const.get?_eq [BEq α] [Hashable α] {m : Raw α (fun _ => β)} (h : m.WF) {a : α} : Raw.Const.get? m a = Raw₀.Const.get? ⟨m, h.size_buckets_pos⟩ a := by simp [Raw.Const.get?, h.size_buckets_pos] diff --git a/src/Std/Data/DHashMap/RawLemmas.lean b/src/Std/Data/DHashMap/RawLemmas.lean index 4a323de994ce..0012aeab1a04 100644 --- a/src/Std/Data/DHashMap/RawLemmas.lean +++ b/src/Std/Data/DHashMap/RawLemmas.lean @@ -61,7 +61,8 @@ private def baseNames : Array Name := ``getKeyD_eq, ``getKeyD_val, ``insertMany_eq, ``insertMany_val, ``Const.insertMany_eq, ``Const.insertMany_val, - ``Const.insertManyIfNewUnit_eq, ``Const.insertManyIfNewUnit_val] + ``Const.insertManyIfNewUnit_eq, ``Const.insertManyIfNewUnit_val, + ``ofList_eq, ``Const.ofList_eq, ``Const.unitOfList_eq] /-- Internal implementation detail of the hash map -/ scoped syntax "simp_to_raw" ("using" term)? : tactic @@ -1062,11 +1063,13 @@ theorem insertMany_list_singleton {k : α} {v : β k} [EquivBEq α] [LawfulHasha simp_to_raw rw [Raw₀.insertMany_list_singleton] -theorem insertMany_cons {l : List ((a : α) × β a)} {k : α} {v : β k} [EquivBEq α] [LawfulHashable α] (h : m.WF) : +theorem insertMany_cons {l : List ((a : α) × β a)} {k : α} {v : β k} [EquivBEq α] [LawfulHashable α] + (h : m.WF) : m.insertMany (⟨k, v⟩ :: l) = (m.insert k v).insertMany l := by simp_to_raw rw [Raw₀.insertMany_cons] +@[simp] theorem contains_insertMany_list [EquivBEq α] [LawfulHashable α] (h : m.WF) {l : List ((a : α) × β a)} {k : α} : (m.insertMany l).contains k = (m.contains k || (l.map Sigma.fst).contains k) := by @@ -1202,6 +1205,7 @@ theorem size_insertMany_list_le [EquivBEq α] [LawfulHashable α] (h : m.WF) (m.insertMany l).size ≤ m.size + l.length := by simp_to_raw using Raw₀.size_insertMany_list_le +@[simp] theorem isEmpty_insertMany_list [EquivBEq α] [LawfulHashable α] (h : m.WF) {l : List ((a : α) × β a)} : (m.insertMany l).isEmpty = (m.isEmpty && l.isEmpty) := by @@ -1218,17 +1222,18 @@ theorem insertMany_nil [EquivBEq α] [LawfulHashable α] (h : m.WF) : @[simp] theorem insertMany_list_singleton [EquivBEq α] [LawfulHashable α] (h : m.WF) - {k : α} {v : β} : + {k : α} {v : β} : insertMany m [⟨k, v⟩] = m.insert k v := by simp_to_raw rw [Raw₀.Const.insertMany_list_singleton] theorem insertMany_cons [EquivBEq α] [LawfulHashable α] (h : m.WF) {l : List (α × β)} - {k : α} {v : β} : + {k : α} {v : β} : insertMany m (⟨k, v⟩ :: l) = insertMany (m.insert k v) l := by simp_to_raw rw [Raw₀.Const.insertMany_cons] +@[simp] theorem contains_insertMany_list [EquivBEq α] [LawfulHashable α] (h : m.WF) {l : List (α × β)} {k : α} : (insertMany m l).contains k = (m.contains k || (l.map Prod.fst).contains k) := by @@ -1246,7 +1251,7 @@ theorem getKey?_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashab simp_to_raw using Raw₀.Const.getKey?_insertMany_list_of_contains_eq_false theorem getKey?_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] (h : m.WF) - {l : List (α × β)} + {l : List (α × β)} {k k' : α} (k_beq : k == k') (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) (mem : k ∈ l.map Prod.fst) : @@ -1310,6 +1315,7 @@ theorem size_insertMany_list_le [EquivBEq α] [LawfulHashable α] (h : m.WF) (insertMany m l).size ≤ m.size + l.length := by simp_to_raw using Raw₀.Const.size_insertMany_list_le +@[simp] theorem isEmpty_insertMany_list [EquivBEq α] [LawfulHashable α] (h : m.WF) {l : List (α × β)} : (insertMany m l).isEmpty = (m.isEmpty && l.isEmpty) := by @@ -1341,7 +1347,7 @@ theorem get_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] (h : m.WF) simp_to_raw using Raw₀.Const.get_insertMany_list_of_mem theorem get!_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] - [Inhabited β] (h : m.WF) {l : List (α × β)} {k : α} + [Inhabited β] (h : m.WF) {l : List (α × β)} {k : α} (h' : (l.map Prod.fst).contains k = false) : get! (insertMany m l) k = get! m k := by simp_to_raw using Raw₀.Const.get!_insertMany_list_of_contains_eq_false @@ -1383,6 +1389,7 @@ theorem insertManyIfNewUnit_cons [EquivBEq α] [LawfulHashable α] (h : m.WF) {l simp_to_raw rw [Raw₀.Const.insertManyIfNewUnit_cons] +@[simp] theorem contains_insertManyIfNewUnit_list [EquivBEq α] [LawfulHashable α] (h : m.WF) {l : List α} {k : α} : (insertManyIfNewUnit m l).contains k = (m.contains k || l.contains k) := by @@ -1450,7 +1457,7 @@ theorem getKey!_insertManyIfNewUnit_list_mem_of_contains_of_contains [EquivBEq [Inhabited α] (h : m.WF) {l : List α} {k : α} (distinct : l.Pairwise (fun a b => (a == b) = false)) (h' : l.contains k) : m.contains k → - getKey! (insertManyIfNewUnit m l) k = getKey! m k := by + getKey! (insertManyIfNewUnit m l) k = getKey! m k := by simp_to_raw using Raw₀.Const.getKey!_insertManyIfNewUnit_list_mem_of_contains_of_contains theorem getKeyD_insertManyIfNewUnit_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] @@ -1485,27 +1492,32 @@ theorem size_insertManyIfNewUnit_list_le [EquivBEq α] [LawfulHashable α] (h : (insertManyIfNewUnit m l).size ≤ m.size + l.length := by simp_to_raw using Raw₀.Const.size_insertManyIfNewUnit_list_le +@[simp] theorem isEmpty_insertManyIfNewUnit_list [EquivBEq α] [LawfulHashable α] (h : m.WF) {l : List α} : (insertManyIfNewUnit m l).isEmpty = (m.isEmpty && l.isEmpty) := by simp_to_raw using Raw₀.Const.isEmpty_insertManyIfNewUnit_list +@[simp] theorem get?_insertManyIfNewUnit_list [EquivBEq α] [LawfulHashable α] (h : m.WF) {l : List α} {k : α} : get? (insertManyIfNewUnit m l) k = if m.contains k ∨ l.contains k then some () else none := by simp_to_raw using Raw₀.Const.get?_insertManyIfNewUnit_list +@[simp] theorem get_insertManyIfNewUnit_list {l : List α} {k : α} {h} : - get (insertManyIfNewUnit m l) k h = () := by + get (insertManyIfNewUnit m l) k h = () := by simp +@[simp] theorem get!_insertManyIfNewUnit_list {l : List α} {k : α} : - get! (insertManyIfNewUnit m l) k = () := by + get! (insertManyIfNewUnit m l) k = () := by simp +@[simp] theorem getD_insertManyIfNewUnit_list {l : List α} {k : α} {fallback : Unit} : getD (insertManyIfNewUnit m l) k fallback = () := by @@ -1516,43 +1528,43 @@ end Raw namespace Raw variable [BEq α] [Hashable α] +open Internal.Raw Internal.Raw₀ @[simp] theorem ofList_nil [EquivBEq α] [LawfulHashable α] : ofList ([] : List ((a : α) × (β a))) = empty := by - simp only [ofList, WF.emptyc, insertMany_nil] - rfl + simp_to_raw + rw [Raw₀.insertMany_empty_list_nil] @[simp] theorem ofList_singleton {k : α} {v : β k} [EquivBEq α] [LawfulHashable α] : ofList [⟨k, v⟩] = empty.insert k v := by - simp only [ofList, WF.emptyc, insertMany_list_singleton] - rfl + simp_to_raw + rw [Raw₀.insertMany_empty_list_singleton] + theorem ofList_cons [EquivBEq α] [LawfulHashable α] {hd : (a : α) × (β a)} {tl : List ((a : α) × (β a))} : ofList (hd::tl) = ((∅ : Raw α β).insert hd.1 hd.2).insertMany tl := by - simp only [ofList] - rw [insertMany_cons WF.emptyc] + simp_to_raw + rw [Raw₀.insertMany_empty_list_cons] +@[simp] theorem contains_ofList [EquivBEq α] [LawfulHashable α] {l : List ((a : α) × β a)} {k : α} : (ofList l).contains k = (l.map Sigma.fst).contains k := by - simp only [ofList, WF.emptyc, contains_insertMany_list, contains_emptyc, Bool.false_or] + simp_to_raw using Raw₀.contains_insertMany_empty_list theorem get?_ofList_of_contains_eq_false [LawfulBEq α] {l : List ((a : α) × β a)} {k : α} (h : (l.map Sigma.fst).contains k = false) : (ofList l).get? k = none := by - simp only [ofList] - rw [get?_insertMany_list_of_contains_eq_false WF.emptyc h] - apply get?_emptyc + simp_to_raw using Raw₀.get?_insertMany_empty_list_of_contains_eq_false theorem get?_ofList_of_mem [LawfulBEq α] {l : List ((a : α) × β a)} {k k' : α} (k_beq : k == k') {v : β k} (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) (mem : ⟨k, v⟩ ∈ l) : (ofList l).get? k' = some (cast (by congr; apply LawfulBEq.eq_of_beq k_beq) v) := by - simp only [ofList] - rw [get?_insertMany_list_of_mem WF.emptyc k_beq distinct mem] + simp_to_raw using Raw₀.get?_insertMany_empty_list_of_mem theorem get_ofList_of_mem [LawfulBEq α] {l : List ((a : α) × β a)} {k k' : α} (k_beq : k == k') {v : β k} @@ -1560,48 +1572,39 @@ theorem get_ofList_of_mem [LawfulBEq α] (mem : ⟨k, v⟩ ∈ l) {h} : (ofList l).get k' h = cast (by congr; apply LawfulBEq.eq_of_beq k_beq) v := by - simp only [ofList] - rw [get_insertMany_list_of_mem WF.emptyc k_beq distinct mem] + simp_to_raw using Raw₀.get_insertMany_empty_list_of_mem theorem get!_ofList_of_contains_eq_false [LawfulBEq α] {l : List ((a : α) × β a)} {k : α} [Inhabited (β k)] (h : (l.map Sigma.fst).contains k = false) : (ofList l).get! k = default := by - simp only [ofList] - rw [get!_insertMany_list_of_contains_eq_false WF.emptyc h] - apply get!_emptyc + simp_to_raw using get!_insertMany_empty_list_of_contains_eq_false theorem get!_ofList_of_mem [LawfulBEq α] {l : List ((a : α) × β a)} {k k' : α} (k_beq : k == k') {v : β k} [Inhabited (β k')] (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) (mem : ⟨k, v⟩ ∈ l) : (ofList l).get! k' = cast (by congr; apply LawfulBEq.eq_of_beq k_beq) v := by - simp only [ofList] - rw [get!_insertMany_list_of_mem WF.emptyc k_beq distinct mem] + simp_to_raw using Raw₀.get!_insertMany_empty_list_of_mem theorem getD_ofList_of_contains_eq_false [LawfulBEq α] {l : List ((a : α) × β a)} {k : α} {fallback : β k} (contains_eq_false : (l.map Sigma.fst).contains k = false) : (ofList l).getD k fallback = fallback := by - simp only [ofList] - rw [getD_insertMany_list_of_contains_eq_false WF.emptyc contains_eq_false] - apply getD_emptyc + simp_to_raw using Raw₀.getD_insertMany_empty_list_of_contains_eq_false theorem getD_ofList_of_mem [LawfulBEq α] {l : List ((a : α) × β a)} {k k' : α} (k_beq : k == k') {v : β k} {fallback : β k'} (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) (mem : ⟨k, v⟩ ∈ l) : (ofList l).getD k' fallback = cast (by congr; apply LawfulBEq.eq_of_beq k_beq) v := by - simp only [ofList] - rw [getD_insertMany_list_of_mem WF.emptyc k_beq distinct mem] + simp_to_raw using Raw₀.getD_insertMany_empty_list_of_mem theorem getKey?_ofList_of_contains_eq_false [EquivBEq α] [LawfulHashable α] {l : List ((a : α) × β a)} {k : α} (h : (l.map Sigma.fst).contains k = false) : (ofList l).getKey? k = none := by - simp only [ofList] - rw [getKey?_insertMany_list_of_contains_eq_false WF.emptyc h] - apply getKey?_emptyc + simp_to_raw using Raw₀.getKey?_insertMany_empty_list_of_contains_eq_false theorem getKey?_ofList_of_mem [EquivBEq α] [LawfulHashable α] {l : List ((a : α) × β a)} @@ -1609,8 +1612,7 @@ theorem getKey?_ofList_of_mem [EquivBEq α] [LawfulHashable α] (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) (mem : k ∈ l.map Sigma.fst) : (ofList l).getKey? k' = some k := by - simp only [ofList] - rw [getKey?_insertMany_list_of_mem WF.emptyc k_beq distinct mem] + simp_to_raw using Raw₀.getKey?_insertMany_empty_list_of_mem theorem getKey_ofList_of_mem [EquivBEq α] [LawfulHashable α] {l : List ((a : α) × β a)} @@ -1619,16 +1621,13 @@ theorem getKey_ofList_of_mem [EquivBEq α] [LawfulHashable α] (mem : k ∈ l.map Sigma.fst) {h'} : (ofList l).getKey k' h' = k := by - simp only [ofList] - rw [getKey_insertMany_list_of_mem WF.emptyc k_beq distinct mem] + simp_to_raw using Raw₀.getKey_insertMany_empty_list_of_mem theorem getKey!_ofList_of_contains_eq_false [EquivBEq α] [LawfulHashable α] [Inhabited α] {l : List ((a : α) × β a)} {k : α} (h : (l.map Sigma.fst).contains k = false) : (ofList l).getKey! k = default := by - simp only [ofList] - rw [getKey!_insertMany_list_of_contains_eq_false WF.emptyc h] - apply getKey!_emptyc + simp_to_raw using Raw₀.getKey!_insertMany_empty_list_of_contains_eq_false theorem getKey!_ofList_of_mem [EquivBEq α] [LawfulHashable α] [Inhabited α] {l : List ((a : α) × β a)} @@ -1636,16 +1635,13 @@ theorem getKey!_ofList_of_mem [EquivBEq α] [LawfulHashable α] [Inhabited α] (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) (mem : k ∈ l.map Sigma.fst) : (ofList l).getKey! k' = k := by - simp only [ofList] - rw [getKey!_insertMany_list_of_mem WF.emptyc k_beq distinct mem] + simp_to_raw using Raw₀.getKey!_insertMany_empty_list_of_mem theorem getKeyD_ofList_of_contains_eq_false [EquivBEq α] [LawfulHashable α] {l : List ((a : α) × β a)} {k fallback : α} (h : (l.map Sigma.fst).contains k = false) : (ofList l).getKeyD k fallback = fallback := by - simp only [ofList] - rw [getKeyD_insertMany_list_of_contains_eq_false WF.emptyc h] - apply getKeyD_emptyc + simp_to_raw using Raw₀.getKeyD_insertMany_empty_list_of_contains_eq_false theorem getKeyD_ofList_of_mem [EquivBEq α] [LawfulHashable α] {l : List ((a : α) × β a)} @@ -1653,30 +1649,23 @@ theorem getKeyD_ofList_of_mem [EquivBEq α] [LawfulHashable α] (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) (mem : k ∈ l.map Sigma.fst) : (ofList l).getKeyD k' fallback = k := by - simp only [ofList] - rw [getKeyD_insertMany_list_of_mem WF.emptyc k_beq distinct mem] + simp_to_raw using Raw₀.getKeyD_insertMany_empty_list_of_mem theorem size_ofList [EquivBEq α] [LawfulHashable α] {l : List ((a : α) × β a)} (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) : (ofList l).size = l.length := by - simp only [ofList] - rw [size_insertMany_list WF.emptyc distinct] - · simp only [size_emptyc, Nat.zero_add] - · simp only [contains_emptyc, Bool.false_eq_true, false_implies, implies_true] + simp_to_raw using Raw₀.size_insertMany_empty_list theorem size_ofList_le [EquivBEq α] [LawfulHashable α] {l : List ((a : α) × β a)} : - (ofList l).size ≤ l.length := by - simp only [ofList] - rw [← Nat.zero_add l.length, ← size_emptyc] - apply (size_insertMany_list_le WF.emptyc) + (ofList l).size ≤ l.length := by + simp_to_raw using Raw₀.size_insertMany_empty_list_le +@[simp] theorem isEmpty_ofList [EquivBEq α] [LawfulHashable α] {l : List ((a : α) × β a)} : (ofList l).isEmpty = l.isEmpty := by - simp only [ofList] - rw [isEmpty_insertMany_list WF.emptyc] - simp + simp_to_raw using Raw₀.isEmpty_insertMany_empty_list namespace Const variable {β : Type v} @@ -1684,40 +1673,38 @@ variable {β : Type v} @[simp] theorem ofList_nil [EquivBEq α] [LawfulHashable α] : ofList ([] : List (α × β)) = empty := by - simp only [ofList, WF.emptyc, insertMany_nil] - rfl + simp_to_raw + simp @[simp] theorem ofList_singleton {k : α} {v : β} [EquivBEq α] [LawfulHashable α] : ofList [⟨k, v⟩] = empty.insert k v := by - simp only [ofList, WF.emptyc, insertMany_list_singleton] - rfl + simp_to_raw + simp theorem ofList_cons [EquivBEq α] [LawfulHashable α] {hd : α × β} {tl : List (α × β)} : ofList (hd::tl) = insertMany ((∅ : Raw α (fun _ => β)).insert hd.1 hd.2) tl := by - simp only [ofList] - rw [insertMany_cons WF.emptyc] + simp_to_raw + rw [Raw₀.Const.insertMany_empty_list_cons] +@[simp] theorem contains_ofList [EquivBEq α] [LawfulHashable α] {l : List (α × β)} {k : α} : (ofList l).contains k = (l.map Prod.fst).contains k := by - simp only [ofList, WF.emptyc, contains_insertMany_list, contains_emptyc, Bool.false_or] + simp_to_raw using Raw₀.Const.contains_insertMany_empty_list theorem get?_ofList_of_contains_eq_false [LawfulBEq α] {l : List (α × β)} {k : α} (h : (l.map Prod.fst).contains k = false) : get? (ofList l) k = none := by - simp only [ofList] - rw [get?_insertMany_list_of_contains_eq_false WF.emptyc h] - apply get?_emptyc + simp_to_raw using Raw₀.Const.get?_insertMany_empty_list_of_contains_eq_false theorem get?_ofList_of_mem [LawfulBEq α] {l : List (α × β)} {k k' : α} (k_beq : k == k') {v : β} (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) (mem : ⟨k, v⟩ ∈ l) : get? (ofList l) k' = some v := by - simp only [ofList] - rw [get?_insertMany_list_of_mem WF.emptyc k_beq distinct mem] + simp_to_raw using Raw₀.Const.get?_insertMany_empty_list_of_mem theorem get_ofList_of_mem [LawfulBEq α] {l : List (α × β)} {k k' : α} (k_beq : k == k') {v : β} @@ -1725,48 +1712,39 @@ theorem get_ofList_of_mem [LawfulBEq α] (mem : ⟨k, v⟩ ∈ l) {h} : get (ofList l) k' h = v := by - simp only [ofList] - rw [get_insertMany_list_of_mem WF.emptyc k_beq distinct mem] + simp_to_raw using Raw₀.Const.get_insertMany_empty_list_of_mem theorem get!_ofList_of_contains_eq_false [LawfulBEq α] {l : List (α × β)} {k : α} [Inhabited β] (h : (l.map Prod.fst).contains k = false) : get! (ofList l) k = default := by - simp only [ofList] - rw [get!_insertMany_list_of_contains_eq_false WF.emptyc h] - apply get!_emptyc + simp_to_raw using Raw₀.Const.get!_insertMany_empty_list_of_contains_eq_false theorem get!_ofList_of_mem [LawfulBEq α] {l : List (α × β)} {k k' : α} (k_beq : k == k') {v : β} [Inhabited β] (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) (mem : ⟨k, v⟩ ∈ l) : get! (ofList l) k' = v := by - simp only [ofList] - rw [get!_insertMany_list_of_mem WF.emptyc k_beq distinct mem] + simp_to_raw using Raw₀.Const.get!_insertMany_empty_list_of_mem theorem getD_ofList_of_contains_eq_false [LawfulBEq α] {l : List (α × β)} {k : α} {fallback : β} (contains_eq_false : (l.map Prod.fst).contains k = false) : getD (ofList l) k fallback = fallback := by - simp only [ofList] - rw [getD_insertMany_list_of_contains_eq_false WF.emptyc contains_eq_false] - apply getD_emptyc + simp_to_raw using Raw₀.Const.getD_insertMany_empty_list_of_contains_eq_false theorem getD_ofList_of_mem [LawfulBEq α] {l : List (α × β)} {k k' : α} (k_beq : k == k') {v : β} {fallback : β} (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) (mem : ⟨k, v⟩ ∈ l) : getD (ofList l) k' fallback = v := by - simp only [ofList] - rw [getD_insertMany_list_of_mem WF.emptyc k_beq distinct mem] + simp_to_raw using Raw₀.Const.getD_insertMany_empty_list_of_mem theorem getKey?_ofList_of_contains_eq_false [EquivBEq α] [LawfulHashable α] {l : List (α × β)} {k : α} (h : (l.map Prod.fst).contains k = false) : (ofList l).getKey? k = none := by - simp only [ofList] - rw [getKey?_insertMany_list_of_contains_eq_false WF.emptyc h] - apply getKey?_emptyc + simp_to_raw using Raw₀.Const.getKey?_insertMany_empty_list_of_contains_eq_false theorem getKey?_ofList_of_mem [EquivBEq α] [LawfulHashable α] {l : List (α × β)} @@ -1774,8 +1752,7 @@ theorem getKey?_ofList_of_mem [EquivBEq α] [LawfulHashable α] (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) (mem : k ∈ l.map Prod.fst) : (ofList l).getKey? k' = some k := by - simp only [ofList] - rw [getKey?_insertMany_list_of_mem WF.emptyc k_beq distinct mem] + simp_to_raw using Raw₀.Const.getKey?_insertMany_empty_list_of_mem theorem getKey_ofList_of_mem [EquivBEq α] [LawfulHashable α] {l : List (α × β)} @@ -1784,16 +1761,13 @@ theorem getKey_ofList_of_mem [EquivBEq α] [LawfulHashable α] (mem : k ∈ l.map Prod.fst) {h'} : (ofList l).getKey k' h' = k := by - simp only [ofList] - rw [getKey_insertMany_list_of_mem WF.emptyc k_beq distinct mem] + simp_to_raw using Raw₀.Const.getKey_insertMany_empty_list_of_mem theorem getKey!_ofList_of_contains_eq_false [EquivBEq α] [LawfulHashable α] [Inhabited α] {l : List (α × β)} {k : α} (h : (l.map Prod.fst).contains k = false) : (ofList l).getKey! k = default := by - simp only [ofList] - rw [getKey!_insertMany_list_of_contains_eq_false WF.emptyc h] - apply getKey!_emptyc + simp_to_raw using Raw₀.Const.getKey!_insertMany_empty_list_of_contains_eq_false theorem getKey!_ofList_of_mem [EquivBEq α] [LawfulHashable α] [Inhabited α] {l : List (α × β)} @@ -1801,16 +1775,13 @@ theorem getKey!_ofList_of_mem [EquivBEq α] [LawfulHashable α] [Inhabited α] (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) (mem : k ∈ l.map Prod.fst) : (ofList l).getKey! k' = k := by - simp only [ofList] - rw [getKey!_insertMany_list_of_mem WF.emptyc k_beq distinct mem] + simp_to_raw using Raw₀.Const.getKey!_insertMany_empty_list_of_mem theorem getKeyD_ofList_of_contains_eq_false [EquivBEq α] [LawfulHashable α] {l : List (α × β)} {k fallback : α} (h : (l.map Prod.fst).contains k = false) : (ofList l).getKeyD k fallback = fallback := by - simp only [ofList] - rw [getKeyD_insertMany_list_of_contains_eq_false WF.emptyc h] - apply getKeyD_emptyc + simp_to_raw using Raw₀.Const.getKeyD_insertMany_empty_list_of_contains_eq_false theorem getKeyD_ofList_of_mem [EquivBEq α] [LawfulHashable α] {l : List (α × β)} @@ -1818,62 +1789,57 @@ theorem getKeyD_ofList_of_mem [EquivBEq α] [LawfulHashable α] (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) (mem : k ∈ l.map Prod.fst) : (ofList l).getKeyD k' fallback = k := by - simp only [ofList] - rw [getKeyD_insertMany_list_of_mem WF.emptyc k_beq distinct mem] + simp_to_raw using Raw₀.Const.getKeyD_insertMany_empty_list_of_mem theorem size_ofList [EquivBEq α] [LawfulHashable α] {l : List (α × β)} (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) : (ofList l).size = l.length := by - simp only [ofList] - rw [size_insertMany_list WF.emptyc distinct] - · simp only [size_emptyc, Nat.zero_add] - · simp only [contains_emptyc, Bool.false_eq_true, false_implies, implies_true] + simp_to_raw using Raw₀.Const.size_insertMany_empty_list theorem size_ofList_le [EquivBEq α] [LawfulHashable α] {l : List (α × β)} : - (ofList l).size ≤ l.length := by - simp only [ofList] - rw [← Nat.zero_add l.length, ← size_emptyc] - apply (size_insertMany_list_le WF.emptyc) + (ofList l).size ≤ l.length := by + simp_to_raw using Raw₀.Const.size_insertMany_empty_list_le +@[simp] theorem isEmpty_ofList [EquivBEq α] [LawfulHashable α] {l : List (α × β)} : (ofList l).isEmpty = l.isEmpty := by - simp [ofList, isEmpty_insertMany_list WF.emptyc] + simp_to_raw using Raw₀.Const.isEmpty_insertMany_empty_list @[simp] theorem unitOfList_nil [EquivBEq α] [LawfulHashable α] : unitOfList ([] : List α) = empty := by - simp only [unitOfList, insertManyIfNewUnit_nil WF.emptyc] - rfl + simp_to_raw + simp @[simp] theorem unitOfList_singleton [EquivBEq α] [LawfulHashable α] {k : α} : unitOfList [k] = empty.insertIfNew k () := by - simp only [unitOfList, insertManyIfNewUnit_list_singleton WF.emptyc] - rfl + simp_to_raw + simp theorem unitOfList_cons [EquivBEq α] [LawfulHashable α] {hd : α} {tl : List α} : unitOfList (hd::tl) = insertManyIfNewUnit ((∅ : Raw α (fun _ => Unit)).insertIfNew hd ()) tl := by - simp only [unitOfList] - rw [insertManyIfNewUnit_cons WF.emptyc] + simp_to_raw + rw [Raw₀.Const.insertManyIfNewUnit_empty_list_cons] +@[simp] theorem contains_unitOfList [EquivBEq α] [LawfulHashable α] {l : List α} {k : α} : (unitOfList l).contains k = l.contains k := by - simp [unitOfList, contains_insertManyIfNewUnit_list WF.emptyc] + simp_to_raw using Raw₀.Const.contains_insertManyIfNewUnit_empty_list theorem getKey?_unitOfList_of_contains_eq_false [EquivBEq α] [LawfulHashable α] {l : List α} {k : α} (h' : l.contains k = false) : getKey? (unitOfList l) k = none := by - simp [unitOfList, getKey?_insertManyIfNewUnit_list_of_contains_eq_false WF.emptyc h'] + simp_to_raw using Raw₀.Const.getKey?_insertManyIfNewUnit_empty_list_of_contains_eq_false theorem getKey?_unitOfList_of_mem [EquivBEq α] [LawfulHashable α] {l : List α} {k k' : α} (k_beq : k == k') (distinct : l.Pairwise (fun a b => (a == b) = false)) (mem : k ∈ l) : getKey? (unitOfList l) k' = some k := by - simp [unitOfList, getKey?_insertManyIfNewUnit_list_of_mem_of_contains_eq_false WF.emptyc k_beq - distinct mem contains_emptyc] + simp_to_raw using Raw₀.Const.getKey?_insertManyIfNewUnit_empty_list_of_mem theorem getKey_unitOfList_of_mem [EquivBEq α] [LawfulHashable α] {l : List α} @@ -1881,79 +1847,71 @@ theorem getKey_unitOfList_of_mem [EquivBEq α] [LawfulHashable α] (distinct : l.Pairwise (fun a b => (a == b) = false)) (mem : k ∈ l) {h'} : getKey (unitOfList l) k' h' = k := by - simp [unitOfList, getKey_insertManyIfNewUnit_list_of_mem_of_contains_eq_false WF.emptyc k_beq - distinct mem contains_emptyc] + simp_to_raw using Raw₀.Const.getKey_insertManyIfNewUnit_empty_list_of_mem theorem getKey!_unitOfList_of_contains_eq_false [EquivBEq α] [LawfulHashable α] [Inhabited α] {l : List α} {k : α} (h' : l.contains k = false) : getKey! (unitOfList l) k = default := by - simp [unitOfList, getKey!_insertManyIfNewUnit_list_of_contains_eq_false WF.emptyc h'] + simp_to_raw using Raw₀.Const.getKey!_insertManyIfNewUnit_empty_list_of_contains_eq_false theorem getKey!_unitOfList_of_mem [EquivBEq α] [LawfulHashable α] [Inhabited α] {l : List α} {k k' : α} (k_beq : k == k') (distinct : l.Pairwise (fun a b => (a == b) = false)) (mem : k ∈ l) : getKey! (unitOfList l) k' = k := by - simp [unitOfList, getKey!_insertManyIfNewUnit_list_of_mem_of_contains_eq_false WF.emptyc k_beq - distinct mem contains_emptyc] + simp_to_raw using Raw₀.Const.getKey!_insertManyIfNewUnit_empty_list_of_mem theorem getKeyD_unitOfList_of_contains_eq_false [EquivBEq α] [LawfulHashable α] {l : List α} {k fallback : α} (h' : l.contains k = false) : getKeyD (unitOfList l) k fallback = fallback := by - simp [unitOfList, getKeyD_insertManyIfNewUnit_list_of_contains_eq_false WF.emptyc h'] + simp_to_raw using Raw₀.Const.getKeyD_insertManyIfNewUnit_empty_list_of_contains_eq_false theorem getKeyD_unitOfList_of_mem [EquivBEq α] [LawfulHashable α] {l : List α} {k k' fallback : α} (k_beq : k == k') (distinct : l.Pairwise (fun a b => (a == b) = false)) {mem : k ∈ l } : getKeyD (unitOfList l) k' fallback = k := by - simp [unitOfList, getKeyD_insertManyIfNewUnit_list_of_mem_of_contains_eq_false WF.emptyc k_beq - distinct mem contains_emptyc] + simp_to_raw using Raw₀.Const.getKeyD_insertManyIfNewUnit_empty_list_of_mem theorem size_unitOfList [EquivBEq α] [LawfulHashable α] {l : List α} (distinct : l.Pairwise (fun a b => (a == b) = false)) : (unitOfList l).size = l.length := by - simp only [unitOfList] - rw [size_insertManyIfNewUnit_list WF.emptyc distinct] - · simp - · simp + simp_to_raw using Raw₀.Const.size_insertManyIfNewUnit_empty_list theorem size_unitOfList_le [EquivBEq α] [LawfulHashable α] {l : List α} : (unitOfList l).size ≤ l.length := by - simp only [unitOfList] - apply Nat.le_trans (size_insertManyIfNewUnit_list_le WF.emptyc) - simp - + simp_to_raw using Raw₀.Const.size_insertManyIfNewUnit_empty_list_le +@[simp] theorem isEmpty_unitOfList [EquivBEq α] [LawfulHashable α] {l : List α} : (unitOfList l).isEmpty = l.isEmpty := by - simp only [unitOfList] - rw [isEmpty_insertManyIfNewUnit_list WF.emptyc] - simp + simp_to_raw using Raw₀.Const.isEmpty_insertManyIfNewUnit_empty_list +@[simp] theorem get?_unitOfList [EquivBEq α] [LawfulHashable α] {l : List α} {k : α} : get? (unitOfList l) k = if l.contains k then some () else none := by - simp only [unitOfList] - rw [get?_insertManyIfNewUnit_list WF.emptyc] - simp + simp_to_raw using Raw₀.Const.get?_insertManyIfNewUnit_empty_list +@[simp] theorem get_unitOfList {l : List α} {k : α} {h} : - get (unitOfList l) k h = () := by + get (unitOfList l) k h = () := by simp +@[simp] theorem get!_unitOfList {l : List α} {k : α} : - get! (unitOfList l) k = () := by + get! (unitOfList l) k = () := by simp +@[simp] theorem getD_unitOfList {l : List α} {k : α} {fallback : Unit} : getD (unitOfList l) k fallback = () := by From 78c35f12709694b1ba90f0113bc9c7daf186e574 Mon Sep 17 00:00:00 2001 From: jt0202 Date: Mon, 6 Jan 2025 11:33:03 +0100 Subject: [PATCH 63/83] Small style fix --- src/Std/Data/DHashMap/RawLemmas.lean | 10 +++++----- 1 file changed, 5 insertions(+), 5 deletions(-) diff --git a/src/Std/Data/DHashMap/RawLemmas.lean b/src/Std/Data/DHashMap/RawLemmas.lean index 0012aeab1a04..9195be17d9c8 100644 --- a/src/Std/Data/DHashMap/RawLemmas.lean +++ b/src/Std/Data/DHashMap/RawLemmas.lean @@ -1542,8 +1542,8 @@ theorem ofList_singleton {k : α} {v : β k} [EquivBEq α] [LawfulHashable α] : rw [Raw₀.insertMany_empty_list_singleton] -theorem ofList_cons [EquivBEq α] [LawfulHashable α] {hd : (a : α) × (β a)} {tl : List ((a : α) × (β a))} : - ofList (hd::tl) = ((∅ : Raw α β).insert hd.1 hd.2).insertMany tl := by +theorem ofList_cons [EquivBEq α] [LawfulHashable α] {k : α} {v : β k} {tl : List ((a : α) × (β a))} : + ofList (⟨k, v⟩ :: tl) = ((∅ : Raw α β).insert k v).insertMany tl := by simp_to_raw rw [Raw₀.insertMany_empty_list_cons] @@ -1682,8 +1682,8 @@ theorem ofList_singleton {k : α} {v : β} [EquivBEq α] [LawfulHashable α] : simp_to_raw simp -theorem ofList_cons [EquivBEq α] [LawfulHashable α] {hd : α × β} {tl : List (α × β)} : - ofList (hd::tl) = insertMany ((∅ : Raw α (fun _ => β)).insert hd.1 hd.2) tl := by +theorem ofList_cons [EquivBEq α] [LawfulHashable α] {k : α} {v : β} {tl : List (α × β)} : + ofList (⟨k, v⟩ :: tl) = insertMany ((∅ : Raw α (fun _ => β)).insert k v) tl := by simp_to_raw rw [Raw₀.Const.insertMany_empty_list_cons] @@ -1820,7 +1820,7 @@ theorem unitOfList_singleton [EquivBEq α] [LawfulHashable α] {k : α} : simp theorem unitOfList_cons [EquivBEq α] [LawfulHashable α] {hd : α} {tl : List α} : - unitOfList (hd::tl) = insertManyIfNewUnit ((∅ : Raw α (fun _ => Unit)).insertIfNew hd ()) tl := by + unitOfList (hd :: tl) = insertManyIfNewUnit ((∅ : Raw α (fun _ => Unit)).insertIfNew hd ()) tl := by simp_to_raw rw [Raw₀.Const.insertManyIfNewUnit_empty_list_cons] From eee46468cdb19bfa28356a03af2d2846678412e7 Mon Sep 17 00:00:00 2001 From: jt0202 Date: Wed, 8 Jan 2025 10:41:52 +0100 Subject: [PATCH 64/83] Style fix --- .../DHashMap/Internal/List/Associative.lean | 2 +- src/Std/Data/DHashMap/Internal/Model.lean | 1 + src/Std/Data/DHashMap/Internal/RawLemmas.lean | 61 ++++++++++--------- src/Std/Data/DHashMap/Internal/WF.lean | 30 ++++----- src/Std/Data/DHashMap/Lemmas.lean | 57 +++++++++-------- src/Std/Data/DHashMap/RawLemmas.lean | 4 +- 6 files changed, 78 insertions(+), 77 deletions(-) diff --git a/src/Std/Data/DHashMap/Internal/List/Associative.lean b/src/Std/Data/DHashMap/Internal/List/Associative.lean index 2a77d4aacfd8..a85566438c08 100644 --- a/src/Std/Data/DHashMap/Internal/List/Associative.lean +++ b/src/Std/Data/DHashMap/Internal/List/Associative.lean @@ -923,7 +923,7 @@ theorem containsKey_eq_false_iff_forall_mem_keys [BEq α] [PartialEquivBEq α] (containsKey a l) = false ↔ ∀ a' ∈ keys l, (a == a') = false := by simp only [Bool.eq_false_iff, ne_eq, containsKey_iff_exists, not_exists, not_and] -theorem containsKey_eq_false_iff [BEq α][PartialEquivBEq α] {l : List ((a : α) × β a)} {a : α} : +theorem containsKey_eq_false_iff [BEq α] [PartialEquivBEq α] {l : List ((a : α) × β a)} {a : α} : containsKey a l = false ↔ ∀ (b : ((a : α) × β a)), b ∈ l → (a == b.fst) = false := by simp [containsKey_eq_false_iff_forall_mem_keys, keys_eq_map] diff --git a/src/Std/Data/DHashMap/Internal/Model.lean b/src/Std/Data/DHashMap/Internal/Model.lean index a98dc78a218e..cfffc0ed0f72 100644 --- a/src/Std/Data/DHashMap/Internal/Model.lean +++ b/src/Std/Data/DHashMap/Internal/Model.lean @@ -543,6 +543,7 @@ theorem Const.insertManyIfNewUnit_eq_insertListIfNewUnitₘ [BEq α] [Hashable | cons hd tl ih => simp only [List.foldl_cons,insertListIfNewUnitₘ] apply ih + end end Raw₀ diff --git a/src/Std/Data/DHashMap/Internal/RawLemmas.lean b/src/Std/Data/DHashMap/Internal/RawLemmas.lean index c942d16c70d6..a545538cc0b4 100644 --- a/src/Std/Data/DHashMap/Internal/RawLemmas.lean +++ b/src/Std/Data/DHashMap/Internal/RawLemmas.lean @@ -1001,8 +1001,8 @@ theorem isEmpty_insertMany_list [EquivBEq α] [LawfulHashable α] (h : m.1.WF) (m.insertMany l).1.1.isEmpty = (m.1.isEmpty && l.isEmpty) := by simp_to_model using isEmpty_insertList - namespace Const + variable {β : Type v} (m : Raw₀ α (fun _ => β)) @[simp] @@ -1118,8 +1118,7 @@ theorem get?_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable theorem get?_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l : List (α × β)} {k k' : α} (k_beq : k == k') {v : β} - (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) (mem : ⟨k, v⟩ ∈ l) - : + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) (mem : ⟨k, v⟩ ∈ l) : get? (insertMany m l).1 k' = v := by simp_to_model using getValue?_insertListConst_of_mem @@ -1144,7 +1143,7 @@ theorem get!_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable theorem get!_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] [Inhabited β] (h : m.1.WF) {l : List (α × β)} {k k' : α} (k_beq : k == k') {v : β} - (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) (mem : ⟨k, v⟩ ∈ l): + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) (mem : ⟨k, v⟩ ∈ l) : get! (insertMany m l).1 k' = v := by simp_to_model using getValue!_insertListConst_of_mem @@ -1219,8 +1218,7 @@ theorem getKey_insertManyIfNewUnit_list_of_mem_of_contains_eq_false [EquivBEq α {k k' : α} (k_beq : k == k') (distinct : l.Pairwise (fun a b => (a == b) = false)) (mem : k ∈ l) {h'} : - m.contains k = false → - getKey (insertManyIfNewUnit m l).1 k' h' = k := by + m.contains k = false → getKey (insertManyIfNewUnit m l).1 k' h' = k := by simp_to_model using getKey_insertListIfNewUnit_of_mem_of_contains_eq_false theorem getKey_insertManyIfNewUnit_list_mem_of_contains_of_contains [EquivBEq α] [LawfulHashable α] @@ -1239,15 +1237,15 @@ theorem getKey!_insertManyIfNewUnit_list_of_contains_eq_false [EquivBEq α] [Law theorem getKey!_insertManyIfNewUnit_list_of_mem_of_contains_eq_false [EquivBEq α] [LawfulHashable α] [Inhabited α] (h : m.1.WF) {l : List α} {k k' : α} (k_beq : k == k') (distinct : l.Pairwise (fun a b => (a == b) = false)) - (mem : k ∈ l) : contains m k = false → - getKey! (insertManyIfNewUnit m l).1 k' = k := by + (mem : k ∈ l) : + contains m k = false → getKey! (insertManyIfNewUnit m l).1 k' = k := by simp_to_model using getKey!_insertListIfNewUnit_of_mem_of_contains_eq_false theorem getKey!_insertManyIfNewUnit_list_mem_of_contains_of_contains [EquivBEq α] [LawfulHashable α] [Inhabited α] (h : m.1.WF) {l : List α} {k : α} (distinct : l.Pairwise (fun a b => (a == b) = false)) - (h' : l.contains k) : m.contains k → - getKey! (insertManyIfNewUnit m l).1 k = getKey! m k := by + (h' : l.contains k) : + m.contains k → getKey! (insertManyIfNewUnit m l).1 k = getKey! m k := by simp_to_model using getKey!_insertListIfNewUnit_of_contains_of_contains theorem getKeyD_insertManyIfNewUnit_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] @@ -1259,15 +1257,15 @@ theorem getKeyD_insertManyIfNewUnit_list_of_contains_eq_false [EquivBEq α] [Law theorem getKeyD_insertManyIfNewUnit_list_of_mem_of_contains_eq_false [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l : List α} {k k' fallback : α} (k_beq : k == k') (distinct : l.Pairwise (fun a b => (a == b) = false)) - (mem : k ∈ l ) : m.contains k = false → - getKeyD (insertManyIfNewUnit m l).1 k' fallback = k := by + (mem : k ∈ l ) : + m.contains k = false → getKeyD (insertManyIfNewUnit m l).1 k' fallback = k := by simp_to_model using getKeyD_insertListIfNewUnit_of_mem_of_contains_eq_false theorem getKeyD_insertManyIfNewUnit_list_of_contains_of_contains [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l : List α} {k fallback : α} (distinct : l.Pairwise (fun a b => (a == b) = false)) - (h' : l.contains k) : m.contains k → - getKeyD (insertManyIfNewUnit m l).1 k fallback = getKeyD m k fallback := by + (h' : l.contains k) : + m.contains k → getKeyD (insertManyIfNewUnit m l).1 k fallback = getKeyD m k fallback := by simp_to_model using getKeyD_insertListIfNewUnit_of_contains_of_contains theorem size_insertManyIfNewUnit_list [EquivBEq α] [LawfulHashable α] (h : m.1.WF) @@ -1310,10 +1308,13 @@ theorem getD_insertManyIfNewUnit_list simp end Const + end Raw₀ namespace Raw₀ + variable [BEq α] [Hashable α] + @[simp] theorem insertMany_empty_list_nil [EquivBEq α] [LawfulHashable α] : (insertMany empty ([] : List ((a : α) × (β a)))).1 = empty := by @@ -1324,7 +1325,8 @@ theorem insertMany_empty_list_singleton {k : α} {v : β k} [EquivBEq α] [Lawfu (insertMany empty [⟨k, v⟩]).1 = empty.insert k v := by simp -theorem insertMany_empty_list_cons [EquivBEq α] [LawfulHashable α] {k : α} {v : β k} {tl : List ((a : α) × (β a))} : +theorem insertMany_empty_list_cons [EquivBEq α] [LawfulHashable α] {k : α} {v : β k} + {tl : List ((a : α) × (β a))} : (insertMany empty (⟨k, v⟩ :: tl)).1 = ((empty.insert k v).insertMany tl).1 := by rw [insertMany_cons] @@ -1358,7 +1360,7 @@ theorem get!_insertMany_empty_list_of_contains_eq_false [LawfulBEq α] {l : List ((a : α) × β a)} {k : α} [Inhabited (β k)] (h : (l.map Sigma.fst).contains k = false) : (insertMany empty l).1.get! k = default := by - simp [get!_insertMany_list_of_contains_eq_false _ Raw.WF.empty₀ h] + simp only [get!_insertMany_list_of_contains_eq_false _ Raw.WF.empty₀ h] apply get!_empty theorem get!_insertMany_empty_list_of_mem [LawfulBEq α] @@ -1379,7 +1381,8 @@ theorem getD_insertMany_empty_list_of_mem [LawfulBEq α] {l : List ((a : α) × β a)} {k k' : α} (k_beq : k == k') {v : β k} {fallback : β k'} (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) (mem : ⟨k, v⟩ ∈ l) : - (insertMany empty l).1.getD k' fallback = cast (by congr; apply LawfulBEq.eq_of_beq k_beq) v := by + (insertMany empty l).1.getD k' fallback = + cast (by congr; apply LawfulBEq.eq_of_beq k_beq) v := by rw [getD_insertMany_list_of_mem _ Raw.WF.empty₀ k_beq distinct mem] theorem getKey?_insertMany_empty_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] @@ -1406,8 +1409,8 @@ theorem getKey_insertMany_empty_list_of_mem [EquivBEq α] [LawfulHashable α] (insertMany empty l).1.getKey k' h' = k := by rw [getKey_insertMany_list_of_mem _ Raw.WF.empty₀ k_beq distinct mem] -theorem getKey!_insertMany_empty_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] [Inhabited α] - {l : List ((a : α) × β a)} {k : α} +theorem getKey!_insertMany_empty_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] + [Inhabited α] {l : List ((a : α) × β a)} {k : α} (h : (l.map Sigma.fst).contains k = false) : (insertMany empty l).1.getKey! k = default := by rw [getKey!_insertMany_list_of_contains_eq_false _ Raw.WF.empty₀ h] @@ -1467,7 +1470,8 @@ theorem insertMany_empty_list_singleton {k : α} {v : β} [EquivBEq α] [LawfulH (insertMany empty [⟨k, v⟩]).1 = empty.insert k v := by simp only [insertMany_list_singleton] -theorem insertMany_empty_list_cons [EquivBEq α] [LawfulHashable α] {k : α} {v : β} {tl : List (α × β)} : +theorem insertMany_empty_list_cons [EquivBEq α] [LawfulHashable α] {k : α} {v : β} + {tl : List (α × β)} : (insertMany empty (⟨k, v⟩ :: tl)) = (insertMany (empty.insert k v) tl).1 := by rw [insertMany_cons] @@ -1496,14 +1500,12 @@ theorem get_insertMany_empty_list_of_mem [LawfulBEq α] (mem : ⟨k, v⟩ ∈ l) {h} : get (insertMany (empty : Raw₀ α (fun _ => β)) l) k' h = v := by - rw [get_insertMany_list_of_mem _ Raw.WF.empty₀ k_beq distinct mem] theorem get!_insertMany_empty_list_of_contains_eq_false [LawfulBEq α] {l : List (α × β)} {k : α} [Inhabited β] (h : (l.map Prod.fst).contains k = false) : get! (insertMany (empty : Raw₀ α (fun _ => β)) l) k = (default : β) := by - rw [get!_insertMany_list_of_contains_eq_false _ Raw.WF.empty₀ h] apply get!_empty @@ -1512,14 +1514,12 @@ theorem get!_insertMany_empty_list_of_mem [LawfulBEq α] (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) (mem : ⟨k, v⟩ ∈ l) : get! (insertMany (empty : Raw₀ α (fun _ => β)) l) k' = v := by - rw [get!_insertMany_list_of_mem _ Raw.WF.empty₀ k_beq distinct mem] theorem getD_insertMany_empty_list_of_contains_eq_false [LawfulBEq α] {l : List (α × β)} {k : α} {fallback : β} (contains_eq_false : (l.map Prod.fst).contains k = false) : getD (insertMany (empty : Raw₀ α (fun _ => β)) l) k fallback = fallback := by - rw [getD_insertMany_list_of_contains_eq_false _ Raw.WF.empty₀ contains_eq_false] apply getD_empty @@ -1552,7 +1552,6 @@ theorem getKey_insertMany_empty_list_of_mem [EquivBEq α] [LawfulHashable α] (mem : k ∈ l.map Prod.fst) {h'} : (insertMany (empty : Raw₀ α (fun _ => β)) l).1.getKey k' h' = k := by - rw [getKey_insertMany_list_of_mem _ Raw.WF.empty₀ k_beq distinct mem] theorem getKey!_insertMany_empty_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] @@ -1594,7 +1593,7 @@ theorem size_insertMany_empty_list [EquivBEq α] [LawfulHashable α] theorem size_insertMany_empty_list_le [EquivBEq α] [LawfulHashable α] {l : List (α × β)} : - (insertMany (empty : Raw₀ α (fun _ => β)) l).1.1.size ≤ l.length := by + (insertMany (empty : Raw₀ α (fun _ => β)) l).1.1.size ≤ l.length := by rw [← Nat.zero_add l.length] apply (size_insertMany_list_le _ Raw.WF.empty₀) @@ -1605,8 +1604,8 @@ theorem isEmpty_insertMany_empty_list [EquivBEq α] [LawfulHashable α] @[simp] theorem insertManyIfNewUnit_empty_list_nil [EquivBEq α] [LawfulHashable α] : - insertManyIfNewUnit (empty : Raw₀ α (fun _ => Unit)) ([] : List α) = - (empty : Raw₀ α (fun _ => Unit)) := by + insertManyIfNewUnit (empty : Raw₀ α (fun _ => Unit)) ([] : List α) = + (empty : Raw₀ α (fun _ => Unit)) := by simp @[simp] @@ -1616,7 +1615,7 @@ theorem insertManyIfNewUnit_empty_list_singleton [EquivBEq α] [LawfulHashable theorem insertManyIfNewUnit_empty_list_cons [EquivBEq α] [LawfulHashable α] {hd : α} {tl : List α} : insertManyIfNewUnit (empty : Raw₀ α (fun _ => Unit)) (hd :: tl) = - (insertManyIfNewUnit (empty.insertIfNew hd ()) tl).1 := by + (insertManyIfNewUnit (empty.insertIfNew hd ()) tl).1 := by rw [insertManyIfNewUnit_cons] theorem contains_insertManyIfNewUnit_empty_list [EquivBEq α] [LawfulHashable α] @@ -1696,7 +1695,7 @@ theorem isEmpty_insertManyIfNewUnit_empty_list [EquivBEq α] [LawfulHashable α] theorem get?_insertManyIfNewUnit_empty_list [EquivBEq α] [LawfulHashable α] {l : List α} {k : α} : get? (insertManyIfNewUnit (empty : Raw₀ α (fun _ => Unit)) l) k = - if l.contains k then some () else none := by + if l.contains k then some () else none := by rw [get?_insertManyIfNewUnit_list _ Raw.WF.empty₀] simp @@ -1716,5 +1715,7 @@ theorem getD_insertManyIfNewUnit_empty_list simp end Const + end Raw₀ + end Std.DHashMap.Internal diff --git a/src/Std/Data/DHashMap/Internal/WF.lean b/src/Std/Data/DHashMap/Internal/WF.lean index 29af63a6805b..853a488c675d 100644 --- a/src/Std/Data/DHashMap/Internal/WF.lean +++ b/src/Std/Data/DHashMap/Internal/WF.lean @@ -740,11 +740,8 @@ theorem toListModel_insertListₘ [BEq α] [Hashable α] [EquivBEq α][LawfulHas simp [insertListₘ, List.insertList] | cons hd tl ih => simp only [insertListₘ, List.insertList] - apply Perm.trans - apply ih (wfImp_insert h) - apply List.insertList_perm_of_perm_first - apply toListModel_insert h - apply (wfImp_insert h).distinct + apply Perm.trans (ih (wfImp_insert h)) + apply List.insertList_perm_of_perm_first (toListModel_insert h) (wfImp_insert h).distinct end Raw₀ @@ -778,30 +775,27 @@ theorem wfImp_insertMany [BEq α] [Hashable α] [EquivBEq α] [LawfulHashable α theorem toListModel_insertMany_list [BEq α] [Hashable α] [EquivBEq α] [LawfulHashable α] {m : Raw₀ α β} {l : List ((a : α) × (β a))} (h : Raw.WFImp m.1) : Perm (toListModel (insertMany m l).1.1.buckets) (List.insertList (toListModel m.1.buckets) l) := by - rw[insertMany_eq_insertListₘ] + rw [insertMany_eq_insertListₘ] apply toListModel_insertListₘ exact h /-! # `Const.insertListₘ` -/ theorem Const.toListModel_insertListₘ {β : Type v} [BEq α] [Hashable α] [EquivBEq α] - [LawfulHashable α] {m : Raw₀ α (fun _ => β)}{l : List (α × β)} (h : Raw.WFImp m.1): + [LawfulHashable α] {m : Raw₀ α (fun _ => β)} {l : List (α × β)} (h : Raw.WFImp m.1) : Perm (toListModel (Const.insertListₘ m l).1.buckets) (insertListConst (toListModel m.1.buckets) l) := by induction l generalizing m with | nil => simp [Const.insertListₘ, insertListConst, insertList] | cons hd tl ih => simp only [Const.insertListₘ, insertListConst] - apply Perm.trans - apply ih (wfImp_insert h) + apply Perm.trans (ih (wfImp_insert h)) unfold insertListConst - apply List.insertList_perm_of_perm_first - apply toListModel_insert h - apply (wfImp_insert h).distinct + apply List.insertList_perm_of_perm_first (toListModel_insert h) (wfImp_insert h).distinct /-! # `Const.insertMany` -/ theorem Const.toListModel_insertMany_list {β : Type v} [BEq α] [Hashable α] [EquivBEq α] - [LawfulHashable α] {m : Raw₀ α (fun _ => β)} {l : List (α × β)} (h : Raw.WFImp m.1): + [LawfulHashable α] {m : Raw₀ α (fun _ => β)} {l : List (α × β)} (h : Raw.WFImp m.1) : Perm (toListModel (Const.insertMany m l).1.1.buckets) (insertListConst (toListModel m.1.buckets) l) := by rw [Const.insertMany_eq_insertListₘ] apply toListModel_insertListₘ h @@ -814,23 +808,21 @@ theorem Const.wfImp_insertMany {β : Type v} [BEq α] [Hashable α] [EquivBEq α /-! # `Const.insertListIfNewUnitₘ` -/ theorem Const.toListModel_insertListIfNewUnitₘ [BEq α] [Hashable α] [EquivBEq α] [LawfulHashable α] - {m : Raw₀ α (fun _ => Unit)} {l : List α} (h : Raw.WFImp m.1): + {m : Raw₀ α (fun _ => Unit)} {l : List α} (h : Raw.WFImp m.1) : Perm (toListModel (Const.insertListIfNewUnitₘ m l).1.buckets) (List.insertListIfNewUnit (toListModel m.1.buckets) l) := by induction l generalizing m with | nil => simp [insertListIfNewUnitₘ, List.insertListIfNewUnit] | cons hd tl ih => simp only [insertListIfNewUnitₘ, insertListIfNewUnit] - apply Perm.trans - apply ih (wfImp_insertIfNew h) - apply List.insertListIfNewUnit_perm_of_perm_first - apply toListModel_insertIfNew h + apply Perm.trans (ih (wfImp_insertIfNew h)) + apply List.insertListIfNewUnit_perm_of_perm_first (toListModel_insertIfNew h) apply (wfImp_insertIfNew h).distinct /-! # `Const.insertManyIfNewUnit` -/ theorem Const.toListModel_insertManyIfNewUnit_list [BEq α] [Hashable α] [EquivBEq α] - [LawfulHashable α] {m : Raw₀ α (fun _ => Unit)} {l : List α} (h : Raw.WFImp m.1): + [LawfulHashable α] {m : Raw₀ α (fun _ => Unit)} {l : List α} (h : Raw.WFImp m.1) : Perm (toListModel (Const.insertManyIfNewUnit m l).1.1.buckets) (List.insertListIfNewUnit (toListModel m.1.buckets) l) := by rw [Const.insertManyIfNewUnit_eq_insertListIfNewUnitₘ] diff --git a/src/Std/Data/DHashMap/Lemmas.lean b/src/Std/Data/DHashMap/Lemmas.lean index 70cdaeb52349..35b2270a38e0 100644 --- a/src/Std/Data/DHashMap/Lemmas.lean +++ b/src/Std/Data/DHashMap/Lemmas.lean @@ -913,7 +913,6 @@ theorem getKeyD_insertIfNew [EquivBEq α] [LawfulHashable α] {k a fallback : α simp [mem_iff_contains, contains_insertIfNew] exact Raw₀.getKeyD_insertIfNew ⟨m.1, _⟩ m.2 - @[simp] theorem getThenInsertIfNew?_fst [LawfulBEq α] {k : α} {v : β k} : (m.getThenInsertIfNew? k v).1 = m.get? k := @@ -982,7 +981,7 @@ theorem insertMany_cons {l : List ((a : α) × β a)} {k : α} {v : β k} : theorem contains_insertMany_list [EquivBEq α] [LawfulHashable α] {l : List ((a : α) × β a)} {k : α} : (m.insertMany l).contains k = (m.contains k || (l.map Sigma.fst).contains k) := - Raw₀.contains_insertMany_list ⟨m.1, _⟩ m.2 + Raw₀.contains_insertMany_list ⟨m.1, _⟩ m.2 theorem contains_of_contains_insertMany_list [EquivBEq α] [LawfulHashable α] {l : List ((a : α) × β a)} {k : α} : @@ -1000,7 +999,7 @@ theorem get?_insertMany_list_of_mem [LawfulBEq α] (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) (mem : ⟨k, v⟩ ∈ l) : (m.insertMany l).get? k' = some (cast (by congr; apply LawfulBEq.eq_of_beq k_beq) v) := - Raw₀.get?_insertMany_list_of_mem ⟨m.1, _⟩ m.2 k_beq distinct mem + Raw₀.get?_insertMany_list_of_mem ⟨m.1, _⟩ m.2 k_beq distinct mem theorem get_insertMany_list_of_contains_eq_false [LawfulBEq α] {l : List ((a : α) × β a)} {k : α} @@ -1016,7 +1015,7 @@ theorem get_insertMany_list_of_mem [LawfulBEq α] (mem : ⟨k, v⟩ ∈ l) {h'} : (m.insertMany l).get k' h' = cast (by congr; apply LawfulBEq.eq_of_beq k_beq) v := - Raw₀.get_insertMany_list_of_mem ⟨m.1, _⟩ m.2 k_beq distinct mem + Raw₀.get_insertMany_list_of_mem ⟨m.1, _⟩ m.2 k_beq distinct mem theorem get!_insertMany_list_of_contains_eq_false [LawfulBEq α] {l : List ((a : α) × β a)} {k : α} [Inhabited (β k)] @@ -1029,7 +1028,7 @@ theorem get!_insertMany_list_of_mem [LawfulBEq α] (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) (mem : ⟨k, v⟩ ∈ l) : (m.insertMany l).get! k' = cast (by congr; apply LawfulBEq.eq_of_beq k_beq) v := - Raw₀.get!_insertMany_list_of_mem ⟨m.1, _⟩ m.2 k_beq distinct mem + Raw₀.get!_insertMany_list_of_mem ⟨m.1, _⟩ m.2 k_beq distinct mem theorem getD_insertMany_list_of_contains_eq_false [LawfulBEq α] {l : List ((a : α) × β a)} {k : α} {fallback : β k} @@ -1042,7 +1041,7 @@ theorem getD_insertMany_list_of_mem [LawfulBEq α] (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) (mem : ⟨k, v⟩ ∈ l) : (m.insertMany l).getD k' fallback = cast (by congr; apply LawfulBEq.eq_of_beq k_beq) v := - Raw₀.getD_insertMany_list_of_mem ⟨m.1, _⟩ m.2 k_beq distinct mem + Raw₀.getD_insertMany_list_of_mem ⟨m.1, _⟩ m.2 k_beq distinct mem theorem getKey?_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] {l : List ((a : α) × β a)} {k : α} @@ -1076,7 +1075,7 @@ theorem getKey_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] Raw₀.getKey_insertMany_list_of_mem ⟨m.1, _⟩ m.2 k_beq distinct mem theorem getKey!_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] [Inhabited α] - {l : List ((a : α) × β a)} {k : α} + {l : List ((a : α) × β a)} {k : α} (h' : (l.map Sigma.fst).contains k = false) : (m.insertMany l).getKey! k = m.getKey! k := Raw₀.getKey!_insertMany_list_of_contains_eq_false ⟨m.1, _⟩ m.2 h' @@ -1121,21 +1120,23 @@ theorem isEmpty_insertMany_list [EquivBEq α] [LawfulHashable α] Raw₀.isEmpty_insertMany_list ⟨m.1, _⟩ m.2 namespace Const + variable {β : Type v} (m : DHashMap α (fun _ => β)) @[simp] theorem insertMany_nil : insertMany m [] = m := - Subtype.eq (congrArg Subtype.val (Raw₀.Const.insertMany_nil ⟨m.1, m.2.size_buckets_pos⟩) :) + Subtype.eq (congrArg Subtype.val (Raw₀.Const.insertMany_nil ⟨m.1, m.2.size_buckets_pos⟩) :) @[simp] theorem insertMany_list_singleton {k : α} {v : β} : insertMany m [⟨k, v⟩] = m.insert k v := - Subtype.eq (congrArg Subtype.val (Raw₀.Const.insertMany_list_singleton ⟨m.1, m.2.size_buckets_pos⟩) :) + Subtype.eq (congrArg Subtype.val + (Raw₀.Const.insertMany_list_singleton ⟨m.1, m.2.size_buckets_pos⟩) :) theorem insertMany_cons {l : List (α × β)} {k : α} {v : β} : insertMany m (⟨k, v⟩ :: l) = insertMany (m.insert k v) l := - Subtype.eq (congrArg Subtype.val (Raw₀.Const.insertMany_cons ⟨m.1, m.2.size_buckets_pos⟩) :) + Subtype.eq (congrArg Subtype.val (Raw₀.Const.insertMany_cons ⟨m.1, m.2.size_buckets_pos⟩) :) @[simp] theorem contains_insertMany_list [EquivBEq α] [LawfulHashable α] @@ -1180,7 +1181,7 @@ theorem getKey_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] Raw₀.Const.getKey_insertMany_list_of_mem ⟨m.1, _⟩ m.2 k_beq distinct mem theorem getKey!_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] [Inhabited α] - {l : List (α × β)} {k : α} + {l : List (α × β)} {k : α} (h' : (l.map Prod.fst).contains k = false) : (insertMany m l).getKey! k = m.getKey! k := Raw₀.Const.getKey!_insertMany_list_of_contains_eq_false ⟨m.1, _⟩ m.2 h' @@ -1233,8 +1234,7 @@ theorem get?_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable theorem get?_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] {l : List (α × β)} {k k' : α} (k_beq : k == k') {v : β} - (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) (mem : ⟨k, v⟩ ∈ l) - : + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) (mem : ⟨k, v⟩ ∈ l) : get? (insertMany m l) k' = v := Raw₀.Const.get?_insertMany_list_of_mem ⟨m.1, _⟩ m.2 k_beq distinct mem @@ -1259,7 +1259,7 @@ theorem get!_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable theorem get!_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] [Inhabited β] {l : List (α × β)} {k k' : α} (k_beq : k == k') {v : β} - (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) (mem : ⟨k, v⟩ ∈ l): + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) (mem : ⟨k, v⟩ ∈ l) : get! (insertMany m l) k' = v := Raw₀.Const.get!_insertMany_list_of_mem ⟨m.1, _⟩ m.2 k_beq distinct mem @@ -1279,17 +1279,19 @@ variable (m : DHashMap α (fun _ => Unit)) @[simp] theorem insertManyIfNewUnit_nil : insertManyIfNewUnit m [] = m := - Subtype.eq (congrArg Subtype.val (Raw₀.Const.insertManyIfNewUnit_nil ⟨m.1, m.2.size_buckets_pos⟩) :) + Subtype.eq (congrArg Subtype.val + (Raw₀.Const.insertManyIfNewUnit_nil ⟨m.1, m.2.size_buckets_pos⟩) :) @[simp] theorem insertManyIfNewUnit_list_singleton [EquivBEq α] [LawfulHashable α] {k : α} : insertManyIfNewUnit m [k] = m.insertIfNew k () := Subtype.eq (congrArg Subtype.val - (Raw₀.Const.insertManyIfNewUnit_list_singleton ⟨m.1, m.2.size_buckets_pos⟩) :) + (Raw₀.Const.insertManyIfNewUnit_list_singleton ⟨m.1, m.2.size_buckets_pos⟩) :) theorem insertManyIfNewUnit_cons {l : List α} {k : α} : insertManyIfNewUnit m (k :: l) = insertManyIfNewUnit (m.insertIfNew k ()) l := - Subtype.eq (congrArg Subtype.val (Raw₀.Const.insertManyIfNewUnit_cons ⟨m.1, m.2.size_buckets_pos⟩) :) + Subtype.eq (congrArg Subtype.val + (Raw₀.Const.insertManyIfNewUnit_cons ⟨m.1, m.2.size_buckets_pos⟩) :) @[simp] theorem contains_insertManyIfNewUnit_list [EquivBEq α] [LawfulHashable α] @@ -1303,32 +1305,32 @@ theorem contains_of_contains_insertManyIfNewUnit_list [EquivBEq α] [LawfulHasha Raw₀.Const.contains_of_contains_insertManyIfNewUnit_list ⟨m.1, _⟩ m.2 h₂ theorem getKey?_insertManyIfNewUnit_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] - {l : List α} {k : α} (h' : l.contains k = false) : + {l : List α} {k : α} (h' : l.contains k = false) : getKey? (insertManyIfNewUnit m l) k = getKey? m k := Raw₀.Const.getKey?_insertManyIfNewUnit_list_of_contains_eq_false ⟨m.1, _⟩ m.2 h' theorem getKey?_insertManyIfNewUnit_list_of_mem_of_contains_eq_false [EquivBEq α] [LawfulHashable α] - {l : List α} {k k' : α} (k_beq : k == k') + {l : List α} {k k' : α} (k_beq : k == k') (distinct : l.Pairwise (fun a b => (a == b) = false)) (mem : k ∈ l) : m.contains k = false → getKey? (insertManyIfNewUnit m l) k' = some k := Raw₀.Const.getKey?_insertManyIfNewUnit_list_of_mem_of_contains_eq_false ⟨m.1, _⟩ m.2 k_beq distinct mem theorem getKey?_insertManyIfNewUnit_list_of_contains_of_contains [EquivBEq α] [LawfulHashable α] - {l : List α} {k : α} + {l : List α} {k : α} (distinct : l.Pairwise (fun a b => (a == b) = false)) (h' : l.contains k) : m.contains k → getKey? (insertManyIfNewUnit m l) k = getKey? m k := Raw₀.Const.getKey?_insertManyIfNewUnit_list_of_contains_of_contains ⟨m.1, _⟩ m.2 distinct h' theorem getKey_insertManyIfNewUnit_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] - {l : List α} {k : α} (h₁ : l.contains k = false) {h'} : + {l : List α} {k : α} (h₁ : l.contains k = false) {h'} : getKey (insertManyIfNewUnit m l) k h' = getKey m k (contains_of_contains_insertManyIfNewUnit_list m h₁ h') := Raw₀.Const.getKey_insertManyIfNewUnit_list_of_contains_eq_false ⟨m.1, _⟩ m.2 h₁ theorem getKey_insertManyIfNewUnit_list_of_mem_of_contains_eq_false [EquivBEq α] [LawfulHashable α] - {l : List α} + {l : List α} {k k' : α} (k_beq : k == k') (distinct : l.Pairwise (fun a b => (a == b) = false)) (mem : k ∈ l) {h'} : @@ -1338,7 +1340,7 @@ theorem getKey_insertManyIfNewUnit_list_of_mem_of_contains_eq_false [EquivBEq α distinct mem theorem getKey_insertManyIfNewUnit_list_mem_of_contains_of_contains [EquivBEq α] [LawfulHashable α] - {l : List α} {k : α} + {l : List α} {k : α} (distinct : l.Pairwise (fun a b => (a == b) = false)) (h₁ : l.contains k) {mem' : m.contains k} {h'} : getKey (insertManyIfNewUnit m l) k h' = getKey m k mem' := @@ -1366,13 +1368,13 @@ theorem getKey!_insertManyIfNewUnit_list_mem_of_contains_of_contains [EquivBEq Raw₀.Const.getKey!_insertManyIfNewUnit_list_mem_of_contains_of_contains ⟨m.1, _⟩ m.2 distinct h' theorem getKeyD_insertManyIfNewUnit_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] - {l : List α} {k fallback : α} + {l : List α} {k fallback : α} (h' : l.contains k = false) : getKeyD (insertManyIfNewUnit m l) k fallback = getKeyD m k fallback := Raw₀.Const.getKeyD_insertManyIfNewUnit_list_of_contains_eq_false ⟨m.1, _⟩ m.2 h' theorem getKeyD_insertManyIfNewUnit_list_of_mem_of_contains_eq_false [EquivBEq α] [LawfulHashable α] - {l : List α} {k k' fallback : α} (k_beq : k == k') + {l : List α} {k k' fallback : α} (k_beq : k == k') (distinct : l.Pairwise (fun a b => (a == b) = false)) (mem : k ∈ l ) : m.contains k = false → getKeyD (insertManyIfNewUnit m l) k' fallback = k := @@ -1380,7 +1382,7 @@ theorem getKeyD_insertManyIfNewUnit_list_of_mem_of_contains_eq_false [EquivBEq distinct mem theorem getKeyD_insertManyIfNewUnit_list_of_contains_of_contains [EquivBEq α] [LawfulHashable α] - {l : List α} {k fallback : α} + {l : List α} {k fallback : α} (distinct : l.Pairwise (fun a b => (a == b) = false)) (h' : l.contains k) : m.contains k → getKeyD (insertManyIfNewUnit m l) k fallback = getKeyD m k fallback := @@ -1429,6 +1431,7 @@ end Const end DHashMap namespace DHashMap + @[simp] theorem ofList_nil [EquivBEq α] [LawfulHashable α] : ofList ([] : List ((a : α) × (β a))) = ∅ := @@ -1564,7 +1567,9 @@ theorem isEmpty_ofList[EquivBEq α] [LawfulHashable α] Raw₀.isEmpty_insertMany_empty_list namespace Const + variable {β : Type v} + @[simp] theorem ofList_nil [EquivBEq α] [LawfulHashable α] : ofList ([] : List (α × β)) = ∅ := diff --git a/src/Std/Data/DHashMap/RawLemmas.lean b/src/Std/Data/DHashMap/RawLemmas.lean index 9195be17d9c8..eaa5848476a8 100644 --- a/src/Std/Data/DHashMap/RawLemmas.lean +++ b/src/Std/Data/DHashMap/RawLemmas.lean @@ -1212,6 +1212,7 @@ theorem isEmpty_insertMany_list [EquivBEq α] [LawfulHashable α] (h : m.WF) simp_to_raw using Raw₀.isEmpty_insertMany_list namespace Const + variable {β : Type v} {m : Raw α (fun _ => β)} @[simp] @@ -1527,7 +1528,9 @@ end Const end Raw namespace Raw + variable [BEq α] [Hashable α] + open Internal.Raw Internal.Raw₀ @[simp] theorem ofList_nil [EquivBEq α] [LawfulHashable α] : @@ -1541,7 +1544,6 @@ theorem ofList_singleton {k : α} {v : β k} [EquivBEq α] [LawfulHashable α] : simp_to_raw rw [Raw₀.insertMany_empty_list_singleton] - theorem ofList_cons [EquivBEq α] [LawfulHashable α] {k : α} {v : β k} {tl : List ((a : α) × (β a))} : ofList (⟨k, v⟩ :: tl) = ((∅ : Raw α β).insert k v).insertMany tl := by simp_to_raw From b54f7be6acb1a79da7e4a140bdd56ed27cb63e39 Mon Sep 17 00:00:00 2001 From: jt0202 Date: Wed, 8 Jan 2025 10:46:19 +0100 Subject: [PATCH 65/83] Removed colons from two model functions --- src/Std/Data/DHashMap/Internal/Model.lean | 6 ++++-- 1 file changed, 4 insertions(+), 2 deletions(-) diff --git a/src/Std/Data/DHashMap/Internal/Model.lean b/src/Std/Data/DHashMap/Internal/Model.lean index cfffc0ed0f72..4487e60a7b36 100644 --- a/src/Std/Data/DHashMap/Internal/Model.lean +++ b/src/Std/Data/DHashMap/Internal/Model.lean @@ -349,13 +349,15 @@ def Const.get!ₘ [BEq α] [Hashable α] [Inhabited β] (m : Raw₀ α (fun _ => (Const.get?ₘ m a).get! /-- Internal implementation detail of the hash map -/ -def Const.insertListₘ [BEq α] [Hashable α](m : Raw₀ α (fun _ => β)) (l: List (α × β)): Raw₀ α (fun _ => β) := +def Const.insertListₘ [BEq α] [Hashable α](m : Raw₀ α (fun _ => β)) (l: List (α × β)) : + Raw₀ α (fun _ => β) := match l with | .nil => m | .cons hd tl => insertListₘ (m.insert hd.1 hd.2) tl /-- Internal implementation detail of the hash map -/ -def Const.insertListIfNewUnitₘ [BEq α] [Hashable α](m : Raw₀ α (fun _ => Unit)) (l: List α): Raw₀ α (fun _ => Unit) := +def Const.insertListIfNewUnitₘ [BEq α] [Hashable α](m : Raw₀ α (fun _ => Unit)) (l: List α) : + Raw₀ α (fun _ => Unit) := match l with | .nil => m | .cons hd tl => insertListIfNewUnitₘ (m.insertIfNew hd ()) tl From b51500d355f182e2b0c8e09296bcf5d0f7ae59fd Mon Sep 17 00:00:00 2001 From: jt0202 Date: Wed, 8 Jan 2025 11:54:18 +0100 Subject: [PATCH 66/83] Style fixes --- src/Std/Data/DHashMap/Lemmas.lean | 45 ++++++++++++++-------------- src/Std/Data/DHashMap/RawLemmas.lean | 38 +++++++++++------------ 2 files changed, 40 insertions(+), 43 deletions(-) diff --git a/src/Std/Data/DHashMap/Lemmas.lean b/src/Std/Data/DHashMap/Lemmas.lean index 35b2270a38e0..4a113c4cba2d 100644 --- a/src/Std/Data/DHashMap/Lemmas.lean +++ b/src/Std/Data/DHashMap/Lemmas.lean @@ -1006,7 +1006,7 @@ theorem get_insertMany_list_of_contains_eq_false [LawfulBEq α] (contains : (l.map Sigma.fst).contains k = false) {h'} : (m.insertMany l).get k h' = - m.get k (contains_of_contains_insertMany_list h' contains) := + m.get k (contains_of_contains_insertMany_list h' contains) := Raw₀.get_insertMany_list_of_contains_eq_false ⟨m.1, _⟩ m.2 contains theorem get_insertMany_list_of_mem [LawfulBEq α] @@ -1062,7 +1062,7 @@ theorem getKey_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashabl (h₁ : (l.map Sigma.fst).contains k = false) {h'} : (m.insertMany l).getKey k h' = - m.getKey k (contains_of_contains_insertMany_list h' h₁) := + m.getKey k (contains_of_contains_insertMany_list h' h₁) := Raw₀.getKey_insertMany_list_of_contains_eq_false ⟨m.1, _⟩ m.2 h₁ theorem getKey_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] @@ -1105,7 +1105,7 @@ theorem getKeyD_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] theorem size_insertMany_list [EquivBEq α] [LawfulHashable α] {l : List ((a : α) × β a)} (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) : (∀ (a : α), m.contains a → (l.map Sigma.fst).contains a = false) → - (m.insertMany l).size = m.size + l.length := + (m.insertMany l).size = m.size + l.length := Raw₀.size_insertMany_list ⟨m.1, _⟩ m.2 distinct theorem size_insertMany_list_le [EquivBEq α] [LawfulHashable α] @@ -1168,7 +1168,7 @@ theorem getKey_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashabl (h₁ : (l.map Prod.fst).contains k = false) {h'} : (insertMany m l).getKey k h' = - m.getKey k (contains_of_contains_insertMany_list _ h' h₁) := + m.getKey k (contains_of_contains_insertMany_list _ h' h₁) := Raw₀.Const.getKey_insertMany_list_of_contains_eq_false ⟨m.1, _⟩ m.2 h₁ theorem getKey_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] @@ -1212,7 +1212,7 @@ theorem size_insertMany_list [EquivBEq α] [LawfulHashable α] {l : List (α × β)} (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) : (∀ (a : α), m.contains a → (l.map Prod.fst).contains a = false) → - (insertMany m l).size = m.size + l.length := + (insertMany m l).size = m.size + l.length := Raw₀.Const.size_insertMany_list ⟨m.1, _⟩ m.2 distinct theorem size_insertMany_list_le [EquivBEq α] [LawfulHashable α] @@ -1276,6 +1276,7 @@ theorem getD_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] Raw₀.Const.getD_insertMany_list_of_mem ⟨m.1, _⟩ m.2 k_beq distinct mem variable (m : DHashMap α (fun _ => Unit)) + @[simp] theorem insertManyIfNewUnit_nil : insertManyIfNewUnit m [] = m := @@ -1334,10 +1335,9 @@ theorem getKey_insertManyIfNewUnit_list_of_mem_of_contains_eq_false [EquivBEq α {k k' : α} (k_beq : k == k') (distinct : l.Pairwise (fun a b => (a == b) = false)) (mem : k ∈ l) {h'} : - m.contains k = false → - getKey (insertManyIfNewUnit m l) k' h' = k := + m.contains k = false → getKey (insertManyIfNewUnit m l) k' h' = k := Raw₀.Const.getKey_insertManyIfNewUnit_list_of_mem_of_contains_eq_false ⟨m.1, _⟩ m.2 k_beq - distinct mem + distinct mem theorem getKey_insertManyIfNewUnit_list_mem_of_contains_of_contains [EquivBEq α] [LawfulHashable α] {l : List α} {k : α} @@ -1355,16 +1355,16 @@ theorem getKey!_insertManyIfNewUnit_list_of_contains_eq_false [EquivBEq α] [Law theorem getKey!_insertManyIfNewUnit_list_of_mem_of_contains_eq_false [EquivBEq α] [LawfulHashable α] [Inhabited α] {l : List α} {k k' : α} (k_beq : k == k') (distinct : l.Pairwise (fun a b => (a == b) = false)) - (mem : k ∈ l) : contains m k = false → - getKey! (insertManyIfNewUnit m l) k' = k := + (mem : k ∈ l) : + contains m k = false → getKey! (insertManyIfNewUnit m l) k' = k := Raw₀.Const.getKey!_insertManyIfNewUnit_list_of_mem_of_contains_eq_false ⟨m.1, _⟩ m.2 k_beq - distinct mem + distinct mem theorem getKey!_insertManyIfNewUnit_list_mem_of_contains_of_contains [EquivBEq α] [LawfulHashable α] [Inhabited α] {l : List α} {k : α} (distinct : l.Pairwise (fun a b => (a == b) = false)) - (h' : l.contains k) : m.contains k → - getKey! (insertManyIfNewUnit m l) k = getKey! m k := + (h' : l.contains k) : + m.contains k → getKey! (insertManyIfNewUnit m l) k = getKey! m k := Raw₀.Const.getKey!_insertManyIfNewUnit_list_mem_of_contains_of_contains ⟨m.1, _⟩ m.2 distinct h' theorem getKeyD_insertManyIfNewUnit_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] @@ -1376,23 +1376,23 @@ theorem getKeyD_insertManyIfNewUnit_list_of_contains_eq_false [EquivBEq α] [Law theorem getKeyD_insertManyIfNewUnit_list_of_mem_of_contains_eq_false [EquivBEq α] [LawfulHashable α] {l : List α} {k k' fallback : α} (k_beq : k == k') (distinct : l.Pairwise (fun a b => (a == b) = false)) - (mem : k ∈ l ) : m.contains k = false → - getKeyD (insertManyIfNewUnit m l) k' fallback = k := + (mem : k ∈ l ) : + m.contains k = false → getKeyD (insertManyIfNewUnit m l) k' fallback = k := Raw₀.Const.getKeyD_insertManyIfNewUnit_list_of_mem_of_contains_eq_false ⟨m.1, _⟩ m.2 k_beq - distinct mem + distinct mem theorem getKeyD_insertManyIfNewUnit_list_of_contains_of_contains [EquivBEq α] [LawfulHashable α] {l : List α} {k fallback : α} (distinct : l.Pairwise (fun a b => (a == b) = false)) - (h' : l.contains k) : m.contains k → - getKeyD (insertManyIfNewUnit m l) k fallback = getKeyD m k fallback := + (h' : l.contains k) : + m.contains k → getKeyD (insertManyIfNewUnit m l) k fallback = getKeyD m k fallback := Raw₀.Const.getKeyD_insertManyIfNewUnit_list_of_contains_of_contains ⟨m.1, _⟩ m.2 distinct h' theorem size_insertManyIfNewUnit_list [EquivBEq α] [LawfulHashable α] {l : List α} (distinct : l.Pairwise (fun a b => (a == b) = false)) : (∀ (a : α), m.contains a → l.contains a = false) → - (insertManyIfNewUnit m l).size = m.size + l.length := + (insertManyIfNewUnit m l).size = m.size + l.length := Raw₀.Const.size_insertManyIfNewUnit_list ⟨m.1, _⟩ m.2 distinct theorem size_insertManyIfNewUnit_list_le [EquivBEq α] [LawfulHashable α] @@ -1409,7 +1409,7 @@ theorem isEmpty_insertManyIfNewUnit_list [EquivBEq α] [LawfulHashable α] theorem get?_insertManyIfNewUnit_list [EquivBEq α] [LawfulHashable α] {l : List α} {k : α} : get? (insertManyIfNewUnit m l) k = - if m.contains k ∨ l.contains k then some () else none := + if m.contains k ∨ l.contains k then some () else none := Raw₀.Const.get?_insertManyIfNewUnit_list ⟨m.1, _⟩ m.2 theorem get_insertManyIfNewUnit_list @@ -1582,7 +1582,7 @@ theorem ofList_singleton {k : α} {v : β} [EquivBEq α] [LawfulHashable α] : theorem ofList_cons [EquivBEq α] [LawfulHashable α] {k : α} {v : β} {tl : List (α × β)} : insertMany empty (⟨k, v⟩ :: tl) = insertMany (empty.insert k v) tl := - Subtype.eq (congrArg Subtype.val (Raw₀.Const.insertMany_empty_list_cons (α:= α)) :) + Subtype.eq (congrArg Subtype.val (Raw₀.Const.insertMany_empty_list_cons (α:= α)) :) @[simp] theorem contains_ofList [EquivBEq α] [LawfulHashable α] @@ -1715,8 +1715,7 @@ theorem unitOfList_singleton [EquivBEq α] [LawfulHashable α] {k : α} : Subtype.eq (congrArg Subtype.val (Raw₀.Const.insertManyIfNewUnit_empty_list_singleton (α := α)) :) theorem unitOfList_cons [EquivBEq α] [LawfulHashable α] {hd : α} {tl : List α} : - unitOfList (hd :: tl) = - insertManyIfNewUnit (empty.insertIfNew hd ()) tl := + unitOfList (hd :: tl) = insertManyIfNewUnit (empty.insertIfNew hd ()) tl := Subtype.eq (congrArg Subtype.val (Raw₀.Const.insertManyIfNewUnit_empty_list_cons (α := α)) :) @[simp] diff --git a/src/Std/Data/DHashMap/RawLemmas.lean b/src/Std/Data/DHashMap/RawLemmas.lean index eaa5848476a8..684e0b551528 100644 --- a/src/Std/Data/DHashMap/RawLemmas.lean +++ b/src/Std/Data/DHashMap/RawLemmas.lean @@ -1097,8 +1097,7 @@ theorem get_insertMany_list_of_contains_eq_false [LawfulBEq α] (h : m.WF) {l : List ((a : α) × β a)} {k : α} (contains : (l.map Sigma.fst).contains k = false) {h'} : - (m.insertMany l).get k h' = - m.get k (contains_of_contains_insertMany_list h h' contains) := by + (m.insertMany l).get k h' = m.get k (contains_of_contains_insertMany_list h h' contains) := by simp_to_raw using Raw₀.get_insertMany_list_of_contains_eq_false theorem get_insertMany_list_of_mem [LawfulBEq α] (h : m.WF) @@ -1153,8 +1152,7 @@ theorem getKey_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashabl {l : List ((a : α) × β a)} {k : α} (h₁ : (l.map Sigma.fst).contains k = false) {h'} : - (m.insertMany l).getKey k h' = - m.getKey k (contains_of_contains_insertMany_list h h' h₁) := by + (m.insertMany l).getKey k h' = m.getKey k (contains_of_contains_insertMany_list h h' h₁) := by simp_to_raw using Raw₀.getKey_insertMany_list_of_contains_eq_false theorem getKey_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] (h : m.WF) @@ -1197,7 +1195,7 @@ theorem getKeyD_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] (h : m. theorem size_insertMany_list [EquivBEq α] [LawfulHashable α] (h : m.WF) {l : List ((a : α) × β a)} (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) : (∀ (a : α), m.contains a → (l.map Sigma.fst).contains a = false) → - (m.insertMany l).size = m.size + l.length := by + (m.insertMany l).size = m.size + l.length := by simp_to_raw using Raw₀.size_insertMany_list theorem size_insertMany_list_le [EquivBEq α] [LawfulHashable α] (h : m.WF) @@ -1263,8 +1261,7 @@ theorem getKey_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashabl {l : List (α × β)} {k : α} {h₁ : (l.map Prod.fst).contains k = false} {h'} : - (insertMany m l).getKey k h' = - m.getKey k (contains_of_contains_insertMany_list h h' h₁) := by + (insertMany m l).getKey k h' = m.getKey k (contains_of_contains_insertMany_list h h' h₁) := by simp_to_raw using Raw₀.Const.getKey_insertMany_list_of_contains_eq_false theorem getKey_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] (h : m.WF) @@ -1308,7 +1305,7 @@ theorem size_insertMany_list [EquivBEq α] [LawfulHashable α] (h : m.WF) {l : List (α × β)} (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) : (∀ (a : α), m.contains a → (l.map Prod.fst).contains a = false) → - (insertMany m l).size = m.size + l.length := by + (insertMany m l).size = m.size + l.length := by simp_to_raw using Raw₀.Const.size_insertMany_list theorem size_insertMany_list_le [EquivBEq α] [LawfulHashable α] (h : m.WF) @@ -1430,8 +1427,7 @@ theorem getKey_insertManyIfNewUnit_list_of_mem_of_contains_eq_false [EquivBEq α {k k' : α} (k_beq : k == k') (distinct : l.Pairwise (fun a b => (a == b) = false)) (mem : k ∈ l) {h'} : - m.contains k = false → - getKey (insertManyIfNewUnit m l) k' h' = k := by + m.contains k = false → getKey (insertManyIfNewUnit m l) k' h' = k := by simp_to_raw using Raw₀.Const.getKey_insertManyIfNewUnit_list_of_mem_of_contains_eq_false theorem getKey_insertManyIfNewUnit_list_mem_of_contains_of_contains [EquivBEq α] [LawfulHashable α] @@ -1450,15 +1446,15 @@ theorem getKey!_insertManyIfNewUnit_list_of_contains_eq_false [EquivBEq α] [Law theorem getKey!_insertManyIfNewUnit_list_of_mem_of_contains_eq_false [EquivBEq α] [LawfulHashable α] [Inhabited α] (h : m.WF) {l : List α} {k k' : α} (k_beq : k == k') (distinct : l.Pairwise (fun a b => (a == b) = false)) - (mem : k ∈ l) : contains m k = false → - getKey! (insertManyIfNewUnit m l) k' = k := by + (mem : k ∈ l) : + contains m k = false → getKey! (insertManyIfNewUnit m l) k' = k := by simp_to_raw using Raw₀.Const.getKey!_insertManyIfNewUnit_list_of_mem_of_contains_eq_false theorem getKey!_insertManyIfNewUnit_list_mem_of_contains_of_contains [EquivBEq α] [LawfulHashable α] [Inhabited α] (h : m.WF) {l : List α} {k : α} (distinct : l.Pairwise (fun a b => (a == b) = false)) - (h' : l.contains k) : m.contains k → - getKey! (insertManyIfNewUnit m l) k = getKey! m k := by + (h' : l.contains k) : + m.contains k → getKey! (insertManyIfNewUnit m l) k = getKey! m k := by simp_to_raw using Raw₀.Const.getKey!_insertManyIfNewUnit_list_mem_of_contains_of_contains theorem getKeyD_insertManyIfNewUnit_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] @@ -1470,22 +1466,22 @@ theorem getKeyD_insertManyIfNewUnit_list_of_contains_eq_false [EquivBEq α] [Law theorem getKeyD_insertManyIfNewUnit_list_of_mem_of_contains_eq_false [EquivBEq α] [LawfulHashable α] (h : m.WF) {l : List α} {k k' fallback : α} (k_beq : k == k') (distinct : l.Pairwise (fun a b => (a == b) = false)) - (mem : k ∈ l) : m.contains k = false → - getKeyD (insertManyIfNewUnit m l) k' fallback = k := by + (mem : k ∈ l) : + m.contains k = false → getKeyD (insertManyIfNewUnit m l) k' fallback = k := by simp_to_raw using Raw₀.Const.getKeyD_insertManyIfNewUnit_list_of_mem_of_contains_eq_false theorem getKeyD_insertManyIfNewUnit_list_of_contains_of_contains [EquivBEq α] [LawfulHashable α] (h : m.WF) {l : List α} {k fallback : α} (distinct : l.Pairwise (fun a b => (a == b) = false)) - (h' : l.contains k) : m.contains k → - getKeyD (insertManyIfNewUnit m l) k fallback = getKeyD m k fallback := by + (h' : l.contains k) : + m.contains k → getKeyD (insertManyIfNewUnit m l) k fallback = getKeyD m k fallback := by simp_to_raw using Raw₀.Const.getKeyD_insertManyIfNewUnit_list_of_contains_of_contains theorem size_insertManyIfNewUnit_list [EquivBEq α] [LawfulHashable α] (h : m.WF) {l : List α} (distinct : l.Pairwise (fun a b => (a == b) = false)) : (∀ (a : α), m.contains a → l.contains a = false) → - (insertManyIfNewUnit m l).size = m.size + l.length := by + (insertManyIfNewUnit m l).size = m.size + l.length := by simp_to_raw using Raw₀.Const.size_insertManyIfNewUnit_list theorem size_insertManyIfNewUnit_list_le [EquivBEq α] [LawfulHashable α] (h : m.WF) @@ -1503,7 +1499,7 @@ theorem isEmpty_insertManyIfNewUnit_list [EquivBEq α] [LawfulHashable α] (h : theorem get?_insertManyIfNewUnit_list [EquivBEq α] [LawfulHashable α] (h : m.WF) {l : List α} {k : α} : get? (insertManyIfNewUnit m l) k = - if m.contains k ∨ l.contains k then some () else none := by + if m.contains k ∨ l.contains k then some () else none := by simp_to_raw using Raw₀.Const.get?_insertManyIfNewUnit_list @[simp] @@ -1525,6 +1521,7 @@ theorem getD_insertManyIfNewUnit_list simp end Const + end Raw namespace Raw @@ -1670,6 +1667,7 @@ theorem isEmpty_ofList [EquivBEq α] [LawfulHashable α] simp_to_raw using Raw₀.isEmpty_insertMany_empty_list namespace Const + variable {β : Type v} @[simp] From 62e03c3a03fc8c24bf74c0fd3aad3dcccae53cf1 Mon Sep 17 00:00:00 2001 From: jt0202 Date: Wed, 8 Jan 2025 11:59:10 +0100 Subject: [PATCH 67/83] Style fixes --- .../Data/DHashMap/Internal/List/Associative.lean | 2 ++ src/Std/Data/DHashMap/Internal/Model.lean | 6 ++++-- src/Std/Data/DHashMap/Internal/WF.lean | 16 ++++++++++------ src/Std/Data/DHashMap/Lemmas.lean | 2 ++ 4 files changed, 18 insertions(+), 8 deletions(-) diff --git a/src/Std/Data/DHashMap/Internal/List/Associative.lean b/src/Std/Data/DHashMap/Internal/List/Associative.lean index a85566438c08..a948706d35e0 100644 --- a/src/Std/Data/DHashMap/Internal/List/Associative.lean +++ b/src/Std/Data/DHashMap/Internal/List/Associative.lean @@ -2219,6 +2219,7 @@ theorem isEmpty_insertList [BEq α] simp section + variable {β : Type v} /-- Internal implementation detail of the hash map -/ @@ -2875,4 +2876,5 @@ theorem getValueD_insertListIfNewUnit [BEq α] [PartialEquivBEq α] rw [getValueD_list_unit] end + end List diff --git a/src/Std/Data/DHashMap/Internal/Model.lean b/src/Std/Data/DHashMap/Internal/Model.lean index 4487e60a7b36..0dc01e257312 100644 --- a/src/Std/Data/DHashMap/Internal/Model.lean +++ b/src/Std/Data/DHashMap/Internal/Model.lean @@ -518,7 +518,8 @@ theorem Const.getThenInsertIfNew?_eq_get?ₘ [BEq α] [Hashable α] (m : Raw₀ dsimp only [Array.ugetElem_eq_getElem, Array.uset] split <;> simp_all [-getValue?_eq_none] -theorem Const.insertMany_eq_insertListₘ [BEq α] [Hashable α] (m : Raw₀ α (fun _ => β)) (l: List (α × β)): +theorem Const.insertMany_eq_insertListₘ [BEq α] [Hashable α] (m : Raw₀ α (fun _ => β)) + (l: List (α × β)): (Const.insertMany m l).1 = Const.insertListₘ m l := by simp only [insertMany, Id.run, Id.pure_eq, Id.bind_eq, List.forIn_yield_eq_foldl] suffices ∀ (t : { m' // ∀ (P : Raw₀ α (fun _ => β) → Prop), @@ -532,7 +533,8 @@ theorem Const.insertMany_eq_insertListₘ [BEq α] [Hashable α] (m : Raw₀ α simp only [List.foldl_cons,insertListₘ] apply ih -theorem Const.insertManyIfNewUnit_eq_insertListIfNewUnitₘ [BEq α] [Hashable α] (m : Raw₀ α (fun _ => Unit)) (l: List α): +theorem Const.insertManyIfNewUnit_eq_insertListIfNewUnitₘ [BEq α] [Hashable α] + (m : Raw₀ α (fun _ => Unit)) (l: List α): (Const.insertManyIfNewUnit m l).1 = Const.insertListIfNewUnitₘ m l := by simp only [insertManyIfNewUnit, Id.run, Id.pure_eq, Id.bind_eq, List.forIn_yield_eq_foldl] suffices ∀ (t : { m' // ∀ (P : Raw₀ α (fun _ => Unit) → Prop), diff --git a/src/Std/Data/DHashMap/Internal/WF.lean b/src/Std/Data/DHashMap/Internal/WF.lean index 853a488c675d..45f32c87ec15 100644 --- a/src/Std/Data/DHashMap/Internal/WF.lean +++ b/src/Std/Data/DHashMap/Internal/WF.lean @@ -734,7 +734,8 @@ theorem wfImp_filter [BEq α] [Hashable α] [EquivBEq α] [LawfulHashable α] {m theorem toListModel_insertListₘ [BEq α] [Hashable α] [EquivBEq α][LawfulHashable α] {m : Raw₀ α β} {l : List ((a : α) × β a)} (h : Raw.WFImp m.1) : - Perm (toListModel (insertListₘ m l).1.buckets) (List.insertList (toListModel m.1.buckets) l) := by + Perm (toListModel (insertListₘ m l).1.buckets) + (List.insertList (toListModel m.1.buckets) l) := by induction l generalizing m with | nil => simp [insertListₘ, List.insertList] @@ -774,7 +775,8 @@ theorem wfImp_insertMany [BEq α] [Hashable α] [EquivBEq α] [LawfulHashable α theorem toListModel_insertMany_list [BEq α] [Hashable α] [EquivBEq α] [LawfulHashable α] {m : Raw₀ α β} {l : List ((a : α) × (β a))} (h : Raw.WFImp m.1) : - Perm (toListModel (insertMany m l).1.1.buckets) (List.insertList (toListModel m.1.buckets) l) := by + Perm (toListModel (insertMany m l).1.1.buckets) + (List.insertList (toListModel m.1.buckets) l) := by rw [insertMany_eq_insertListₘ] apply toListModel_insertListₘ exact h @@ -783,7 +785,8 @@ theorem toListModel_insertMany_list [BEq α] [Hashable α] [EquivBEq α] [Lawful theorem Const.toListModel_insertListₘ {β : Type v} [BEq α] [Hashable α] [EquivBEq α] [LawfulHashable α] {m : Raw₀ α (fun _ => β)} {l : List (α × β)} (h : Raw.WFImp m.1) : - Perm (toListModel (Const.insertListₘ m l).1.buckets) (insertListConst (toListModel m.1.buckets) l) := by + Perm (toListModel (Const.insertListₘ m l).1.buckets) + (insertListConst (toListModel m.1.buckets) l) := by induction l generalizing m with | nil => simp [Const.insertListₘ, insertListConst, insertList] | cons hd tl ih => @@ -796,7 +799,8 @@ theorem Const.toListModel_insertListₘ {β : Type v} [BEq α] [Hashable α] [Eq theorem Const.toListModel_insertMany_list {β : Type v} [BEq α] [Hashable α] [EquivBEq α] [LawfulHashable α] {m : Raw₀ α (fun _ => β)} {l : List (α × β)} (h : Raw.WFImp m.1) : - Perm (toListModel (Const.insertMany m l).1.1.buckets) (insertListConst (toListModel m.1.buckets) l) := by + Perm (toListModel (Const.insertMany m l).1.1.buckets) + (insertListConst (toListModel m.1.buckets) l) := by rw [Const.insertMany_eq_insertListₘ] apply toListModel_insertListₘ h @@ -810,7 +814,7 @@ theorem Const.wfImp_insertMany {β : Type v} [BEq α] [Hashable α] [EquivBEq α theorem Const.toListModel_insertListIfNewUnitₘ [BEq α] [Hashable α] [EquivBEq α] [LawfulHashable α] {m : Raw₀ α (fun _ => Unit)} {l : List α} (h : Raw.WFImp m.1) : Perm (toListModel (Const.insertListIfNewUnitₘ m l).1.buckets) - (List.insertListIfNewUnit (toListModel m.1.buckets) l) := by + (List.insertListIfNewUnit (toListModel m.1.buckets) l) := by induction l generalizing m with | nil => simp [insertListIfNewUnitₘ, List.insertListIfNewUnit] | cons hd tl ih => @@ -824,7 +828,7 @@ theorem Const.toListModel_insertListIfNewUnitₘ [BEq α] [Hashable α] [EquivBE theorem Const.toListModel_insertManyIfNewUnit_list [BEq α] [Hashable α] [EquivBEq α] [LawfulHashable α] {m : Raw₀ α (fun _ => Unit)} {l : List α} (h : Raw.WFImp m.1) : Perm (toListModel (Const.insertManyIfNewUnit m l).1.1.buckets) - (List.insertListIfNewUnit (toListModel m.1.buckets) l) := by + (List.insertListIfNewUnit (toListModel m.1.buckets) l) := by rw [Const.insertManyIfNewUnit_eq_insertListIfNewUnitₘ] apply toListModel_insertListIfNewUnitₘ h diff --git a/src/Std/Data/DHashMap/Lemmas.lean b/src/Std/Data/DHashMap/Lemmas.lean index 4a113c4cba2d..ed14d7e05d70 100644 --- a/src/Std/Data/DHashMap/Lemmas.lean +++ b/src/Std/Data/DHashMap/Lemmas.lean @@ -1428,6 +1428,7 @@ theorem getD_insertManyIfNewUnit_list simp end Const + end DHashMap namespace DHashMap @@ -1812,4 +1813,5 @@ theorem getD_unitOfList simp end Const + end Std.DHashMap From fffcc8fb8d75dd00994e0ffef13e25d51f43d8e8 Mon Sep 17 00:00:00 2001 From: jt0202 Date: Wed, 8 Jan 2025 18:50:12 +0100 Subject: [PATCH 68/83] Style fixes --- src/Std/Data/DHashMap/Lemmas.lean | 276 +++++++++++++++------------ src/Std/Data/DHashMap/RawLemmas.lean | 130 ++++++++----- 2 files changed, 235 insertions(+), 171 deletions(-) diff --git a/src/Std/Data/DHashMap/Lemmas.lean b/src/Std/Data/DHashMap/Lemmas.lean index ed14d7e05d70..9a244f5bf7f8 100644 --- a/src/Std/Data/DHashMap/Lemmas.lean +++ b/src/Std/Data/DHashMap/Lemmas.lean @@ -983,16 +983,23 @@ theorem contains_insertMany_list [EquivBEq α] [LawfulHashable α] (m.insertMany l).contains k = (m.contains k || (l.map Sigma.fst).contains k) := Raw₀.contains_insertMany_list ⟨m.1, _⟩ m.2 -theorem contains_of_contains_insertMany_list [EquivBEq α] [LawfulHashable α] +@[simp] +theorem mem_insertMany_list [EquivBEq α] [LawfulHashable α] {l : List ((a : α) × β a)} {k : α} : - (m.insertMany l).contains k → (l.map Sigma.fst).contains k = false → m.contains k := - Raw₀.contains_of_contains_insertMany_list ⟨m.1, _⟩ m.2 + k ∈ m.insertMany l ↔ k ∈ m ∨ (l.map Sigma.fst).contains k := by + simp [mem_iff_contains] + +theorem contains_of_contains_insertMany_list [EquivBEq α] [LawfulHashable α] + {l : List ((a : α) × β a)} {k : α} (contains : (m.insertMany l).contains k) + (contains_eq_false: (l.map Sigma.fst).contains k = false) : + m.contains k := + Raw₀.contains_of_contains_insertMany_list ⟨m.1, _⟩ m.2 contains contains_eq_false theorem get?_insertMany_list_of_contains_eq_false [LawfulBEq α] {l : List ((a : α) × β a)} {k : α} - (h' : (l.map Sigma.fst).contains k = false) : + (contains_eq_false : (l.map Sigma.fst).contains k = false) : (m.insertMany l).get? k = m.get? k := - Raw₀.get?_insertMany_list_of_contains_eq_false ⟨m.1, _⟩ m.2 h' + Raw₀.get?_insertMany_list_of_contains_eq_false ⟨m.1, _⟩ m.2 contains_eq_false theorem get?_insertMany_list_of_mem [LawfulBEq α] {l : List ((a : α) × β a)} {k k' : α} (k_beq : k == k') {v : β k} @@ -1003,25 +1010,25 @@ theorem get?_insertMany_list_of_mem [LawfulBEq α] theorem get_insertMany_list_of_contains_eq_false [LawfulBEq α] {l : List ((a : α) × β a)} {k : α} - (contains : (l.map Sigma.fst).contains k = false) - {h'} : - (m.insertMany l).get k h' = - m.get k (contains_of_contains_insertMany_list h' contains) := - Raw₀.get_insertMany_list_of_contains_eq_false ⟨m.1, _⟩ m.2 contains + (contains_eq_false : (l.map Sigma.fst).contains k = false) + {h} : + (m.insertMany l).get k h = + m.get k (contains_of_contains_insertMany_list h contains_eq_false) := + Raw₀.get_insertMany_list_of_contains_eq_false ⟨m.1, _⟩ m.2 contains_eq_false theorem get_insertMany_list_of_mem [LawfulBEq α] {l : List ((a : α) × β a)} {k k' : α} (k_beq : k == k') {v : β k} (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) (mem : ⟨k, v⟩ ∈ l) - {h'} : - (m.insertMany l).get k' h' = cast (by congr; apply LawfulBEq.eq_of_beq k_beq) v := + {h} : + (m.insertMany l).get k' h = cast (by congr; apply LawfulBEq.eq_of_beq k_beq) v := Raw₀.get_insertMany_list_of_mem ⟨m.1, _⟩ m.2 k_beq distinct mem theorem get!_insertMany_list_of_contains_eq_false [LawfulBEq α] {l : List ((a : α) × β a)} {k : α} [Inhabited (β k)] - (h' : (l.map Sigma.fst).contains k = false) : + (contains_eq_false : (l.map Sigma.fst).contains k = false) : (m.insertMany l).get! k = m.get! k := - Raw₀.get!_insertMany_list_of_contains_eq_false ⟨m.1, _⟩ m.2 h' + Raw₀.get!_insertMany_list_of_contains_eq_false ⟨m.1, _⟩ m.2 contains_eq_false theorem get!_insertMany_list_of_mem [LawfulBEq α] {l : List ((a : α) × β a)} {k k' : α} (k_beq : k == k') {v : β k} [Inhabited (β k')] @@ -1032,9 +1039,9 @@ theorem get!_insertMany_list_of_mem [LawfulBEq α] theorem getD_insertMany_list_of_contains_eq_false [LawfulBEq α] {l : List ((a : α) × β a)} {k : α} {fallback : β k} - (not_mem : (l.map Sigma.fst).contains k = false) : + (contains_eq_false : (l.map Sigma.fst).contains k = false) : (m.insertMany l).getD k fallback = m.getD k fallback := - Raw₀.getD_insertMany_list_of_contains_eq_false ⟨m.1, _⟩ m.2 not_mem + Raw₀.getD_insertMany_list_of_contains_eq_false ⟨m.1, _⟩ m.2 contains_eq_false theorem getD_insertMany_list_of_mem [LawfulBEq α] {l : List ((a : α) × β a)} {k k' : α} (k_beq : k == k') {v : β k} {fallback : β k'} @@ -1045,9 +1052,9 @@ theorem getD_insertMany_list_of_mem [LawfulBEq α] theorem getKey?_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] {l : List ((a : α) × β a)} {k : α} - (h' : (l.map Sigma.fst).contains k = false) : + (contains_eq_false : (l.map Sigma.fst).contains k = false) : (m.insertMany l).getKey? k = m.getKey? k := - Raw₀.getKey?_insertMany_list_of_contains_eq_false ⟨m.1, _⟩ m.2 h' + Raw₀.getKey?_insertMany_list_of_contains_eq_false ⟨m.1, _⟩ m.2 contains_eq_false theorem getKey?_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] {l : List ((a : α) × β a)} @@ -1059,26 +1066,26 @@ theorem getKey?_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] theorem getKey_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] {l : List ((a : α) × β a)} {k : α} - (h₁ : (l.map Sigma.fst).contains k = false) - {h'} : - (m.insertMany l).getKey k h' = - m.getKey k (contains_of_contains_insertMany_list h' h₁) := - Raw₀.getKey_insertMany_list_of_contains_eq_false ⟨m.1, _⟩ m.2 h₁ + (contains_eq_false : (l.map Sigma.fst).contains k = false) + {h} : + (m.insertMany l).getKey k h = + m.getKey k (contains_of_contains_insertMany_list h contains_eq_false) := + Raw₀.getKey_insertMany_list_of_contains_eq_false ⟨m.1, _⟩ m.2 contains_eq_false theorem getKey_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] {l : List ((a : α) × β a)} {k k' : α} (k_beq : k == k') (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) (mem : k ∈ l.map Sigma.fst) - {h'} : - (m.insertMany l).getKey k' h' = k := + {h} : + (m.insertMany l).getKey k' h = k := Raw₀.getKey_insertMany_list_of_mem ⟨m.1, _⟩ m.2 k_beq distinct mem theorem getKey!_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] [Inhabited α] {l : List ((a : α) × β a)} {k : α} - (h' : (l.map Sigma.fst).contains k = false) : + (contains_eq_false : (l.map Sigma.fst).contains k = false) : (m.insertMany l).getKey! k = m.getKey! k := - Raw₀.getKey!_insertMany_list_of_contains_eq_false ⟨m.1, _⟩ m.2 h' + Raw₀.getKey!_insertMany_list_of_contains_eq_false ⟨m.1, _⟩ m.2 contains_eq_false theorem getKey!_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] [Inhabited α] {l : List ((a : α) × β a)} @@ -1090,9 +1097,9 @@ theorem getKey!_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] [Inhabi theorem getKeyD_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] {l : List ((a : α) × β a)} {k fallback : α} - (h' : (l.map Sigma.fst).contains k = false) : + (contains_eq_false : (l.map Sigma.fst).contains k = false) : (m.insertMany l).getKeyD k fallback = m.getKeyD k fallback := - Raw₀.getKeyD_insertMany_list_of_contains_eq_false ⟨m.1, _⟩ m.2 h' + Raw₀.getKeyD_insertMany_list_of_contains_eq_false ⟨m.1, _⟩ m.2 contains_eq_false theorem getKeyD_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] {l : List ((a : α) × β a)} @@ -1144,16 +1151,23 @@ theorem contains_insertMany_list [EquivBEq α] [LawfulHashable α] (Const.insertMany m l).contains k = (m.contains k || (l.map Prod.fst).contains k) := Raw₀.Const.contains_insertMany_list ⟨m.1, _⟩ m.2 +@[simp] +theorem mem_insertMany_list [EquivBEq α] [LawfulHashable α] + {l : List (α × β)} {k : α} : + k ∈ insertMany m l ↔ k ∈ m ∨ (l.map Prod.fst).contains k := by + simp [mem_iff_contains] + theorem contains_of_contains_insertMany_list [EquivBEq α] [LawfulHashable α] - {l : List ( α × β )} {k : α} : - (insertMany m l).contains k → (l.map Prod.fst).contains k = false → m.contains k := - Raw₀.Const.contains_of_contains_insertMany_list ⟨m.1, _⟩ m.2 + {l : List ( α × β )} {k : α} (contains : (insertMany m l).contains k) + (contains_eq_false : (l.map Prod.fst).contains k = false) : + m.contains k := + Raw₀.Const.contains_of_contains_insertMany_list ⟨m.1, _⟩ m.2 contains contains_eq_false theorem getKey?_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] {l : List (α × β)} {k : α} - (h' : (l.map Prod.fst).contains k = false) : + (contains_eq_false : (l.map Prod.fst).contains k = false) : (insertMany m l).getKey? k = m.getKey? k := - Raw₀.Const.getKey?_insertMany_list_of_contains_eq_false ⟨m.1, _⟩ m.2 h' + Raw₀.Const.getKey?_insertMany_list_of_contains_eq_false ⟨m.1, _⟩ m.2 contains_eq_false theorem getKey?_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] {l : List (α × β)} @@ -1165,26 +1179,26 @@ theorem getKey?_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] theorem getKey_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] {l : List (α × β)} {k : α} - (h₁ : (l.map Prod.fst).contains k = false) - {h'} : - (insertMany m l).getKey k h' = - m.getKey k (contains_of_contains_insertMany_list _ h' h₁) := - Raw₀.Const.getKey_insertMany_list_of_contains_eq_false ⟨m.1, _⟩ m.2 h₁ + (contains_eq_false : (l.map Prod.fst).contains k = false) + {h} : + (insertMany m l).getKey k h = + m.getKey k (contains_of_contains_insertMany_list _ h contains_eq_false) := + Raw₀.Const.getKey_insertMany_list_of_contains_eq_false ⟨m.1, _⟩ m.2 contains_eq_false theorem getKey_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] {l : List (α × β)} {k k' : α} (k_beq : k == k') (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) (mem : k ∈ l.map Prod.fst) - {h'} : - (insertMany m l).getKey k' h' = k := + {h} : + (insertMany m l).getKey k' h = k := Raw₀.Const.getKey_insertMany_list_of_mem ⟨m.1, _⟩ m.2 k_beq distinct mem theorem getKey!_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] [Inhabited α] {l : List (α × β)} {k : α} - (h' : (l.map Prod.fst).contains k = false) : + (contains_eq_false : (l.map Prod.fst).contains k = false) : (insertMany m l).getKey! k = m.getKey! k := - Raw₀.Const.getKey!_insertMany_list_of_contains_eq_false ⟨m.1, _⟩ m.2 h' + Raw₀.Const.getKey!_insertMany_list_of_contains_eq_false ⟨m.1, _⟩ m.2 contains_eq_false theorem getKey!_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] [Inhabited α] {l : List (α × β)} @@ -1196,9 +1210,9 @@ theorem getKey!_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] [Inhabi theorem getKeyD_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] {l : List (α × β)} {k fallback : α} - (h' : (l.map Prod.fst).contains k = false) : + (contains_eq_false : (l.map Prod.fst).contains k = false) : (insertMany m l).getKeyD k fallback = m.getKeyD k fallback := - Raw₀.Const.getKeyD_insertMany_list_of_contains_eq_false ⟨m.1, _⟩ m.2 h' + Raw₀.Const.getKeyD_insertMany_list_of_contains_eq_false ⟨m.1, _⟩ m.2 contains_eq_false theorem getKeyD_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] {l : List (α × β)} @@ -1228,9 +1242,9 @@ theorem isEmpty_insertMany_list [EquivBEq α] [LawfulHashable α] theorem get?_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] {l : List (α × β)} {k : α} - (h' : (l.map Prod.fst).contains k = false) : + (contains_eq_false : (l.map Prod.fst).contains k = false) : get? (insertMany m l) k = get? m k := - Raw₀.Const.get?_insertMany_list_of_contains_eq_false ⟨m.1, _⟩ m.2 h' + Raw₀.Const.get?_insertMany_list_of_contains_eq_false ⟨m.1, _⟩ m.2 contains_eq_false theorem get?_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] {l : List (α × β)} {k k' : α} (k_beq : k == k') {v : β} @@ -1240,22 +1254,22 @@ theorem get?_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] theorem get_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] {l : List (α × β)} {k : α} - (h₁ : (l.map Prod.fst).contains k = false) - {h'} : - get (insertMany m l) k h' = get m k (contains_of_contains_insertMany_list _ h' h₁) := - Raw₀.Const.get_insertMany_list_of_contains_eq_false ⟨m.1, _⟩ m.2 h₁ + (contains_eq_false : (l.map Prod.fst).contains k = false) + {h} : + get (insertMany m l) k h = get m k (contains_of_contains_insertMany_list _ h contains_eq_false) := + Raw₀.Const.get_insertMany_list_of_contains_eq_false ⟨m.1, _⟩ m.2 contains_eq_false theorem get_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] {l : List (α × β)} {k k' : α} (k_beq : k == k') {v : β} - (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) (mem : ⟨k, v⟩ ∈ l) {h'} : - get (insertMany m l) k' h' = v := + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) (mem : ⟨k, v⟩ ∈ l) {h} : + get (insertMany m l) k' h = v := Raw₀.Const.get_insertMany_list_of_mem ⟨m.1, _⟩ m.2 k_beq distinct mem theorem get!_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] [Inhabited β] {l : List (α × β)} {k : α} - (h' : (l.map Prod.fst).contains k = false) : + (contains_eq_false : (l.map Prod.fst).contains k = false) : get! (insertMany m l) k = get! m k := - Raw₀.Const.get!_insertMany_list_of_contains_eq_false ⟨m.1, _⟩ m.2 h' + Raw₀.Const.get!_insertMany_list_of_contains_eq_false ⟨m.1, _⟩ m.2 contains_eq_false theorem get!_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] [Inhabited β] {l : List (α × β)} {k k' : α} (k_beq : k == k') {v : β} @@ -1265,9 +1279,9 @@ theorem get!_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] [Inhabited theorem getD_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] {l : List (α × β)} {k : α} {fallback : β} - (h' : (l.map Prod.fst).contains k = false) : + (contains_eq_false : (l.map Prod.fst).contains k = false) : getD (insertMany m l) k fallback = getD m k fallback := - Raw₀.Const.getD_insertMany_list_of_contains_eq_false ⟨m.1, _⟩ m.2 h' + Raw₀.Const.getD_insertMany_list_of_contains_eq_false ⟨m.1, _⟩ m.2 contains_eq_false theorem getD_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] {l : List (α × β)} {k k' : α} (k_beq : k == k') {v fallback : β} @@ -1300,57 +1314,66 @@ theorem contains_insertManyIfNewUnit_list [EquivBEq α] [LawfulHashable α] (insertManyIfNewUnit m l).contains k = (m.contains k || l.contains k) := Raw₀.Const.contains_insertManyIfNewUnit_list ⟨m.1, _⟩ m.2 +@[simp] +theorem mem_insertManyIfNewUnit_list [EquivBEq α] [LawfulHashable α] + {l : List α} {k : α} : + k ∈ insertManyIfNewUnit m l ↔ k ∈ m ∨ l.contains k := by + simp [mem_iff_contains] + theorem contains_of_contains_insertManyIfNewUnit_list [EquivBEq α] [LawfulHashable α] - {l : List α} {k : α} (h₂ : l.contains k = false) : + {l : List α} {k : α} (contains_eq_false : l.contains k = false) : (insertManyIfNewUnit m l).contains k → m.contains k := - Raw₀.Const.contains_of_contains_insertManyIfNewUnit_list ⟨m.1, _⟩ m.2 h₂ + Raw₀.Const.contains_of_contains_insertManyIfNewUnit_list ⟨m.1, _⟩ m.2 contains_eq_false theorem getKey?_insertManyIfNewUnit_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] - {l : List α} {k : α} (h' : l.contains k = false) : + {l : List α} {k : α} (contains_eq_false : l.contains k = false) : getKey? (insertManyIfNewUnit m l) k = getKey? m k := - Raw₀.Const.getKey?_insertManyIfNewUnit_list_of_contains_eq_false ⟨m.1, _⟩ m.2 h' + Raw₀.Const.getKey?_insertManyIfNewUnit_list_of_contains_eq_false ⟨m.1, _⟩ m.2 contains_eq_false theorem getKey?_insertManyIfNewUnit_list_of_mem_of_contains_eq_false [EquivBEq α] [LawfulHashable α] {l : List α} {k k' : α} (k_beq : k == k') - (distinct : l.Pairwise (fun a b => (a == b) = false)) (mem : k ∈ l) : - m.contains k = false → getKey? (insertManyIfNewUnit m l) k' = some k := + (distinct : l.Pairwise (fun a b => (a == b) = false)) (mem : k ∈ l) + (contains_eq_false : m.contains k = false) : + getKey? (insertManyIfNewUnit m l) k' = some k := Raw₀.Const.getKey?_insertManyIfNewUnit_list_of_mem_of_contains_eq_false ⟨m.1, _⟩ - m.2 k_beq distinct mem + m.2 k_beq distinct mem contains_eq_false theorem getKey?_insertManyIfNewUnit_list_of_contains_of_contains [EquivBEq α] [LawfulHashable α] {l : List α} {k : α} (distinct : l.Pairwise (fun a b => (a == b) = false)) - (h' : l.contains k) : - m.contains k → getKey? (insertManyIfNewUnit m l) k = getKey? m k := - Raw₀.Const.getKey?_insertManyIfNewUnit_list_of_contains_of_contains ⟨m.1, _⟩ m.2 distinct h' + (h : l.contains k) + (h' : m.contains k) : + getKey? (insertManyIfNewUnit m l) k = getKey? m k := + Raw₀.Const.getKey?_insertManyIfNewUnit_list_of_contains_of_contains ⟨m.1, _⟩ m.2 distinct h h' theorem getKey_insertManyIfNewUnit_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] - {l : List α} {k : α} (h₁ : l.contains k = false) {h'} : - getKey (insertManyIfNewUnit m l) k h' = - getKey m k (contains_of_contains_insertManyIfNewUnit_list m h₁ h') := - Raw₀.Const.getKey_insertManyIfNewUnit_list_of_contains_eq_false ⟨m.1, _⟩ m.2 h₁ + {l : List α} {k : α} (contains_eq_false : l.contains k = false) {h} : + getKey (insertManyIfNewUnit m l) k h = + getKey m k (contains_of_contains_insertManyIfNewUnit_list m contains_eq_false h) := + Raw₀.Const.getKey_insertManyIfNewUnit_list_of_contains_eq_false ⟨m.1, _⟩ m.2 contains_eq_false theorem getKey_insertManyIfNewUnit_list_of_mem_of_contains_eq_false [EquivBEq α] [LawfulHashable α] {l : List α} {k k' : α} (k_beq : k == k') (distinct : l.Pairwise (fun a b => (a == b) = false)) - (mem : k ∈ l) {h'} : - m.contains k = false → getKey (insertManyIfNewUnit m l) k' h' = k := + (mem : k ∈ l) {h} + (contains_eq_false : m.contains k = false) : + getKey (insertManyIfNewUnit m l) k' h = k := Raw₀.Const.getKey_insertManyIfNewUnit_list_of_mem_of_contains_eq_false ⟨m.1, _⟩ m.2 k_beq - distinct mem + distinct mem contains_eq_false theorem getKey_insertManyIfNewUnit_list_mem_of_contains_of_contains [EquivBEq α] [LawfulHashable α] {l : List α} {k : α} (distinct : l.Pairwise (fun a b => (a == b) = false)) - (h₁ : l.contains k) {mem' : m.contains k} {h'} : - getKey (insertManyIfNewUnit m l) k h' = getKey m k mem' := + (h₁ : l.contains k) {h₂ : m.contains k} {h} : + getKey (insertManyIfNewUnit m l) k h = getKey m k h₂ := Raw₀.Const.getKey_insertManyIfNewUnit_list_mem_of_contains_of_contains ⟨m.1, _⟩ m.2 distinct h₁ theorem getKey!_insertManyIfNewUnit_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] [Inhabited α] {l : List α} {k : α} - (h' : l.contains k = false) : + (contains_eq_false : l.contains k = false) : getKey! (insertManyIfNewUnit m l) k = getKey! m k := - Raw₀.Const.getKey!_insertManyIfNewUnit_list_of_contains_eq_false ⟨m.1, _⟩ m.2 h' + Raw₀.Const.getKey!_insertManyIfNewUnit_list_of_contains_eq_false ⟨m.1, _⟩ m.2 contains_eq_false theorem getKey!_insertManyIfNewUnit_list_of_mem_of_contains_eq_false [EquivBEq α] [LawfulHashable α] [Inhabited α] {l : List α} {k k' : α} (k_beq : k == k') @@ -1363,30 +1386,30 @@ theorem getKey!_insertManyIfNewUnit_list_of_mem_of_contains_eq_false [EquivBEq theorem getKey!_insertManyIfNewUnit_list_mem_of_contains_of_contains [EquivBEq α] [LawfulHashable α] [Inhabited α] {l : List α} {k : α} (distinct : l.Pairwise (fun a b => (a == b) = false)) - (h' : l.contains k) : + (h : l.contains k) : m.contains k → getKey! (insertManyIfNewUnit m l) k = getKey! m k := - Raw₀.Const.getKey!_insertManyIfNewUnit_list_mem_of_contains_of_contains ⟨m.1, _⟩ m.2 distinct h' + Raw₀.Const.getKey!_insertManyIfNewUnit_list_mem_of_contains_of_contains ⟨m.1, _⟩ m.2 distinct h theorem getKeyD_insertManyIfNewUnit_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] {l : List α} {k fallback : α} - (h' : l.contains k = false) : + (contains_eq_false : l.contains k = false) : getKeyD (insertManyIfNewUnit m l) k fallback = getKeyD m k fallback := - Raw₀.Const.getKeyD_insertManyIfNewUnit_list_of_contains_eq_false ⟨m.1, _⟩ m.2 h' + Raw₀.Const.getKeyD_insertManyIfNewUnit_list_of_contains_eq_false ⟨m.1, _⟩ m.2 contains_eq_false theorem getKeyD_insertManyIfNewUnit_list_of_mem_of_contains_eq_false [EquivBEq α] [LawfulHashable α] {l : List α} {k k' fallback : α} (k_beq : k == k') (distinct : l.Pairwise (fun a b => (a == b) = false)) - (mem : k ∈ l ) : - m.contains k = false → getKeyD (insertManyIfNewUnit m l) k' fallback = k := + (mem : k ∈ l ) (contains_eq_false : m.contains k = false) : + getKeyD (insertManyIfNewUnit m l) k' fallback = k := Raw₀.Const.getKeyD_insertManyIfNewUnit_list_of_mem_of_contains_eq_false ⟨m.1, _⟩ m.2 k_beq - distinct mem + distinct mem contains_eq_false theorem getKeyD_insertManyIfNewUnit_list_of_contains_of_contains [EquivBEq α] [LawfulHashable α] {l : List α} {k fallback : α} (distinct : l.Pairwise (fun a b => (a == b) = false)) - (h' : l.contains k) : + (h : l.contains k) : m.contains k → getKeyD (insertManyIfNewUnit m l) k fallback = getKeyD m k fallback := - Raw₀.Const.getKeyD_insertManyIfNewUnit_list_of_contains_of_contains ⟨m.1, _⟩ m.2 distinct h' + Raw₀.Const.getKeyD_insertManyIfNewUnit_list_of_contains_of_contains ⟨m.1, _⟩ m.2 distinct h theorem size_insertManyIfNewUnit_list [EquivBEq α] [LawfulHashable α] {l : List α} @@ -1440,11 +1463,11 @@ theorem ofList_nil [EquivBEq α] [LawfulHashable α] : @[simp] theorem ofList_singleton {k : α} {v : β k} [EquivBEq α] [LawfulHashable α] : - ofList [⟨k, v⟩] = empty.insert k v := + ofList [⟨k, v⟩] = (∅: DHashMap α β).insert k v := Subtype.eq (congrArg Subtype.val (Raw₀.insertMany_empty_list_singleton (α := α)) :) theorem ofList_cons [EquivBEq α] [LawfulHashable α] {k : α} {v : β k} {tl : List ((a : α) × (β a))} : - ofList (⟨k, v⟩ :: tl) = (empty.insert k v).insertMany tl := + ofList (⟨k, v⟩ :: tl) = ((∅ : DHashMap α β).insert k v).insertMany tl := Subtype.eq (congrArg Subtype.val (Raw₀.insertMany_empty_list_cons (α := α)) :) @[simp] @@ -1455,9 +1478,9 @@ theorem contains_ofList [EquivBEq α] [LawfulHashable α] theorem get?_ofList_of_contains_eq_false [LawfulBEq α] {l : List ((a : α) × β a)} {k : α} - (h : (l.map Sigma.fst).contains k = false) : + (contains_eq_false : (l.map Sigma.fst).contains k = false) : (ofList l).get? k = none := - Raw₀.get?_insertMany_empty_list_of_contains_eq_false h + Raw₀.get?_insertMany_empty_list_of_contains_eq_false contains_eq_false theorem get?_ofList_of_mem [LawfulBEq α] {l : List ((a : α) × β a)} {k k' : α} (k_beq : k == k') {v : β k} @@ -1476,9 +1499,9 @@ theorem get_ofList_of_mem [LawfulBEq α] theorem get!_ofList_of_contains_eq_false [LawfulBEq α] {l : List ((a : α) × β a)} {k : α} [Inhabited (β k)] - (h : (l.map Sigma.fst).contains k = false) : + (contains_eq_false : (l.map Sigma.fst).contains k = false) : (ofList l).get! k = default := - Raw₀.get!_insertMany_empty_list_of_contains_eq_false h + Raw₀.get!_insertMany_empty_list_of_contains_eq_false contains_eq_false theorem get!_ofList_of_mem [LawfulBEq α] {l : List ((a : α) × β a)} {k k' : α} (k_beq : k == k') {v : β k} [Inhabited (β k')] @@ -1502,9 +1525,9 @@ theorem getD_ofList_of_mem [LawfulBEq α] theorem getKey?_ofList_of_contains_eq_false [EquivBEq α] [LawfulHashable α] {l : List ((a : α) × β a)} {k : α} - (h : (l.map Sigma.fst).contains k = false) : + (contains_eq_false : (l.map Sigma.fst).contains k = false) : (ofList l).getKey? k = none := - Raw₀.getKey?_insertMany_empty_list_of_contains_eq_false h + Raw₀.getKey?_insertMany_empty_list_of_contains_eq_false contains_eq_false theorem getKey?_ofList_of_mem [EquivBEq α] [LawfulHashable α] {l : List ((a : α) × β a)} @@ -1519,15 +1542,15 @@ theorem getKey_ofList_of_mem [EquivBEq α] [LawfulHashable α] {k k' : α} (k_beq : k == k') (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) (mem : k ∈ l.map Sigma.fst) - {h'} : - (ofList l).getKey k' h' = k := + {h} : + (ofList l).getKey k' h = k := Raw₀.getKey_insertMany_empty_list_of_mem k_beq distinct mem theorem getKey!_ofList_of_contains_eq_false [EquivBEq α] [LawfulHashable α] [Inhabited α] {l : List ((a : α) × β a)} {k : α} - (h : (l.map Sigma.fst).contains k = false) : + (contains_eq_false : (l.map Sigma.fst).contains k = false) : (ofList l).getKey! k = default := - Raw₀.getKey!_insertMany_empty_list_of_contains_eq_false h + Raw₀.getKey!_insertMany_empty_list_of_contains_eq_false contains_eq_false theorem getKey!_ofList_of_mem [EquivBEq α] [LawfulHashable α] [Inhabited α] {l : List ((a : α) × β a)} @@ -1539,9 +1562,9 @@ theorem getKey!_ofList_of_mem [EquivBEq α] [LawfulHashable α] [Inhabited α] theorem getKeyD_ofList_of_contains_eq_false [EquivBEq α] [LawfulHashable α] {l : List ((a : α) × β a)} {k fallback : α} - (h : (l.map Sigma.fst).contains k = false) : + (contains_eq_false : (l.map Sigma.fst).contains k = false) : (ofList l).getKeyD k fallback = fallback := - Raw₀.getKeyD_insertMany_empty_list_of_contains_eq_false h + Raw₀.getKeyD_insertMany_empty_list_of_contains_eq_false contains_eq_false theorem getKeyD_ofList_of_mem [EquivBEq α] [LawfulHashable α] {l : List ((a : α) × β a)} @@ -1562,7 +1585,7 @@ theorem size_ofList_le [EquivBEq α] [LawfulHashable α] Raw₀.size_insertMany_empty_list_le @[simp] -theorem isEmpty_ofList[EquivBEq α] [LawfulHashable α] +theorem isEmpty_ofList [EquivBEq α] [LawfulHashable α] {l : List ((a : α) × β a)} : (ofList l).isEmpty = l.isEmpty := Raw₀.isEmpty_insertMany_empty_list @@ -1578,11 +1601,11 @@ theorem ofList_nil [EquivBEq α] [LawfulHashable α] : @[simp] theorem ofList_singleton {k : α} {v : β} [EquivBEq α] [LawfulHashable α] : - ofList [⟨k, v⟩] = empty.insert k v := + ofList [⟨k, v⟩] = (∅ : DHashMap α (fun _ => β)).insert k v := Subtype.eq (congrArg Subtype.val (Raw₀.Const.insertMany_empty_list_singleton (α:= α)) :) theorem ofList_cons [EquivBEq α] [LawfulHashable α] {k : α} {v : β} {tl : List (α × β)} : - insertMany empty (⟨k, v⟩ :: tl) = insertMany (empty.insert k v) tl := + ofList (⟨k, v⟩ :: tl) = insertMany ((∅ : DHashMap α (fun _ => β)).insert k v) tl := Subtype.eq (congrArg Subtype.val (Raw₀.Const.insertMany_empty_list_cons (α:= α)) :) @[simp] @@ -1593,9 +1616,9 @@ theorem contains_ofList [EquivBEq α] [LawfulHashable α] theorem get?_ofList_of_contains_eq_false [LawfulBEq α] {l : List (α × β)} {k : α} - (h : (l.map Prod.fst).contains k = false) : + (contains_eq_false : (l.map Prod.fst).contains k = false) : get? (ofList l) k = none := - Raw₀.Const.get?_insertMany_empty_list_of_contains_eq_false h + Raw₀.Const.get?_insertMany_empty_list_of_contains_eq_false contains_eq_false theorem get?_ofList_of_mem [LawfulBEq α] {l : List (α × β)} {k k' : α} (k_beq : k == k') {v : β} @@ -1614,9 +1637,9 @@ theorem get_ofList_of_mem [LawfulBEq α] theorem get!_ofList_of_contains_eq_false [LawfulBEq α] {l : List (α × β)} {k : α} [Inhabited β] - (h : (l.map Prod.fst).contains k = false) : + (contains_eq_false : (l.map Prod.fst).contains k = false) : get! (ofList l) k = (default : β) := - Raw₀.Const.get!_insertMany_empty_list_of_contains_eq_false h + Raw₀.Const.get!_insertMany_empty_list_of_contains_eq_false contains_eq_false theorem get!_ofList_of_mem [LawfulBEq α] {l : List (α × β)} {k k' : α} (k_beq : k == k') {v : β} [Inhabited β] @@ -1640,9 +1663,9 @@ theorem getD_ofList_of_mem [LawfulBEq α] theorem getKey?_ofList_of_contains_eq_false [EquivBEq α] [LawfulHashable α] {l : List (α × β)} {k : α} - (h : (l.map Prod.fst).contains k = false) : + (contains_eq_false : (l.map Prod.fst).contains k = false) : (ofList l).getKey? k = none := - Raw₀.Const.getKey?_insertMany_empty_list_of_contains_eq_false h + Raw₀.Const.getKey?_insertMany_empty_list_of_contains_eq_false contains_eq_false theorem getKey?_ofList_of_mem [EquivBEq α] [LawfulHashable α] {l : List (α × β)} @@ -1657,15 +1680,15 @@ theorem getKey_ofList_of_mem [EquivBEq α] [LawfulHashable α] {k k' : α} (k_beq : k == k') (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) (mem : k ∈ l.map Prod.fst) - {h'} : - (ofList l).getKey k' h' = k := + {h} : + (ofList l).getKey k' h = k := Raw₀.Const.getKey_insertMany_empty_list_of_mem k_beq distinct mem theorem getKey!_ofList_of_contains_eq_false [EquivBEq α] [LawfulHashable α] [Inhabited α] {l : List (α × β)} {k : α} - (h : (l.map Prod.fst).contains k = false) : + (contains_eq_false : (l.map Prod.fst).contains k = false) : (ofList l).getKey! k = default := - Raw₀.Const.getKey!_insertMany_empty_list_of_contains_eq_false h + Raw₀.Const.getKey!_insertMany_empty_list_of_contains_eq_false contains_eq_false theorem getKey!_ofList_of_mem [EquivBEq α] [LawfulHashable α] [Inhabited α] {l : List (α × β)} @@ -1677,9 +1700,9 @@ theorem getKey!_ofList_of_mem [EquivBEq α] [LawfulHashable α] [Inhabited α] theorem getKeyD_ofList_of_contains_eq_false [EquivBEq α] [LawfulHashable α] {l : List (α × β)} {k fallback : α} - (h : (l.map Prod.fst).contains k = false) : + (contains_eq_false : (l.map Prod.fst).contains k = false) : (ofList l).getKeyD k fallback = fallback := - Raw₀.Const.getKeyD_insertMany_empty_list_of_contains_eq_false h + Raw₀.Const.getKeyD_insertMany_empty_list_of_contains_eq_false contains_eq_false theorem getKeyD_ofList_of_mem [EquivBEq α] [LawfulHashable α] {l : List (α × β)} @@ -1712,11 +1735,12 @@ theorem unitOfList_nil [EquivBEq α] [LawfulHashable α] : @[simp] theorem unitOfList_singleton [EquivBEq α] [LawfulHashable α] {k : α} : - unitOfList [k] = empty.insertIfNew k () := + unitOfList [k] = (∅ : DHashMap α (fun _ => Unit)).insertIfNew k () := Subtype.eq (congrArg Subtype.val (Raw₀.Const.insertManyIfNewUnit_empty_list_singleton (α := α)) :) theorem unitOfList_cons [EquivBEq α] [LawfulHashable α] {hd : α} {tl : List α} : - unitOfList (hd :: tl) = insertManyIfNewUnit (empty.insertIfNew hd ()) tl := + unitOfList (hd :: tl) = + insertManyIfNewUnit ((∅ : DHashMap α (fun _ => Unit)).insertIfNew hd ()) tl := Subtype.eq (congrArg Subtype.val (Raw₀.Const.insertManyIfNewUnit_empty_list_cons (α := α)) :) @[simp] @@ -1726,9 +1750,9 @@ theorem contains_unitOfList [EquivBEq α] [LawfulHashable α] Raw₀.Const.contains_insertManyIfNewUnit_empty_list theorem getKey?_unitOfList_of_contains_eq_false [EquivBEq α] [LawfulHashable α] - {l : List α} {k : α} (h' : l.contains k = false) : + {l : List α} {k : α} (contains_eq_false : l.contains k = false) : getKey? (unitOfList l) k = none := - Raw₀.Const.getKey?_insertManyIfNewUnit_empty_list_of_contains_eq_false h' + Raw₀.Const.getKey?_insertManyIfNewUnit_empty_list_of_contains_eq_false contains_eq_false theorem getKey?_unitOfList_of_mem [EquivBEq α] [LawfulHashable α] {l : List α} {k k' : α} (k_beq : k == k') @@ -1740,15 +1764,15 @@ theorem getKey_unitOfList_of_mem [EquivBEq α] [LawfulHashable α] {l : List α} {k k' : α} (k_beq : k == k') (distinct : l.Pairwise (fun a b => (a == b) = false)) - (mem : k ∈ l) {h'} : - getKey (unitOfList l) k' h' = k := + (mem : k ∈ l) {h} : + getKey (unitOfList l) k' h = k := Raw₀.Const.getKey_insertManyIfNewUnit_empty_list_of_mem k_beq distinct mem theorem getKey!_unitOfList_of_contains_eq_false [EquivBEq α] [LawfulHashable α] [Inhabited α] {l : List α} {k : α} - (h' : l.contains k = false) : + (contains_eq_false : l.contains k = false) : getKey! (unitOfList l) k = default := - Raw₀.Const.getKey!_insertManyIfNewUnit_empty_list_of_contains_eq_false h' + Raw₀.Const.getKey!_insertManyIfNewUnit_empty_list_of_contains_eq_false contains_eq_false theorem getKey!_unitOfList_of_mem [EquivBEq α] [LawfulHashable α] [Inhabited α] {l : List α} {k k' : α} (k_beq : k == k') @@ -1759,9 +1783,9 @@ theorem getKey!_unitOfList_of_mem [EquivBEq α] [LawfulHashable α] theorem getKeyD_unitOfList_of_contains_eq_false [EquivBEq α] [LawfulHashable α] {l : List α} {k fallback : α} - (h' : l.contains k = false) : + (contains_eq_false : l.contains k = false) : getKeyD (unitOfList l) k fallback = fallback := - Raw₀.Const.getKeyD_insertManyIfNewUnit_empty_list_of_contains_eq_false h' + Raw₀.Const.getKeyD_insertManyIfNewUnit_empty_list_of_contains_eq_false contains_eq_false theorem getKeyD_unitOfList_of_mem [EquivBEq α] [LawfulHashable α] {l : List α} {k k' fallback : α} (k_beq : k == k') diff --git a/src/Std/Data/DHashMap/RawLemmas.lean b/src/Std/Data/DHashMap/RawLemmas.lean index 684e0b551528..6356c1c4235f 100644 --- a/src/Std/Data/DHashMap/RawLemmas.lean +++ b/src/Std/Data/DHashMap/RawLemmas.lean @@ -1075,6 +1075,12 @@ theorem contains_insertMany_list [EquivBEq α] [LawfulHashable α] (h : m.WF) (m.insertMany l).contains k = (m.contains k || (l.map Sigma.fst).contains k) := by simp_to_raw using Raw₀.contains_insertMany_list +@[simp] +theorem mem_insertMany_list [EquivBEq α] [LawfulHashable α] (h : m.WF) + {l : List ((a : α) × β a)} {k : α} : + k ∈ (m.insertMany l) ↔ k ∈ m ∨ (l.map Sigma.fst).contains k := by + simp [mem_iff_contains, contains_insertMany_list h] + theorem contains_of_contains_insertMany_list [EquivBEq α] [LawfulHashable α] (h : m.WF) {l : List ((a : α) × β a)} {k : α} : (m.insertMany l).contains k → (l.map Sigma.fst).contains k = false → m.contains k := by @@ -1082,7 +1088,7 @@ theorem contains_of_contains_insertMany_list [EquivBEq α] [LawfulHashable α] ( theorem get?_insertMany_list_of_contains_eq_false [LawfulBEq α] (h : m.WF) {l : List ((a : α) × β a)} {k : α} - (h' : (l.map Sigma.fst).contains k = false) : + (contains_eq_false : (l.map Sigma.fst).contains k = false) : (m.insertMany l).get? k = m.get? k := by simp_to_raw using Raw₀.get?_insertMany_list_of_contains_eq_false @@ -1095,9 +1101,10 @@ theorem get?_insertMany_list_of_mem [LawfulBEq α] (h : m.WF) theorem get_insertMany_list_of_contains_eq_false [LawfulBEq α] (h : m.WF) {l : List ((a : α) × β a)} {k : α} - (contains : (l.map Sigma.fst).contains k = false) + (contains_eq_false : (l.map Sigma.fst).contains k = false) {h'} : - (m.insertMany l).get k h' = m.get k (contains_of_contains_insertMany_list h h' contains) := by + (m.insertMany l).get k h' = + m.get k (contains_of_contains_insertMany_list h h' contains_eq_false) := by simp_to_raw using Raw₀.get_insertMany_list_of_contains_eq_false theorem get_insertMany_list_of_mem [LawfulBEq α] (h : m.WF) @@ -1110,7 +1117,7 @@ theorem get_insertMany_list_of_mem [LawfulBEq α] (h : m.WF) theorem get!_insertMany_list_of_contains_eq_false [LawfulBEq α] (h : m.WF) {l : List ((a : α) × β a)} {k : α} [Inhabited (β k)] - (h' : (l.map Sigma.fst).contains k = false) : + (contains_eq_false : (l.map Sigma.fst).contains k = false) : (m.insertMany l).get! k = m.get! k := by simp_to_raw using Raw₀.get!_insertMany_list_of_contains_eq_false @@ -1136,7 +1143,7 @@ theorem getD_insertMany_list_of_mem [LawfulBEq α] (h : m.WF) theorem getKey?_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] (h : m.WF) {l : List ((a : α) × β a)} {k : α} - (h : (l.map Sigma.fst).contains k = false) : + (contains_eq_false : (l.map Sigma.fst).contains k = false) : (m.insertMany l).getKey? k = m.getKey? k := by simp_to_raw using Raw₀.getKey?_insertMany_list_of_contains_eq_false @@ -1150,9 +1157,10 @@ theorem getKey?_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] (h : m. theorem getKey_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] (h : m.WF) {l : List ((a : α) × β a)} {k : α} - (h₁ : (l.map Sigma.fst).contains k = false) + (contains_eq_false : (l.map Sigma.fst).contains k = false) {h'} : - (m.insertMany l).getKey k h' = m.getKey k (contains_of_contains_insertMany_list h h' h₁) := by + (m.insertMany l).getKey k h' = + m.getKey k (contains_of_contains_insertMany_list h h' contains_eq_false) := by simp_to_raw using Raw₀.getKey_insertMany_list_of_contains_eq_false theorem getKey_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] (h : m.WF) @@ -1166,7 +1174,7 @@ theorem getKey_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] (h : m.W theorem getKey!_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] [Inhabited α] (h : m.WF) {l : List ((a : α) × β a)} {k : α} - (h' : (l.map Sigma.fst).contains k = false) : + (contains_eq_false : (l.map Sigma.fst).contains k = false) : (m.insertMany l).getKey! k = m.getKey! k := by simp_to_raw using Raw₀.getKey!_insertMany_list_of_contains_eq_false @@ -1180,7 +1188,7 @@ theorem getKey!_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] [Inhabi theorem getKeyD_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] (h : m.WF) {l : List ((a : α) × β a)} {k fallback : α} - (h' : (l.map Sigma.fst).contains k = false) : + (contains_eq_false : (l.map Sigma.fst).contains k = false) : (m.insertMany l).getKeyD k fallback = m.getKeyD k fallback := by simp_to_raw using Raw₀.getKeyD_insertMany_list_of_contains_eq_false @@ -1238,6 +1246,12 @@ theorem contains_insertMany_list [EquivBEq α] [LawfulHashable α] (h : m.WF) (insertMany m l).contains k = (m.contains k || (l.map Prod.fst).contains k) := by simp_to_raw using Raw₀.Const.contains_insertMany_list +@[simp] +theorem mem_insertMany_list [EquivBEq α] [LawfulHashable α] (h : m.WF) + {l : List (α × β)} {k : α} : + k ∈ (insertMany m l) ↔ k ∈ m ∨ (l.map Prod.fst).contains k := by + simp [mem_iff_contains, contains_insertMany_list h] + theorem contains_of_contains_insertMany_list [EquivBEq α] [LawfulHashable α] (h : m.WF) {l : List ( α × β )} {k : α} : (insertMany m l).contains k → (l.map Prod.fst).contains k = false → m.contains k := by @@ -1245,7 +1259,7 @@ theorem contains_of_contains_insertMany_list [EquivBEq α] [LawfulHashable α] ( theorem getKey?_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] (h : m.WF) {l : List (α × β)} {k : α} - (h' : (l.map Prod.fst).contains k = false) : + (contains_eq_false : (l.map Prod.fst).contains k = false) : (insertMany m l).getKey? k = m.getKey? k := by simp_to_raw using Raw₀.Const.getKey?_insertMany_list_of_contains_eq_false @@ -1259,9 +1273,10 @@ theorem getKey?_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] (h : m. theorem getKey_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] (h : m.WF) {l : List (α × β)} {k : α} - {h₁ : (l.map Prod.fst).contains k = false} + {contains_eq_false : (l.map Prod.fst).contains k = false} {h'} : - (insertMany m l).getKey k h' = m.getKey k (contains_of_contains_insertMany_list h h' h₁) := by + (insertMany m l).getKey k h' = + m.getKey k (contains_of_contains_insertMany_list h h' contains_eq_false) := by simp_to_raw using Raw₀.Const.getKey_insertMany_list_of_contains_eq_false theorem getKey_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] (h : m.WF) @@ -1275,7 +1290,7 @@ theorem getKey_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] (h : m.W theorem getKey!_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] [Inhabited α] (h : m.WF) {l : List (α × β)} {k : α} - (h' : (l.map Prod.fst).contains k = false) : + (contains_eq_false : (l.map Prod.fst).contains k = false) : (insertMany m l).getKey! k = m.getKey! k := by simp_to_raw using Raw₀.Const.getKey!_insertMany_list_of_contains_eq_false @@ -1289,7 +1304,7 @@ theorem getKey!_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] [Inhabi theorem getKeyD_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] (h : m.WF) {l : List (α × β)} {k fallback : α} - (h' : (l.map Prod.fst).contains k = false) : + (contains_eq_false : (l.map Prod.fst).contains k = false) : (insertMany m l).getKeyD k fallback = m.getKeyD k fallback := by simp_to_raw using Raw₀.Const.getKeyD_insertMany_list_of_contains_eq_false @@ -1321,7 +1336,7 @@ theorem isEmpty_insertMany_list [EquivBEq α] [LawfulHashable α] (h : m.WF) theorem get?_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] (h : m.WF) {l : List (α × β)} {k : α} - (h' : (l.map Prod.fst).contains k = false) : + (contains_eq_false : (l.map Prod.fst).contains k = false) : get? (insertMany m l) k = get? m k := by simp_to_raw using Raw₀.Const.get?_insertMany_list_of_contains_eq_false @@ -1333,9 +1348,10 @@ theorem get?_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] (h : m.WF) theorem get_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] (h : m.WF) {l : List (α × β)} {k : α} - (h₁ : (l.map Prod.fst).contains k = false) + (contains_eq_false : (l.map Prod.fst).contains k = false) {h'} : - get (insertMany m l) k h' = get m k (contains_of_contains_insertMany_list h h' h₁) := by + get (insertMany m l) k h' = + get m k (contains_of_contains_insertMany_list h h' contains_eq_false) := by simp_to_raw using Raw₀.Const.get_insertMany_list_of_contains_eq_false theorem get_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] (h : m.WF) @@ -1346,7 +1362,7 @@ theorem get_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] (h : m.WF) theorem get!_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] [Inhabited β] (h : m.WF) {l : List (α × β)} {k : α} - (h' : (l.map Prod.fst).contains k = false) : + (contains_eq_false : (l.map Prod.fst).contains k = false) : get! (insertMany m l) k = get! m k := by simp_to_raw using Raw₀.Const.get!_insertMany_list_of_contains_eq_false @@ -1358,7 +1374,7 @@ theorem get!_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] [Inhabited theorem getD_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] (h : m.WF) {l : List (α × β)} {k : α} {fallback : β} - (h' : (l.map Prod.fst).contains k = false) : + (contains_eq_false : (l.map Prod.fst).contains k = false) : getD (insertMany m l) k fallback = getD m k fallback := by simp_to_raw using Raw₀.Const.getD_insertMany_list_of_contains_eq_false @@ -1393,13 +1409,19 @@ theorem contains_insertManyIfNewUnit_list [EquivBEq α] [LawfulHashable α] (h : (insertManyIfNewUnit m l).contains k = (m.contains k || l.contains k) := by simp_to_raw using Raw₀.Const.contains_insertManyIfNewUnit_list +@[simp] +theorem mem_insertManyIfNewUnit_list [EquivBEq α] [LawfulHashable α] (h : m.WF) + {l : List α} {k : α} : + k ∈ insertManyIfNewUnit m l ↔ k ∈ m ∨ l.contains k := by + simp [mem_iff_contains, contains_insertManyIfNewUnit_list h] + theorem contains_of_contains_insertManyIfNewUnit_list [EquivBEq α] [LawfulHashable α] (h : m.WF) - {l : List α} {k : α} (h₂ : l.contains k = false) : + {l : List α} {k : α} (contains_eq_false : l.contains k = false) : (insertManyIfNewUnit m l).contains k → m.contains k := by simp_to_raw using Raw₀.Const.contains_of_contains_insertManyIfNewUnit_list theorem getKey?_insertManyIfNewUnit_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] - (h : m.WF) {l : List α} {k : α} (h' : l.contains k = false) : + (h : m.WF) {l : List α} {k : α} (contains_eq_false : l.contains k = false) : getKey? (insertManyIfNewUnit m l) k = getKey? m k := by simp_to_raw using Raw₀.Const.getKey?_insertManyIfNewUnit_list_of_contains_eq_false @@ -1417,9 +1439,9 @@ theorem getKey?_insertManyIfNewUnit_list_of_contains_of_contains [EquivBEq α] [ simp_to_raw using Raw₀.Const.getKey?_insertManyIfNewUnit_list_of_contains_of_contains theorem getKey_insertManyIfNewUnit_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] - (h : m.WF) {l : List α} {k : α} (h₁ : l.contains k = false) {h'} : + (h : m.WF) {l : List α} {k : α} (contains_eq_false : l.contains k = false) {h'} : getKey (insertManyIfNewUnit m l) k h' = - getKey m k (contains_of_contains_insertManyIfNewUnit_list h h₁ h') := by + getKey m k (contains_of_contains_insertManyIfNewUnit_list h contains_eq_false h') := by simp_to_raw using Raw₀.Const.getKey_insertManyIfNewUnit_list_of_contains_eq_false theorem getKey_insertManyIfNewUnit_list_of_mem_of_contains_eq_false [EquivBEq α] [LawfulHashable α] @@ -1439,7 +1461,7 @@ theorem getKey_insertManyIfNewUnit_list_mem_of_contains_of_contains [EquivBEq α theorem getKey!_insertManyIfNewUnit_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] [Inhabited α] (h : m.WF) {l : List α} {k : α} - (h' : l.contains k = false) : + (contains_eq_false : l.contains k = false) : getKey! (insertManyIfNewUnit m l) k = getKey! m k := by simp_to_raw using Raw₀.Const.getKey!_insertManyIfNewUnit_list_of_contains_eq_false @@ -1459,7 +1481,7 @@ theorem getKey!_insertManyIfNewUnit_list_mem_of_contains_of_contains [EquivBEq theorem getKeyD_insertManyIfNewUnit_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] (h : m.WF) {l : List α} {k fallback : α} - (h' : l.contains k = false) : + (contains_eq_false : l.contains k = false) : getKeyD (insertManyIfNewUnit m l) k fallback = getKeyD m k fallback := by simp_to_raw using Raw₀.Const.getKeyD_insertManyIfNewUnit_list_of_contains_eq_false @@ -1531,13 +1553,13 @@ variable [BEq α] [Hashable α] open Internal.Raw Internal.Raw₀ @[simp] theorem ofList_nil [EquivBEq α] [LawfulHashable α] : - ofList ([] : List ((a : α) × (β a))) = empty := by + ofList ([] : List ((a : α) × (β a))) = ∅ := by simp_to_raw rw [Raw₀.insertMany_empty_list_nil] @[simp] theorem ofList_singleton {k : α} {v : β k} [EquivBEq α] [LawfulHashable α] : - ofList [⟨k, v⟩] = empty.insert k v := by + ofList [⟨k, v⟩] = (∅ : Raw α β).insert k v := by simp_to_raw rw [Raw₀.insertMany_empty_list_singleton] @@ -1552,9 +1574,15 @@ theorem contains_ofList [EquivBEq α] [LawfulHashable α] (ofList l).contains k = (l.map Sigma.fst).contains k := by simp_to_raw using Raw₀.contains_insertMany_empty_list +@[simp] +theorem mem_ofList [EquivBEq α] [LawfulHashable α] + {l : List ((a : α) × β a)} {k : α} : + k ∈ ofList l ↔ (l.map Sigma.fst).contains k := by + simp [mem_iff_contains] + theorem get?_ofList_of_contains_eq_false [LawfulBEq α] {l : List ((a : α) × β a)} {k : α} - (h : (l.map Sigma.fst).contains k = false) : + (contains_eq_false : (l.map Sigma.fst).contains k = false) : (ofList l).get? k = none := by simp_to_raw using Raw₀.get?_insertMany_empty_list_of_contains_eq_false @@ -1575,7 +1603,7 @@ theorem get_ofList_of_mem [LawfulBEq α] theorem get!_ofList_of_contains_eq_false [LawfulBEq α] {l : List ((a : α) × β a)} {k : α} [Inhabited (β k)] - (h : (l.map Sigma.fst).contains k = false) : + (contains_eq_false : (l.map Sigma.fst).contains k = false) : (ofList l).get! k = default := by simp_to_raw using get!_insertMany_empty_list_of_contains_eq_false @@ -1601,7 +1629,7 @@ theorem getD_ofList_of_mem [LawfulBEq α] theorem getKey?_ofList_of_contains_eq_false [EquivBEq α] [LawfulHashable α] {l : List ((a : α) × β a)} {k : α} - (h : (l.map Sigma.fst).contains k = false) : + (contains_eq_false : (l.map Sigma.fst).contains k = false) : (ofList l).getKey? k = none := by simp_to_raw using Raw₀.getKey?_insertMany_empty_list_of_contains_eq_false @@ -1624,7 +1652,7 @@ theorem getKey_ofList_of_mem [EquivBEq α] [LawfulHashable α] theorem getKey!_ofList_of_contains_eq_false [EquivBEq α] [LawfulHashable α] [Inhabited α] {l : List ((a : α) × β a)} {k : α} - (h : (l.map Sigma.fst).contains k = false) : + (contains_eq_false : (l.map Sigma.fst).contains k = false) : (ofList l).getKey! k = default := by simp_to_raw using Raw₀.getKey!_insertMany_empty_list_of_contains_eq_false @@ -1638,7 +1666,7 @@ theorem getKey!_ofList_of_mem [EquivBEq α] [LawfulHashable α] [Inhabited α] theorem getKeyD_ofList_of_contains_eq_false [EquivBEq α] [LawfulHashable α] {l : List ((a : α) × β a)} {k fallback : α} - (h : (l.map Sigma.fst).contains k = false) : + (contains_eq_false : (l.map Sigma.fst).contains k = false) : (ofList l).getKeyD k fallback = fallback := by simp_to_raw using Raw₀.getKeyD_insertMany_empty_list_of_contains_eq_false @@ -1672,13 +1700,13 @@ variable {β : Type v} @[simp] theorem ofList_nil [EquivBEq α] [LawfulHashable α] : - ofList ([] : List (α × β)) = empty := by + ofList ([] : List (α × β)) = ∅ := by simp_to_raw simp @[simp] -theorem ofList_singleton {k : α} {v : β} [EquivBEq α] [LawfulHashable α] : - ofList [⟨k, v⟩] = empty.insert k v := by +theorem ofList_singleton [EquivBEq α] [LawfulHashable α] {k : α} {v : β} : + ofList [⟨k, v⟩] = (∅ : Raw α (fun _ => β)).insert k v := by simp_to_raw simp @@ -1693,9 +1721,15 @@ theorem contains_ofList [EquivBEq α] [LawfulHashable α] (ofList l).contains k = (l.map Prod.fst).contains k := by simp_to_raw using Raw₀.Const.contains_insertMany_empty_list +@[simp] +theorem mem_ofList [EquivBEq α] [LawfulHashable α] + {l : List (α × β)} {k : α} : + k ∈ (ofList l) ↔ (l.map Prod.fst).contains k := by + simp [mem_iff_contains] + theorem get?_ofList_of_contains_eq_false [LawfulBEq α] {l : List (α × β)} {k : α} - (h : (l.map Prod.fst).contains k = false) : + (contains_eq_false : (l.map Prod.fst).contains k = false) : get? (ofList l) k = none := by simp_to_raw using Raw₀.Const.get?_insertMany_empty_list_of_contains_eq_false @@ -1716,7 +1750,7 @@ theorem get_ofList_of_mem [LawfulBEq α] theorem get!_ofList_of_contains_eq_false [LawfulBEq α] {l : List (α × β)} {k : α} [Inhabited β] - (h : (l.map Prod.fst).contains k = false) : + (contains_eq_false : (l.map Prod.fst).contains k = false) : get! (ofList l) k = default := by simp_to_raw using Raw₀.Const.get!_insertMany_empty_list_of_contains_eq_false @@ -1742,7 +1776,7 @@ theorem getD_ofList_of_mem [LawfulBEq α] theorem getKey?_ofList_of_contains_eq_false [EquivBEq α] [LawfulHashable α] {l : List (α × β)} {k : α} - (h : (l.map Prod.fst).contains k = false) : + (contains_eq_false : (l.map Prod.fst).contains k = false) : (ofList l).getKey? k = none := by simp_to_raw using Raw₀.Const.getKey?_insertMany_empty_list_of_contains_eq_false @@ -1765,7 +1799,7 @@ theorem getKey_ofList_of_mem [EquivBEq α] [LawfulHashable α] theorem getKey!_ofList_of_contains_eq_false [EquivBEq α] [LawfulHashable α] [Inhabited α] {l : List (α × β)} {k : α} - (h : (l.map Prod.fst).contains k = false) : + (contains_eq_false : (l.map Prod.fst).contains k = false) : (ofList l).getKey! k = default := by simp_to_raw using Raw₀.Const.getKey!_insertMany_empty_list_of_contains_eq_false @@ -1779,7 +1813,7 @@ theorem getKey!_ofList_of_mem [EquivBEq α] [LawfulHashable α] [Inhabited α] theorem getKeyD_ofList_of_contains_eq_false [EquivBEq α] [LawfulHashable α] {l : List (α × β)} {k fallback : α} - (h : (l.map Prod.fst).contains k = false) : + (contains_eq_false : (l.map Prod.fst).contains k = false) : (ofList l).getKeyD k fallback = fallback := by simp_to_raw using Raw₀.Const.getKeyD_insertMany_empty_list_of_contains_eq_false @@ -1809,13 +1843,13 @@ theorem isEmpty_ofList [EquivBEq α] [LawfulHashable α] @[simp] theorem unitOfList_nil [EquivBEq α] [LawfulHashable α] : - unitOfList ([] : List α) = empty := by + unitOfList ([] : List α) = ∅ := by simp_to_raw simp @[simp] theorem unitOfList_singleton [EquivBEq α] [LawfulHashable α] {k : α} : - unitOfList [k] = empty.insertIfNew k () := by + unitOfList [k] = (∅ : Raw α (fun _ => Unit)).insertIfNew k () := by simp_to_raw simp @@ -1830,8 +1864,14 @@ theorem contains_unitOfList [EquivBEq α] [LawfulHashable α] (unitOfList l).contains k = l.contains k := by simp_to_raw using Raw₀.Const.contains_insertManyIfNewUnit_empty_list +@[simp] +theorem mem_unitOfList [EquivBEq α] [LawfulHashable α] + {l : List α} {k : α} : + k ∈ unitOfList l ↔ l.contains k := by + simp [mem_iff_contains] + theorem getKey?_unitOfList_of_contains_eq_false [EquivBEq α] [LawfulHashable α] - {l : List α} {k : α} (h' : l.contains k = false) : + {l : List α} {k : α} (contains_eq_false : l.contains k = false) : getKey? (unitOfList l) k = none := by simp_to_raw using Raw₀.Const.getKey?_insertManyIfNewUnit_empty_list_of_contains_eq_false @@ -1851,7 +1891,7 @@ theorem getKey_unitOfList_of_mem [EquivBEq α] [LawfulHashable α] theorem getKey!_unitOfList_of_contains_eq_false [EquivBEq α] [LawfulHashable α] [Inhabited α] {l : List α} {k : α} - (h' : l.contains k = false) : + (contains_eq_false : l.contains k = false) : getKey! (unitOfList l) k = default := by simp_to_raw using Raw₀.Const.getKey!_insertManyIfNewUnit_empty_list_of_contains_eq_false @@ -1864,7 +1904,7 @@ theorem getKey!_unitOfList_of_mem [EquivBEq α] [LawfulHashable α] theorem getKeyD_unitOfList_of_contains_eq_false [EquivBEq α] [LawfulHashable α] {l : List α} {k fallback : α} - (h' : l.contains k = false) : + (contains_eq_false : l.contains k = false) : getKeyD (unitOfList l) k fallback = fallback := by simp_to_raw using Raw₀.Const.getKeyD_insertManyIfNewUnit_empty_list_of_contains_eq_false From e5ac6491c83f4ef9b9d09ad284b4ea9734f6754b Mon Sep 17 00:00:00 2001 From: jt0202 Date: Thu, 9 Jan 2025 11:38:03 +0100 Subject: [PATCH 69/83] Style fixes --- src/Std/Data/DHashMap/Internal/Model.lean | 6 +++--- src/Std/Data/DHashMap/Internal/Raw.lean | 6 +++--- src/Std/Data/DHashMap/Lemmas.lean | 20 +++++++++++++++++++- 3 files changed, 25 insertions(+), 7 deletions(-) diff --git a/src/Std/Data/DHashMap/Internal/Model.lean b/src/Std/Data/DHashMap/Internal/Model.lean index 0dc01e257312..3ef7cbd46ce4 100644 --- a/src/Std/Data/DHashMap/Internal/Model.lean +++ b/src/Std/Data/DHashMap/Internal/Model.lean @@ -485,7 +485,7 @@ theorem insertMany_eq_insertListₘ [BEq α] [Hashable α](m : Raw₀ α β) (l induction l generalizing m with | nil => simp [insertListₘ] | cons hd tl ih => - simp only [List.foldl_cons,insertListₘ] + simp only [List.foldl_cons, insertListₘ] apply ih section @@ -530,7 +530,7 @@ theorem Const.insertMany_eq_insertListₘ [BEq α] [Hashable α] (m : Raw₀ α induction l generalizing m with | nil => simp [insertListₘ] | cons hd tl ih => - simp only [List.foldl_cons,insertListₘ] + simp only [List.foldl_cons, insertListₘ] apply ih theorem Const.insertManyIfNewUnit_eq_insertListIfNewUnitₘ [BEq α] [Hashable α] @@ -545,7 +545,7 @@ theorem Const.insertManyIfNewUnit_eq_insertListIfNewUnitₘ [BEq α] [Hashable induction l generalizing m with | nil => simp [insertListIfNewUnitₘ] | cons hd tl ih => - simp only [List.foldl_cons,insertListIfNewUnitₘ] + simp only [List.foldl_cons, insertListIfNewUnitₘ] apply ih end diff --git a/src/Std/Data/DHashMap/Internal/Raw.lean b/src/Std/Data/DHashMap/Internal/Raw.lean index 5fd8539388d1..4b28a7153eae 100644 --- a/src/Std/Data/DHashMap/Internal/Raw.lean +++ b/src/Std/Data/DHashMap/Internal/Raw.lean @@ -207,7 +207,7 @@ theorem insertMany_val [BEq α][Hashable α] {m : Raw₀ α β} {ρ : Type w} [F m.val.insertMany l = m.insertMany l := by simp [Raw.insertMany, m.2] -theorem ofList_eq [BEq α] [Hashable α] {l: List ((a : α) × β a)} : +theorem ofList_eq [BEq α] [Hashable α] {l : List ((a : α) × β a)} : Raw.ofList l = Raw₀.insertMany Raw₀.empty l := by simp only [Raw.ofList, Raw.insertMany, (Raw.WF.empty).size_buckets_pos ∅, ↓reduceDIte] congr @@ -224,7 +224,7 @@ theorem Const.insertMany_val [BEq α][Hashable α] {m : Raw₀ α (fun _ => β)} Raw.Const.insertMany m.val l = Raw₀.Const.insertMany m l := by simp [Raw.Const.insertMany, m.2] -theorem Const.ofList_eq [BEq α] [Hashable α] {l: List (α × β)} : +theorem Const.ofList_eq [BEq α] [Hashable α] {l : List (α × β)} : Raw.Const.ofList l = Raw₀.Const.insertMany Raw₀.empty l := by simp only [Raw.Const.ofList, Raw.Const.insertMany, (Raw.WF.empty).size_buckets_pos ∅, ↓reduceDIte] congr @@ -239,7 +239,7 @@ theorem Const.insertManyIfNewUnit_val {ρ : Type w} [ForIn Id ρ α] [BEq α] [H Raw.Const.insertManyIfNewUnit m.val l = Raw₀.Const.insertManyIfNewUnit m l := by simp [Raw.Const.insertManyIfNewUnit, m.2] -theorem Const.unitOfList_eq [BEq α] [Hashable α] {l: List α} : +theorem Const.unitOfList_eq [BEq α] [Hashable α] {l : List α} : Raw.Const.unitOfList l = Raw₀.Const.insertManyIfNewUnit Raw₀.empty l := by simp only [Raw.Const.unitOfList, Raw.Const.insertManyIfNewUnit, (Raw.WF.empty).size_buckets_pos ∅, ↓reduceDIte] diff --git a/src/Std/Data/DHashMap/Lemmas.lean b/src/Std/Data/DHashMap/Lemmas.lean index 9a244f5bf7f8..83fd2af88e31 100644 --- a/src/Std/Data/DHashMap/Lemmas.lean +++ b/src/Std/Data/DHashMap/Lemmas.lean @@ -991,7 +991,7 @@ theorem mem_insertMany_list [EquivBEq α] [LawfulHashable α] theorem contains_of_contains_insertMany_list [EquivBEq α] [LawfulHashable α] {l : List ((a : α) × β a)} {k : α} (contains : (m.insertMany l).contains k) - (contains_eq_false: (l.map Sigma.fst).contains k = false) : + (contains_eq_false : (l.map Sigma.fst).contains k = false) : m.contains k := Raw₀.contains_of_contains_insertMany_list ⟨m.1, _⟩ m.2 contains contains_eq_false @@ -1476,6 +1476,12 @@ theorem contains_ofList [EquivBEq α] [LawfulHashable α] (ofList l).contains k = (l.map Sigma.fst).contains k := Raw₀.contains_insertMany_empty_list +@[simp] +theorem mem_ofList [EquivBEq α] [LawfulHashable α] + {l : List ((a : α) × β a)} {k : α} : + k ∈ ofList l ↔ (l.map Sigma.fst).contains k := by + simp [mem_iff_contains] + theorem get?_ofList_of_contains_eq_false [LawfulBEq α] {l : List ((a : α) × β a)} {k : α} (contains_eq_false : (l.map Sigma.fst).contains k = false) : @@ -1614,6 +1620,12 @@ theorem contains_ofList [EquivBEq α] [LawfulHashable α] (ofList l).contains k = (l.map Prod.fst).contains k := Raw₀.Const.contains_insertMany_empty_list +@[simp] +theorem mem_ofList [EquivBEq α] [LawfulHashable α] + {l : List (α × β)} {k : α} : + k ∈ ofList l ↔ (l.map Prod.fst).contains k := by + simp [mem_iff_contains] + theorem get?_ofList_of_contains_eq_false [LawfulBEq α] {l : List (α × β)} {k : α} (contains_eq_false : (l.map Prod.fst).contains k = false) : @@ -1749,6 +1761,12 @@ theorem contains_unitOfList [EquivBEq α] [LawfulHashable α] (unitOfList l).contains k = l.contains k := Raw₀.Const.contains_insertManyIfNewUnit_empty_list +@[simp] +theorem mem_unitOfList [EquivBEq α] [LawfulHashable α] + {l : List α} {k : α} : + k ∈ unitOfList l ↔ l.contains k := by + simp [mem_iff_contains] + theorem getKey?_unitOfList_of_contains_eq_false [EquivBEq α] [LawfulHashable α] {l : List α} {k : α} (contains_eq_false : l.contains k = false) : getKey? (unitOfList l) k = none := From 447c8e4a7032fc14c3e119a33409479a8299ebea Mon Sep 17 00:00:00 2001 From: jt0202 Date: Thu, 9 Jan 2025 14:35:43 +0100 Subject: [PATCH 70/83] Results for HashMaps --- src/Std/Data/DHashMap/Lemmas.lean | 10 +- src/Std/Data/DHashMap/RawLemmas.lean | 8 +- src/Std/Data/HashMap/Lemmas.lean | 570 ++++++++++++++++++++++++++ src/Std/Data/HashMap/RawLemmas.lean | 584 ++++++++++++++++++++++++++- 4 files changed, 1160 insertions(+), 12 deletions(-) diff --git a/src/Std/Data/DHashMap/Lemmas.lean b/src/Std/Data/DHashMap/Lemmas.lean index 83fd2af88e31..5ca7feee8fb9 100644 --- a/src/Std/Data/DHashMap/Lemmas.lean +++ b/src/Std/Data/DHashMap/Lemmas.lean @@ -1128,7 +1128,7 @@ theorem isEmpty_insertMany_list [EquivBEq α] [LawfulHashable α] namespace Const -variable {β : Type v} (m : DHashMap α (fun _ => β)) +variable {β : Type v} {m : DHashMap α (fun _ => β)} @[simp] theorem insertMany_nil : @@ -1182,7 +1182,7 @@ theorem getKey_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashabl (contains_eq_false : (l.map Prod.fst).contains k = false) {h} : (insertMany m l).getKey k h = - m.getKey k (contains_of_contains_insertMany_list _ h contains_eq_false) := + m.getKey k (contains_of_contains_insertMany_list h contains_eq_false) := Raw₀.Const.getKey_insertMany_list_of_contains_eq_false ⟨m.1, _⟩ m.2 contains_eq_false theorem getKey_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] @@ -1256,7 +1256,7 @@ theorem get_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable {l : List (α × β)} {k : α} (contains_eq_false : (l.map Prod.fst).contains k = false) {h} : - get (insertMany m l) k h = get m k (contains_of_contains_insertMany_list _ h contains_eq_false) := + get (insertMany m l) k h = get m k (contains_of_contains_insertMany_list h contains_eq_false) := Raw₀.Const.get_insertMany_list_of_contains_eq_false ⟨m.1, _⟩ m.2 contains_eq_false theorem get_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] @@ -1289,7 +1289,7 @@ theorem getD_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] getD (insertMany m l) k' fallback = v := Raw₀.Const.getD_insertMany_list_of_mem ⟨m.1, _⟩ m.2 k_beq distinct mem -variable (m : DHashMap α (fun _ => Unit)) +variable {m : DHashMap α (fun _ => Unit)} @[simp] theorem insertManyIfNewUnit_nil : @@ -1349,7 +1349,7 @@ theorem getKey?_insertManyIfNewUnit_list_of_contains_of_contains [EquivBEq α] [ theorem getKey_insertManyIfNewUnit_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] {l : List α} {k : α} (contains_eq_false : l.contains k = false) {h} : getKey (insertManyIfNewUnit m l) k h = - getKey m k (contains_of_contains_insertManyIfNewUnit_list m contains_eq_false h) := + getKey m k (contains_of_contains_insertManyIfNewUnit_list contains_eq_false h) := Raw₀.Const.getKey_insertManyIfNewUnit_list_of_contains_eq_false ⟨m.1, _⟩ m.2 contains_eq_false theorem getKey_insertManyIfNewUnit_list_of_mem_of_contains_eq_false [EquivBEq α] [LawfulHashable α] diff --git a/src/Std/Data/DHashMap/RawLemmas.lean b/src/Std/Data/DHashMap/RawLemmas.lean index 6356c1c4235f..13c014050738 100644 --- a/src/Std/Data/DHashMap/RawLemmas.lean +++ b/src/Std/Data/DHashMap/RawLemmas.lean @@ -1273,7 +1273,7 @@ theorem getKey?_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] (h : m. theorem getKey_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] (h : m.WF) {l : List (α × β)} {k : α} - {contains_eq_false : (l.map Prod.fst).contains k = false} + (contains_eq_false : (l.map Prod.fst).contains k = false) {h'} : (insertMany m l).getKey k h' = m.getKey k (contains_of_contains_insertMany_list h h' contains_eq_false) := by @@ -1455,8 +1455,8 @@ theorem getKey_insertManyIfNewUnit_list_of_mem_of_contains_eq_false [EquivBEq α theorem getKey_insertManyIfNewUnit_list_mem_of_contains_of_contains [EquivBEq α] [LawfulHashable α] (h : m.WF) {l : List α} {k : α} (distinct : l.Pairwise (fun a b => (a == b) = false)) - (h' : l.contains k) {mem' : m.contains k} {h'} : - getKey (insertManyIfNewUnit m l) k h' = getKey m k mem' := by + (h₁ : l.contains k) (h₂ : m.contains k) {h₃} : + getKey (insertManyIfNewUnit m l) k h₃ = getKey m k h₂ := by simp_to_raw using Raw₀.Const.getKey_insertManyIfNewUnit_list_mem_of_contains_of_contains theorem getKey!_insertManyIfNewUnit_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] @@ -1911,7 +1911,7 @@ theorem getKeyD_unitOfList_of_contains_eq_false [EquivBEq α] [LawfulHashable α theorem getKeyD_unitOfList_of_mem [EquivBEq α] [LawfulHashable α] {l : List α} {k k' fallback : α} (k_beq : k == k') (distinct : l.Pairwise (fun a b => (a == b) = false)) - {mem : k ∈ l } : + (mem : k ∈ l ) : getKeyD (unitOfList l) k' fallback = k := by simp_to_raw using Raw₀.Const.getKeyD_insertManyIfNewUnit_empty_list_of_mem diff --git a/src/Std/Data/HashMap/Lemmas.lean b/src/Std/Data/HashMap/Lemmas.lean index 5e9f89e692fa..8fa3164915c7 100644 --- a/src/Std/Data/HashMap/Lemmas.lean +++ b/src/Std/Data/HashMap/Lemmas.lean @@ -698,6 +698,576 @@ theorem distinct_keys [EquivBEq α] [LawfulHashable α] : m.keys.Pairwise (fun a b => (a == b) = false) := DHashMap.distinct_keys +@[simp] +theorem insertMany_nil : + insertMany m [] = m := + ext DHashMap.Const.insertMany_nil + +@[simp] +theorem insertMany_list_singleton {k : α} {v : β} : + insertMany m [⟨k, v⟩] = m.insert k v := + ext DHashMap.Const.insertMany_list_singleton + +theorem insertMany_cons {l : List (α × β)} {k : α} {v : β} : + insertMany m (⟨k, v⟩ :: l) = insertMany (m.insert k v) l := + ext DHashMap.Const.insertMany_cons + +@[simp] +theorem contains_insertMany_list [EquivBEq α] [LawfulHashable α] + {l : List (α × β)} {k : α} : + (insertMany m l).contains k = (m.contains k || (l.map Prod.fst).contains k) := + DHashMap.Const.contains_insertMany_list + +@[simp] +theorem mem_insertMany_list [EquivBEq α] [LawfulHashable α] + {l : List (α × β)} {k : α} : + k ∈ insertMany m l ↔ k ∈ m ∨ (l.map Prod.fst).contains k := + DHashMap.Const.mem_insertMany_list + +theorem contains_of_contains_insertMany_list [EquivBEq α] [LawfulHashable α] + {l : List ( α × β )} {k : α} (contains : (insertMany m l).contains k) + (contains_eq_false : (l.map Prod.fst).contains k = false) : + m.contains k := + DHashMap.Const.contains_of_contains_insertMany_list contains contains_eq_false + +theorem getKey?_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] + {l : List (α × β)} {k : α} + (contains_eq_false : (l.map Prod.fst).contains k = false) : + (insertMany m l).getKey? k = m.getKey? k := + DHashMap.Const.getKey?_insertMany_list_of_contains_eq_false contains_eq_false + +theorem getKey?_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] + {l : List (α × β)} + {k k' : α} (k_beq : k == k') + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) + (mem : k ∈ l.map Prod.fst) : + (insertMany m l).getKey? k' = some k := + DHashMap.Const.getKey?_insertMany_list_of_mem k_beq distinct mem + +theorem getKey_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] + {l : List (α × β)} {k : α} + (contains_eq_false : (l.map Prod.fst).contains k = false) + {h} : + (insertMany m l).getKey k h = + m.getKey k (contains_of_contains_insertMany_list h contains_eq_false) := + DHashMap.Const.getKey_insertMany_list_of_contains_eq_false contains_eq_false + +theorem getKey_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] + {l : List (α × β)} + {k k' : α} (k_beq : k == k') + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) + (mem : k ∈ l.map Prod.fst) + {h} : + (insertMany m l).getKey k' h = k := + DHashMap.Const.getKey_insertMany_list_of_mem k_beq distinct mem + +theorem getKey!_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] [Inhabited α] + {l : List (α × β)} {k : α} + (contains_eq_false : (l.map Prod.fst).contains k = false) : + (insertMany m l).getKey! k = m.getKey! k := + DHashMap.Const.getKey!_insertMany_list_of_contains_eq_false contains_eq_false + +theorem getKey!_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] [Inhabited α] + {l : List (α × β)} + {k k' : α} (k_beq : k == k') + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) + (mem : k ∈ l.map Prod.fst) : + (insertMany m l).getKey! k' = k := + DHashMap.Const.getKey!_insertMany_list_of_mem k_beq distinct mem + +theorem getKeyD_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] + {l : List (α × β)} {k fallback : α} + (contains_eq_false : (l.map Prod.fst).contains k = false) : + (insertMany m l).getKeyD k fallback = m.getKeyD k fallback := + DHashMap.Const.getKeyD_insertMany_list_of_contains_eq_false contains_eq_false + +theorem getKeyD_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] + {l : List (α × β)} + {k k' fallback : α} (k_beq : k == k') + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) + (mem : k ∈ l.map Prod.fst) : + (insertMany m l).getKeyD k' fallback = k := + DHashMap.Const.getKeyD_insertMany_list_of_mem k_beq distinct mem + +theorem size_insertMany_list [EquivBEq α] [LawfulHashable α] + {l : List (α × β)} + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) : + (∀ (a : α), m.contains a → (l.map Prod.fst).contains a = false) → + (insertMany m l).size = m.size + l.length := + DHashMap.Const.size_insertMany_list distinct + +theorem size_insertMany_list_le [EquivBEq α] [LawfulHashable α] + {l : List (α × β)} : + (insertMany m l).size ≤ m.size + l.length := + DHashMap.Const.size_insertMany_list_le + +@[simp] +theorem isEmpty_insertMany_list [EquivBEq α] [LawfulHashable α] + {l : List (α × β)} : + (insertMany m l).isEmpty = (m.isEmpty && l.isEmpty) := + DHashMap.Const.isEmpty_insertMany_list + +theorem get?_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] + {l : List (α × β)} {k : α} + (contains_eq_false : (l.map Prod.fst).contains k = false) : + get? (insertMany m l) k = get? m k := + DHashMap.Const.get?_insertMany_list_of_contains_eq_false contains_eq_false + +theorem get?_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] + {l : List (α × β)} {k k' : α} (k_beq : k == k') {v : β} + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) (mem : ⟨k, v⟩ ∈ l) : + get? (insertMany m l) k' = v := + DHashMap.Const.get?_insertMany_list_of_mem k_beq distinct mem + +theorem get_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] + {l : List (α × β)} {k : α} + (contains_eq_false : (l.map Prod.fst).contains k = false) + {h} : + get (insertMany m l) k h = get m k (contains_of_contains_insertMany_list h contains_eq_false) := + DHashMap.Const.get_insertMany_list_of_contains_eq_false contains_eq_false + +theorem get_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] + {l : List (α × β)} {k k' : α} (k_beq : k == k') {v : β} + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) (mem : ⟨k, v⟩ ∈ l) {h} : + get (insertMany m l) k' h = v := + DHashMap.Const.get_insertMany_list_of_mem k_beq distinct mem + +theorem get!_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] + [Inhabited β] {l : List (α × β)} {k : α} + (contains_eq_false : (l.map Prod.fst).contains k = false) : + get! (insertMany m l) k = get! m k := + DHashMap.Const.get!_insertMany_list_of_contains_eq_false contains_eq_false + +theorem get!_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] [Inhabited β] + {l : List (α × β)} {k k' : α} (k_beq : k == k') {v : β} + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) (mem : ⟨k, v⟩ ∈ l) : + get! (insertMany m l) k' = v := + DHashMap.Const.get!_insertMany_list_of_mem k_beq distinct mem + +theorem getD_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] + {l : List (α × β)} {k : α} {fallback : β} + (contains_eq_false : (l.map Prod.fst).contains k = false) : + getD (insertMany m l) k fallback = getD m k fallback := + DHashMap.Const.getD_insertMany_list_of_contains_eq_false contains_eq_false + +theorem getD_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] + {l : List (α × β)} {k k' : α} (k_beq : k == k') {v fallback : β} + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) (mem : ⟨k, v⟩ ∈ l) : + getD (insertMany m l) k' fallback = v := + DHashMap.Const.getD_insertMany_list_of_mem k_beq distinct mem + +variable {m : HashMap α Unit} + +@[simp] +theorem insertManyIfNewUnit_nil : + insertManyIfNewUnit m [] = m := + ext DHashMap.Const.insertManyIfNewUnit_nil + +@[simp] +theorem insertManyIfNewUnit_list_singleton [EquivBEq α] [LawfulHashable α] {k : α} : + insertManyIfNewUnit m [k] = m.insertIfNew k () := + ext DHashMap.Const.insertManyIfNewUnit_list_singleton + +theorem insertManyIfNewUnit_cons {l : List α} {k : α} : + insertManyIfNewUnit m (k :: l) = insertManyIfNewUnit (m.insertIfNew k ()) l := + ext DHashMap.Const.insertManyIfNewUnit_cons + +@[simp] +theorem contains_insertManyIfNewUnit_list [EquivBEq α] [LawfulHashable α] + {l : List α} {k : α} : + (insertManyIfNewUnit m l).contains k = (m.contains k || l.contains k) := + DHashMap.Const.contains_insertManyIfNewUnit_list + +@[simp] +theorem mem_insertManyIfNewUnit_list [EquivBEq α] [LawfulHashable α] + {l : List α} {k : α} : + k ∈ insertManyIfNewUnit m l ↔ k ∈ m ∨ l.contains k := + DHashMap.Const.mem_insertManyIfNewUnit_list + +theorem contains_of_contains_insertManyIfNewUnit_list [EquivBEq α] [LawfulHashable α] + {l : List α} {k : α} (contains_eq_false : l.contains k = false) : + (insertManyIfNewUnit m l).contains k → m.contains k := + DHashMap.Const.contains_of_contains_insertManyIfNewUnit_list contains_eq_false + +theorem getKey?_insertManyIfNewUnit_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] + {l : List α} {k : α} (contains_eq_false : l.contains k = false) : + getKey? (insertManyIfNewUnit m l) k = getKey? m k := + DHashMap.Const.getKey?_insertManyIfNewUnit_list_of_contains_eq_false contains_eq_false + +theorem getKey?_insertManyIfNewUnit_list_of_mem_of_contains_eq_false [EquivBEq α] [LawfulHashable α] + {l : List α} {k k' : α} (k_beq : k == k') + (distinct : l.Pairwise (fun a b => (a == b) = false)) (mem : k ∈ l) + (contains_eq_false : m.contains k = false) : + getKey? (insertManyIfNewUnit m l) k' = some k := + DHashMap.Const.getKey?_insertManyIfNewUnit_list_of_mem_of_contains_eq_false k_beq distinct mem contains_eq_false + +theorem getKey?_insertManyIfNewUnit_list_of_contains_of_contains [EquivBEq α] [LawfulHashable α] + {l : List α} {k : α} + (distinct : l.Pairwise (fun a b => (a == b) = false)) + (h : l.contains k) + (h' : m.contains k) : + getKey? (insertManyIfNewUnit m l) k = getKey? m k := + DHashMap.Const.getKey?_insertManyIfNewUnit_list_of_contains_of_contains distinct h h' + +theorem getKey_insertManyIfNewUnit_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] + {l : List α} {k : α} (contains_eq_false : l.contains k = false) {h} : + getKey (insertManyIfNewUnit m l) k h = + getKey m k (contains_of_contains_insertManyIfNewUnit_list contains_eq_false h) := + DHashMap.Const.getKey_insertManyIfNewUnit_list_of_contains_eq_false contains_eq_false + +theorem getKey_insertManyIfNewUnit_list_of_mem_of_contains_eq_false [EquivBEq α] [LawfulHashable α] + {l : List α} + {k k' : α} (k_beq : k == k') + (distinct : l.Pairwise (fun a b => (a == b) = false)) + (mem : k ∈ l) {h} + (contains_eq_false : m.contains k = false) : + getKey (insertManyIfNewUnit m l) k' h = k := + DHashMap.Const.getKey_insertManyIfNewUnit_list_of_mem_of_contains_eq_false k_beq distinct mem contains_eq_false + +theorem getKey_insertManyIfNewUnit_list_mem_of_contains_of_contains [EquivBEq α] [LawfulHashable α] + {l : List α} {k : α} + (distinct : l.Pairwise (fun a b => (a == b) = false)) + (h₁ : l.contains k) {h₂ : m.contains k} {h} : + getKey (insertManyIfNewUnit m l) k h = getKey m k h₂ := + DHashMap.Const.getKey_insertManyIfNewUnit_list_mem_of_contains_of_contains distinct h₁ + +theorem getKey!_insertManyIfNewUnit_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] + [Inhabited α] {l : List α} {k : α} + (contains_eq_false : l.contains k = false) : + getKey! (insertManyIfNewUnit m l) k = getKey! m k := + DHashMap.Const.getKey!_insertManyIfNewUnit_list_of_contains_eq_false contains_eq_false + +theorem getKey!_insertManyIfNewUnit_list_of_mem_of_contains_eq_false [EquivBEq α] [LawfulHashable α] + [Inhabited α] {l : List α} {k k' : α} (k_beq : k == k') + (distinct : l.Pairwise (fun a b => (a == b) = false)) + (mem : k ∈ l) : + contains m k = false → getKey! (insertManyIfNewUnit m l) k' = k := + DHashMap.Const.getKey!_insertManyIfNewUnit_list_of_mem_of_contains_eq_false k_beq distinct mem + +theorem getKey!_insertManyIfNewUnit_list_mem_of_contains_of_contains [EquivBEq α] [LawfulHashable α] + [Inhabited α] {l : List α} {k : α} + (distinct : l.Pairwise (fun a b => (a == b) = false)) + (h : l.contains k) : + m.contains k → getKey! (insertManyIfNewUnit m l) k = getKey! m k := + DHashMap.Const.getKey!_insertManyIfNewUnit_list_mem_of_contains_of_contains distinct h + +theorem getKeyD_insertManyIfNewUnit_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] + {l : List α} {k fallback : α} + (contains_eq_false : l.contains k = false) : + getKeyD (insertManyIfNewUnit m l) k fallback = getKeyD m k fallback := + DHashMap.Const.getKeyD_insertManyIfNewUnit_list_of_contains_eq_false contains_eq_false + +theorem getKeyD_insertManyIfNewUnit_list_of_mem_of_contains_eq_false [EquivBEq α] [LawfulHashable α] + {l : List α} {k k' fallback : α} (k_beq : k == k') + (distinct : l.Pairwise (fun a b => (a == b) = false)) + (mem : k ∈ l ) (contains_eq_false : m.contains k = false) : + getKeyD (insertManyIfNewUnit m l) k' fallback = k := + DHashMap.Const.getKeyD_insertManyIfNewUnit_list_of_mem_of_contains_eq_false k_beq distinct mem contains_eq_false + +theorem getKeyD_insertManyIfNewUnit_list_of_contains_of_contains [EquivBEq α] [LawfulHashable α] + {l : List α} {k fallback : α} + (distinct : l.Pairwise (fun a b => (a == b) = false)) + (h : l.contains k) : + m.contains k → getKeyD (insertManyIfNewUnit m l) k fallback = getKeyD m k fallback := + DHashMap.Const.getKeyD_insertManyIfNewUnit_list_of_contains_of_contains distinct h + +theorem size_insertManyIfNewUnit_list [EquivBEq α] [LawfulHashable α] + {l : List α} + (distinct : l.Pairwise (fun a b => (a == b) = false)) : + (∀ (a : α), m.contains a → l.contains a = false) → + (insertManyIfNewUnit m l).size = m.size + l.length := + DHashMap.Const.size_insertManyIfNewUnit_list distinct + +theorem size_insertManyIfNewUnit_list_le [EquivBEq α] [LawfulHashable α] + {l : List α} : + (insertManyIfNewUnit m l).size ≤ m.size + l.length := + DHashMap.Const.size_insertManyIfNewUnit_list_le + +@[simp] +theorem isEmpty_insertManyIfNewUnit_list [EquivBEq α] [LawfulHashable α] + {l : List α} : + (insertManyIfNewUnit m l).isEmpty = (m.isEmpty && l.isEmpty) := + DHashMap.Const.isEmpty_insertManyIfNewUnit_list + +theorem get?_insertManyIfNewUnit_list [EquivBEq α] [LawfulHashable α] + {l : List α} {k : α} : + get? (insertManyIfNewUnit m l) k = + if m.contains k ∨ l.contains k then some () else none := + DHashMap.Const.get?_insertManyIfNewUnit_list + +theorem get_insertManyIfNewUnit_list + {l : List α} {k : α} {h} : + get (insertManyIfNewUnit m l) k h = () := + DHashMap.Const.get_insertManyIfNewUnit_list + +theorem get!_insertManyIfNewUnit_list + {l : List α} {k : α} : + get! (insertManyIfNewUnit m l) k = () := + DHashMap.Const.get!_insertManyIfNewUnit_list + +theorem getD_insertManyIfNewUnit_list + {l : List α} {k : α} {fallback : Unit} : + getD (insertManyIfNewUnit m l) k fallback = () := by + simp + +end + +section + +@[simp] +theorem ofList_nil [EquivBEq α] [LawfulHashable α] : + ofList ([] : List (α × β)) = ∅ := + ext DHashMap.Const.ofList_nil + +@[simp] +theorem ofList_singleton {k : α} {v : β} [EquivBEq α] [LawfulHashable α] : + ofList [⟨k, v⟩] = (∅ : HashMap α β).insert k v := + ext DHashMap.Const.ofList_singleton + +theorem ofList_cons [EquivBEq α] [LawfulHashable α] {k : α} {v : β} {tl : List (α × β)} : + ofList (⟨k, v⟩ :: tl) = insertMany ((∅ : HashMap α β).insert k v) tl := + ext DHashMap.Const.ofList_cons + +@[simp] +theorem contains_ofList [EquivBEq α] [LawfulHashable α] + {l : List (α × β)} {k : α} : + (ofList l).contains k = (l.map Prod.fst).contains k := + DHashMap.Const.contains_ofList + +@[simp] +theorem mem_ofList [EquivBEq α] [LawfulHashable α] + {l : List (α × β)} {k : α} : + k ∈ ofList l ↔ (l.map Prod.fst).contains k := + DHashMap.Const.mem_ofList + +theorem get?_ofList_of_contains_eq_false [LawfulBEq α] + {l : List (α × β)} {k : α} + (contains_eq_false : (l.map Prod.fst).contains k = false) : + get? (ofList l) k = none := + DHashMap.Const.get?_ofList_of_contains_eq_false contains_eq_false + +theorem get?_ofList_of_mem [LawfulBEq α] + {l : List (α × β)} {k k' : α} (k_beq : k == k') {v : β} + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) + (mem : ⟨k, v⟩ ∈ l) : + get? (ofList l) k' = some v := + DHashMap.Const.get?_ofList_of_mem k_beq distinct mem + +theorem get_ofList_of_mem [LawfulBEq α] + {l : List (α × β)} {k k' : α} (k_beq : k == k') {v : β} + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) + (mem : ⟨k, v⟩ ∈ l) + {h} : + get (ofList l) k' h = v := + DHashMap.Const.get_ofList_of_mem k_beq distinct mem + +theorem get!_ofList_of_contains_eq_false [LawfulBEq α] + {l : List (α × β)} {k : α} [Inhabited β] + (contains_eq_false : (l.map Prod.fst).contains k = false) : + get! (ofList l) k = (default : β) := + DHashMap.Const.get!_ofList_of_contains_eq_false contains_eq_false + +theorem get!_ofList_of_mem [LawfulBEq α] + {l : List (α × β)} {k k' : α} (k_beq : k == k') {v : β} [Inhabited β] + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) + (mem : ⟨k, v⟩ ∈ l) : + get! (ofList l) k' = v := + DHashMap.Const.get!_ofList_of_mem k_beq distinct mem + +theorem getD_ofList_of_contains_eq_false [LawfulBEq α] + {l : List (α × β)} {k : α} {fallback : β} + (contains_eq_false : (l.map Prod.fst).contains k = false) : + getD (ofList l) k fallback = fallback := + DHashMap.Const.getD_ofList_of_contains_eq_false contains_eq_false + +theorem getD_ofList_of_mem [LawfulBEq α] + {l : List (α × β)} {k k' : α} (k_beq : k == k') {v : β} {fallback : β} + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) + (mem : ⟨k, v⟩ ∈ l) : + getD (ofList l) k' fallback = v := + DHashMap.Const.getD_ofList_of_mem k_beq distinct mem + +theorem getKey?_ofList_of_contains_eq_false [EquivBEq α] [LawfulHashable α] + {l : List (α × β)} {k : α} + (contains_eq_false : (l.map Prod.fst).contains k = false) : + (ofList l).getKey? k = none := + DHashMap.Const.getKey?_ofList_of_contains_eq_false contains_eq_false + +theorem getKey?_ofList_of_mem [EquivBEq α] [LawfulHashable α] + {l : List (α × β)} + {k k' : α} (k_beq : k == k') + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) + (mem : k ∈ l.map Prod.fst) : + (ofList l).getKey? k' = some k := + DHashMap.Const.getKey?_ofList_of_mem k_beq distinct mem + +theorem getKey_ofList_of_mem [EquivBEq α] [LawfulHashable α] + {l : List (α × β)} + {k k' : α} (k_beq : k == k') + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) + (mem : k ∈ l.map Prod.fst) + {h} : + (ofList l).getKey k' h = k := + DHashMap.Const.getKey_ofList_of_mem k_beq distinct mem + +theorem getKey!_ofList_of_contains_eq_false [EquivBEq α] [LawfulHashable α] + [Inhabited α] {l : List (α × β)} {k : α} + (contains_eq_false : (l.map Prod.fst).contains k = false) : + (ofList l).getKey! k = default := + DHashMap.Const.getKey!_ofList_of_contains_eq_false contains_eq_false + +theorem getKey!_ofList_of_mem [EquivBEq α] [LawfulHashable α] [Inhabited α] + {l : List (α × β)} + {k k' : α} (k_beq : k == k') + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) + (mem : k ∈ l.map Prod.fst) : + (ofList l).getKey! k' = k := + DHashMap.Const.getKey!_ofList_of_mem k_beq distinct mem + +theorem getKeyD_ofList_of_contains_eq_false [EquivBEq α] [LawfulHashable α] + {l : List (α × β)} {k fallback : α} + (contains_eq_false : (l.map Prod.fst).contains k = false) : + (ofList l).getKeyD k fallback = fallback := + DHashMap.Const.getKeyD_ofList_of_contains_eq_false contains_eq_false + +theorem getKeyD_ofList_of_mem [EquivBEq α] [LawfulHashable α] + {l : List (α × β)} + {k k' fallback : α} (k_beq : k == k') + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) + (mem : k ∈ l.map Prod.fst) : + (ofList l).getKeyD k' fallback = k := + DHashMap.Const.getKeyD_ofList_of_mem k_beq distinct mem + +theorem size_ofList [EquivBEq α] [LawfulHashable α] + {l : List (α × β)} (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) : + (ofList l).size = l.length := + DHashMap.Const.size_ofList distinct + +theorem size_ofList_le [EquivBEq α] [LawfulHashable α] + {l : List (α × β)} : + (ofList l).size ≤ l.length := + DHashMap.Const.size_ofList_le + +@[simp] +theorem isEmpty_ofList [EquivBEq α] [LawfulHashable α] + {l : List (α × β)} : + (ofList l).isEmpty = l.isEmpty := + DHashMap.Const.isEmpty_ofList + +@[simp] +theorem unitOfList_nil [EquivBEq α] [LawfulHashable α] : + unitOfList ([] : List α) = ∅ := + ext DHashMap.Const.unitOfList_nil + +@[simp] +theorem unitOfList_singleton [EquivBEq α] [LawfulHashable α] {k : α} : + unitOfList [k] = (∅ : HashMap α Unit).insertIfNew k () := + ext DHashMap.Const.unitOfList_singleton + +theorem unitOfList_cons [EquivBEq α] [LawfulHashable α] {hd : α} {tl : List α} : + unitOfList (hd :: tl) = + insertManyIfNewUnit ((∅ : HashMap α Unit).insertIfNew hd ()) tl := + ext DHashMap.Const.unitOfList_cons + +@[simp] +theorem contains_unitOfList [EquivBEq α] [LawfulHashable α] + {l : List α} {k : α} : + (unitOfList l).contains k = l.contains k := + DHashMap.Const.contains_unitOfList + +@[simp] +theorem mem_unitOfList [EquivBEq α] [LawfulHashable α] + {l : List α} {k : α} : + k ∈ unitOfList l ↔ l.contains k := + DHashMap.Const.mem_unitOfList + +theorem getKey?_unitOfList_of_contains_eq_false [EquivBEq α] [LawfulHashable α] + {l : List α} {k : α} (contains_eq_false : l.contains k = false) : + getKey? (unitOfList l) k = none := + DHashMap.Const.getKey?_unitOfList_of_contains_eq_false contains_eq_false + +theorem getKey?_unitOfList_of_mem [EquivBEq α] [LawfulHashable α] + {l : List α} {k k' : α} (k_beq : k == k') + (distinct : l.Pairwise (fun a b => (a == b) = false)) (mem : k ∈ l) : + getKey? (unitOfList l) k' = some k := + DHashMap.Const.getKey?_unitOfList_of_mem k_beq distinct mem + +theorem getKey_unitOfList_of_mem [EquivBEq α] [LawfulHashable α] + {l : List α} + {k k' : α} (k_beq : k == k') + (distinct : l.Pairwise (fun a b => (a == b) = false)) + (mem : k ∈ l) {h} : + getKey (unitOfList l) k' h = k := + DHashMap.Const.getKey_unitOfList_of_mem k_beq distinct mem + +theorem getKey!_unitOfList_of_contains_eq_false [EquivBEq α] [LawfulHashable α] + [Inhabited α] {l : List α} {k : α} + (contains_eq_false : l.contains k = false) : + getKey! (unitOfList l) k = default := + DHashMap.Const.getKey!_unitOfList_of_contains_eq_false contains_eq_false + +theorem getKey!_unitOfList_of_mem [EquivBEq α] [LawfulHashable α] + [Inhabited α] {l : List α} {k k' : α} (k_beq : k == k') + (distinct : l.Pairwise (fun a b => (a == b) = false)) + (mem : k ∈ l) : + getKey! (unitOfList l) k' = k := + DHashMap.Const.getKey!_unitOfList_of_mem k_beq distinct mem + +theorem getKeyD_unitOfList_of_contains_eq_false [EquivBEq α] [LawfulHashable α] + {l : List α} {k fallback : α} + (contains_eq_false : l.contains k = false) : + getKeyD (unitOfList l) k fallback = fallback := + DHashMap.Const.getKeyD_unitOfList_of_contains_eq_false contains_eq_false + +theorem getKeyD_unitOfList_of_mem [EquivBEq α] [LawfulHashable α] + {l : List α} {k k' fallback : α} (k_beq : k == k') + (distinct : l.Pairwise (fun a b => (a == b) = false)) + (mem : k ∈ l) : + getKeyD (unitOfList l) k' fallback = k := + DHashMap.Const.getKeyD_unitOfList_of_mem k_beq distinct mem + +theorem size_unitOfList [EquivBEq α] [LawfulHashable α] + {l : List α} + (distinct : l.Pairwise (fun a b => (a == b) = false)) : + (unitOfList l).size = l.length := + DHashMap.Const.size_unitOfList distinct + +theorem size_unitOfList_le [EquivBEq α] [LawfulHashable α] + {l : List α} : + (unitOfList l).size ≤ l.length := + DHashMap.Const.size_unitOfList_le + +@[simp] +theorem isEmpty_unitOfList [EquivBEq α] [LawfulHashable α] + {l : List α} : + (unitOfList l).isEmpty = l.isEmpty := + DHashMap.Const.isEmpty_unitOfList + +@[simp] +theorem get?_unitOfList [EquivBEq α] [LawfulHashable α] + {l : List α} {k : α} : + get? (unitOfList l) k = + if l.contains k then some () else none := + DHashMap.Const.get?_unitOfList + +@[simp] +theorem get_unitOfList + {l : List α} {k : α} {h} : + get (unitOfList l) k h = () := + DHashMap.Const.get_unitOfList + +@[simp] +theorem get!_unitOfList + {l : List α} {k : α} : + get! (unitOfList l) k = () := + DHashMap.Const.get!_unitOfList + +@[simp] +theorem getD_unitOfList + {l : List α} {k : α} {fallback : Unit} : + getD (unitOfList l) k fallback = () := by + simp + end end Std.HashMap diff --git a/src/Std/Data/HashMap/RawLemmas.lean b/src/Std/Data/HashMap/RawLemmas.lean index cc0816a966c7..051aa22143f5 100644 --- a/src/Std/Data/HashMap/RawLemmas.lean +++ b/src/Std/Data/HashMap/RawLemmas.lean @@ -671,7 +671,6 @@ theorem getKeyD_insertIfNew [EquivBEq α] [LawfulHashable α] (h : m.WF) {k a fa (m.insertIfNew k v).getKeyD a fallback = if k == a ∧ ¬k ∈ m then k else m.getKeyD a fallback := DHashMap.Raw.getKeyD_insertIfNew h.out - @[simp] theorem getThenInsertIfNew?_fst (h : m.WF) {k : α} {v : β} : (getThenInsertIfNew? m k v).1 = get? m k := @@ -699,13 +698,592 @@ theorem contains_keys [EquivBEq α] [LawfulHashable α] (h : m.WF) {k : α} : @[simp] theorem mem_keys [LawfulBEq α] [LawfulHashable α] (h : m.WF) {k : α} : - k ∈ m.keys ↔ k ∈ m := + k ∈ m.keys ↔ k ∈ m := DHashMap.Raw.mem_keys h.out theorem distinct_keys [EquivBEq α] [LawfulHashable α] (h : m.WF) : - m.keys.Pairwise (fun a b => (a == b) = false) := + m.keys.Pairwise (fun a b => (a == b) = false) := DHashMap.Raw.distinct_keys h.out +@[simp] +theorem insertMany_nil [EquivBEq α] [LawfulHashable α] (h : m.WF) : + insertMany m [] = m := + ext (DHashMap.Raw.Const.insertMany_nil h.out) + +@[simp] +theorem insertMany_list_singleton [EquivBEq α] [LawfulHashable α] (h : m.WF) + {k : α} {v : β} : + insertMany m [⟨k, v⟩] = m.insert k v := + ext (DHashMap.Raw.Const.insertMany_list_singleton h.out) + +theorem insertMany_cons [EquivBEq α] [LawfulHashable α] (h : m.WF) {l : List (α × β)} + {k : α} {v : β} : + insertMany m (⟨k, v⟩ :: l) = insertMany (m.insert k v) l := + ext (DHashMap.Raw.Const.insertMany_cons h.out) + +@[simp] +theorem contains_insertMany_list [EquivBEq α] [LawfulHashable α] (h : m.WF) + {l : List (α × β)} {k : α} : + (insertMany m l).contains k = (m.contains k || (l.map Prod.fst).contains k) := + DHashMap.Raw.Const.contains_insertMany_list h.out + +@[simp] +theorem mem_insertMany_list [EquivBEq α] [LawfulHashable α] (h : m.WF) + {l : List (α × β)} {k : α} : + k ∈ (insertMany m l) ↔ k ∈ m ∨ (l.map Prod.fst).contains k := + DHashMap.Raw.Const.mem_insertMany_list h.out + +theorem contains_of_contains_insertMany_list [EquivBEq α] [LawfulHashable α] (h : m.WF) + {l : List ( α × β )} {k : α} : + (insertMany m l).contains k → (l.map Prod.fst).contains k = false → m.contains k := + DHashMap.Raw.Const.contains_of_contains_insertMany_list h.out + +theorem getKey?_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] + (h : m.WF) {l : List (α × β)} {k : α} + (contains_eq_false : (l.map Prod.fst).contains k = false) : + (insertMany m l).getKey? k = m.getKey? k := + DHashMap.Raw.Const.getKey?_insertMany_list_of_contains_eq_false h.out contains_eq_false + +theorem getKey?_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] + (h : m.WF) {l : List (α × β)} + {k k' : α} (k_beq : k == k') + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) + (mem : k ∈ l.map Prod.fst) : + (insertMany m l).getKey? k' = some k := + DHashMap.Raw.Const.getKey?_insertMany_list_of_mem h.out k_beq distinct mem + +theorem getKey_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] + (h : m.WF) {l : List (α × β)} {k : α} + (contains_eq_false : (l.map Prod.fst).contains k = false) + {h'} : + (insertMany m l).getKey k h' = + m.getKey k (contains_of_contains_insertMany_list h h' contains_eq_false) := + DHashMap.Raw.Const.getKey_insertMany_list_of_contains_eq_false h.out contains_eq_false + +theorem getKey_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] + (h : m.WF) {l : List (α × β)} + {k k' : α} (k_beq : k == k') + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) + (mem : k ∈ l.map Prod.fst) + {h'} : + (insertMany m l).getKey k' h' = k := + DHashMap.Raw.Const.getKey_insertMany_list_of_mem h.out k_beq distinct mem + +theorem getKey!_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] [Inhabited α] + (h : m.WF) {l : List (α × β)} {k : α} + (contains_eq_false : (l.map Prod.fst).contains k = false) : + (insertMany m l).getKey! k = m.getKey! k := + DHashMap.Raw.Const.getKey!_insertMany_list_of_contains_eq_false h.out contains_eq_false + +theorem getKey!_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] [Inhabited α] + (h : m.WF) {l : List (α × β)} + {k k' : α} (k_beq : k == k') + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) + (mem : k ∈ l.map Prod.fst) : + (insertMany m l).getKey! k' = k := + DHashMap.Raw.Const.getKey!_insertMany_list_of_mem h.out k_beq distinct mem + +theorem getKeyD_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] + (h : m.WF) {l : List (α × β)} {k fallback : α} + (contains_eq_false : (l.map Prod.fst).contains k = false) : + (insertMany m l).getKeyD k fallback = m.getKeyD k fallback := + DHashMap.Raw.Const.getKeyD_insertMany_list_of_contains_eq_false h.out contains_eq_false + +theorem getKeyD_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] + (h : m.WF) {l : List (α × β)} + {k k' fallback : α} (k_beq : k == k') + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) + (mem : k ∈ l.map Prod.fst) : + (insertMany m l).getKeyD k' fallback = k := + DHashMap.Raw.Const.getKeyD_insertMany_list_of_mem h.out k_beq distinct mem + +theorem size_insertMany_list [EquivBEq α] [LawfulHashable α] + (h : m.WF) {l : List (α × β)} + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) : + (∀ (a : α), m.contains a → (l.map Prod.fst).contains a = false) → + (insertMany m l).size = m.size + l.length := + DHashMap.Raw.Const.size_insertMany_list h.out distinct + +theorem size_insertMany_list_le [EquivBEq α] [LawfulHashable α] + (h : m.WF) {l : List (α × β)} : + (insertMany m l).size ≤ m.size + l.length := + DHashMap.Raw.Const.size_insertMany_list_le h.out + +@[simp] +theorem isEmpty_insertMany_list [EquivBEq α] [LawfulHashable α] + (h : m.WF) {l : List (α × β)} : + (insertMany m l).isEmpty = (m.isEmpty && l.isEmpty) := + DHashMap.Raw.Const.isEmpty_insertMany_list h.out + +theorem get?_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] + (h : m.WF) {l : List (α × β)} {k : α} + (contains_eq_false : (l.map Prod.fst).contains k = false) : + (insertMany m l)[k]? = m[k]? := + DHashMap.Raw.Const.get?_insertMany_list_of_contains_eq_false h.out contains_eq_false + +theorem get?_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] + (h : m.WF) {l : List (α × β)} {k k' : α} (k_beq : k == k') {v : β} + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) (mem : ⟨k, v⟩ ∈ l) : + (insertMany m l)[k']? = v := + DHashMap.Raw.Const.get?_insertMany_list_of_mem h.out k_beq distinct mem + +theorem get_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] + (h : m.WF) {l : List (α × β)} {k : α} + (contains_eq_false : (l.map Prod.fst).contains k = false) + {h'} : + (insertMany m l)[k] = + m[k]'(contains_of_contains_insertMany_list h h' contains_eq_false) := + DHashMap.Raw.Const.get_insertMany_list_of_contains_eq_false h.out contains_eq_false (h':= h') + +theorem get_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] + (h : m.WF) {l : List (α × β)} {k k' : α} (k_beq : k == k') {v : β} + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) (mem : ⟨k, v⟩ ∈ l) {h'} : + (insertMany m l)[k'] = v := + DHashMap.Raw.Const.get_insertMany_list_of_mem h.out k_beq distinct mem (h' := h') + +theorem get!_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] + [Inhabited β] (h : m.WF) {l : List (α × β)} {k : α} + (contains_eq_false : (l.map Prod.fst).contains k = false) : + (insertMany m l)[k]! = m[k]! := + DHashMap.Raw.Const.get!_insertMany_list_of_contains_eq_false h.out contains_eq_false + +theorem get!_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] [Inhabited β] + (h : m.WF) {l : List (α × β)} {k k' : α} (k_beq : k == k') {v : β} + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) (mem : ⟨k, v⟩ ∈ l) : + (insertMany m l)[k']! = v := + DHashMap.Raw.Const.get!_insertMany_list_of_mem h.out k_beq distinct mem + +theorem getD_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] + (h : m.WF) {l : List (α × β)} {k : α} {fallback : β} + (contains_eq_false : (l.map Prod.fst).contains k = false) : + getD (insertMany m l) k fallback = getD m k fallback := + DHashMap.Raw.Const.getD_insertMany_list_of_contains_eq_false h.out contains_eq_false + +theorem getD_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] + (h : m.WF) {l : List (α × β)} {k k' : α} (k_beq : k == k') {v fallback : β} + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) (mem : ⟨k, v⟩ ∈ l) : + getD (insertMany m l) k' fallback = v := + DHashMap.Raw.Const.getD_insertMany_list_of_mem h.out k_beq distinct mem + +variable {m : Raw α Unit} + +@[simp] +theorem insertManyIfNewUnit_nil [EquivBEq α] [LawfulHashable α] (h : m.WF) : + insertManyIfNewUnit m [] = m := + ext (DHashMap.Raw.Const.insertManyIfNewUnit_nil h.out) + +@[simp] +theorem insertManyIfNewUnit_list_singleton [EquivBEq α] [LawfulHashable α] (h : m.WF) {k : α} : + insertManyIfNewUnit m [k] = m.insertIfNew k () := + ext (DHashMap.Raw.Const.insertManyIfNewUnit_list_singleton h.out) + +theorem insertManyIfNewUnit_cons [EquivBEq α] [LawfulHashable α] (h : m.WF) {l : List α} {k : α} : + (insertManyIfNewUnit m (k :: l)) = (insertManyIfNewUnit (m.insertIfNew k ()) l) := + ext (DHashMap.Raw.Const.insertManyIfNewUnit_cons h.out) + +@[simp] +theorem contains_insertManyIfNewUnit_list [EquivBEq α] [LawfulHashable α] (h : m.WF) + {l : List α} {k : α} : + (insertManyIfNewUnit m l).contains k = (m.contains k || l.contains k) := + DHashMap.Raw.Const.contains_insertManyIfNewUnit_list h.out + +@[simp] +theorem mem_insertManyIfNewUnit_list [EquivBEq α] [LawfulHashable α] (h : m.WF) + {l : List α} {k : α} : + k ∈ insertManyIfNewUnit m l ↔ k ∈ m ∨ l.contains k := + DHashMap.Raw.Const.mem_insertManyIfNewUnit_list h.out + +theorem contains_of_contains_insertManyIfNewUnit_list [EquivBEq α] [LawfulHashable α] (h : m.WF) + {l : List α} {k : α} (contains_eq_false : l.contains k = false) : + (insertManyIfNewUnit m l).contains k → m.contains k := + DHashMap.Raw.Const.contains_of_contains_insertManyIfNewUnit_list h.out contains_eq_false + +theorem getKey?_insertManyIfNewUnit_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] + (h : m.WF) {l : List α} {k : α} (contains_eq_false : l.contains k = false) : + getKey? (insertManyIfNewUnit m l) k = getKey? m k := + DHashMap.Raw.Const.getKey?_insertManyIfNewUnit_list_of_contains_eq_false h.out contains_eq_false + +theorem getKey?_insertManyIfNewUnit_list_of_mem_of_contains_eq_false [EquivBEq α] [LawfulHashable α] + (h : m.WF) {l : List α} {k k' : α} (k_beq : k == k') + (distinct : l.Pairwise (fun a b => (a == b) = false)) (mem : k ∈ l) : + m.contains k = false → getKey? (insertManyIfNewUnit m l) k' = some k := + DHashMap.Raw.Const.getKey?_insertManyIfNewUnit_list_of_mem_of_contains_eq_false + h.out k_beq distinct mem + +theorem getKey?_insertManyIfNewUnit_list_of_contains_of_contains [EquivBEq α] [LawfulHashable α] + (h : m.WF) {l : List α} {k : α} + (distinct : l.Pairwise (fun a b => (a == b) = false)) + (h' : l.contains k) : + m.contains k → getKey? (insertManyIfNewUnit m l) k = getKey? m k := + DHashMap.Raw.Const.getKey?_insertManyIfNewUnit_list_of_contains_of_contains h.out distinct h' + +theorem getKey_insertManyIfNewUnit_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] + (h : m.WF) {l : List α} {k : α} (contains_eq_false : l.contains k = false) {h'} : + getKey (insertManyIfNewUnit m l) k h' = + getKey m k (contains_of_contains_insertManyIfNewUnit_list h contains_eq_false h') := + DHashMap.Raw.Const.getKey_insertManyIfNewUnit_list_of_contains_eq_false h.out contains_eq_false + +theorem getKey_insertManyIfNewUnit_list_of_mem_of_contains_eq_false [EquivBEq α] [LawfulHashable α] + (h : m.WF) {l : List α} + {k k' : α} (k_beq : k == k') + (distinct : l.Pairwise (fun a b => (a == b) = false)) + (mem : k ∈ l) {h'} : + m.contains k = false → getKey (insertManyIfNewUnit m l) k' h' = k := + DHashMap.Raw.Const.getKey_insertManyIfNewUnit_list_of_mem_of_contains_eq_false + h.out k_beq distinct mem + +theorem getKey_insertManyIfNewUnit_list_mem_of_contains_of_contains [EquivBEq α] [LawfulHashable α] + (h : m.WF) {l : List α} {k : α} + (distinct : l.Pairwise (fun a b => (a == b) = false)) + (h₁ : l.contains k) (h₂ : m.contains k) {h₃} : + getKey (insertManyIfNewUnit m l) k h₃ = getKey m k h₂ := + DHashMap.Raw.Const.getKey_insertManyIfNewUnit_list_mem_of_contains_of_contains + h.out distinct h₁ h₂ + +theorem getKey!_insertManyIfNewUnit_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] + [Inhabited α] (h : m.WF) {l : List α} {k : α} + (contains_eq_false : l.contains k = false) : + getKey! (insertManyIfNewUnit m l) k = getKey! m k := + DHashMap.Raw.Const.getKey!_insertManyIfNewUnit_list_of_contains_eq_false h.out contains_eq_false + +theorem getKey!_insertManyIfNewUnit_list_of_mem_of_contains_eq_false [EquivBEq α] [LawfulHashable α] + [Inhabited α] (h : m.WF) {l : List α} {k k' : α} (k_beq : k == k') + (distinct : l.Pairwise (fun a b => (a == b) = false)) + (mem : k ∈ l) : + contains m k = false → getKey! (insertManyIfNewUnit m l) k' = k := + DHashMap.Raw.Const.getKey!_insertManyIfNewUnit_list_of_mem_of_contains_eq_false + h.out k_beq distinct mem + +theorem getKey!_insertManyIfNewUnit_list_mem_of_contains_of_contains [EquivBEq α] [LawfulHashable α] + [Inhabited α] (h : m.WF) {l : List α} {k : α} + (distinct : l.Pairwise (fun a b => (a == b) = false)) + (h' : l.contains k) : + m.contains k → getKey! (insertManyIfNewUnit m l) k = getKey! m k := + DHashMap.Raw.Const.getKey!_insertManyIfNewUnit_list_mem_of_contains_of_contains h.out distinct h' + +theorem getKeyD_insertManyIfNewUnit_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] + (h : m.WF) {l : List α} {k fallback : α} + (contains_eq_false : l.contains k = false) : + getKeyD (insertManyIfNewUnit m l) k fallback = getKeyD m k fallback := + DHashMap.Raw.Const.getKeyD_insertManyIfNewUnit_list_of_contains_eq_false h.out contains_eq_false + +theorem getKeyD_insertManyIfNewUnit_list_of_mem_of_contains_eq_false [EquivBEq α] [LawfulHashable α] + (h : m.WF) {l : List α} {k k' fallback : α} (k_beq : k == k') + (distinct : l.Pairwise (fun a b => (a == b) = false)) + (mem : k ∈ l) : + m.contains k = false → getKeyD (insertManyIfNewUnit m l) k' fallback = k := + DHashMap.Raw.Const.getKeyD_insertManyIfNewUnit_list_of_mem_of_contains_eq_false + h.out k_beq distinct mem + +theorem getKeyD_insertManyIfNewUnit_list_of_contains_of_contains [EquivBEq α] [LawfulHashable α] + (h : m.WF) {l : List α} {k fallback : α} + (distinct : l.Pairwise (fun a b => (a == b) = false)) + (h' : l.contains k) : + m.contains k → getKeyD (insertManyIfNewUnit m l) k fallback = getKeyD m k fallback := + DHashMap.Raw.Const.getKeyD_insertManyIfNewUnit_list_of_contains_of_contains h.out distinct h' + +theorem size_insertManyIfNewUnit_list [EquivBEq α] [LawfulHashable α] (h : m.WF) + {l : List α} + (distinct : l.Pairwise (fun a b => (a == b) = false)) : + (∀ (a : α), m.contains a → l.contains a = false) → + (insertManyIfNewUnit m l).size = m.size + l.length := + DHashMap.Raw.Const.size_insertManyIfNewUnit_list h.out distinct + +theorem size_insertManyIfNewUnit_list_le [EquivBEq α] [LawfulHashable α] (h : m.WF) + {l : List α} : + (insertManyIfNewUnit m l).size ≤ m.size + l.length := + DHashMap.Raw.Const.size_insertManyIfNewUnit_list_le h.out + +@[simp] +theorem isEmpty_insertManyIfNewUnit_list [EquivBEq α] [LawfulHashable α] (h : m.WF) + {l : List α} : + (insertManyIfNewUnit m l).isEmpty = (m.isEmpty && l.isEmpty) := + DHashMap.Raw.Const.isEmpty_insertManyIfNewUnit_list h.out + +@[simp] +theorem get?_insertManyIfNewUnit_list [EquivBEq α] [LawfulHashable α] (h : m.WF) + {l : List α} {k : α} : + (insertManyIfNewUnit m l)[k]? = + if m.contains k ∨ l.contains k then some () else none := + DHashMap.Raw.Const.get?_insertManyIfNewUnit_list h.out + +@[simp] +theorem get_insertManyIfNewUnit_list + {l : List α} {k : α} {h} : + (insertManyIfNewUnit m l)[k] = () := + DHashMap.Raw.Const.get_insertManyIfNewUnit_list (h:=h) + +@[simp] +theorem get!_insertManyIfNewUnit_list + {l : List α} {k : α} : + (insertManyIfNewUnit m l)[k]! = () := + DHashMap.Raw.Const.get!_insertManyIfNewUnit_list + +@[simp] +theorem getD_insertManyIfNewUnit_list + {l : List α} {k : α} {fallback : Unit} : + getD (insertManyIfNewUnit m l) k fallback = () := by + simp + +end Raw + +namespace Raw + +variable [BEq α] [Hashable α] + +@[simp] +theorem ofList_nil [EquivBEq α] [LawfulHashable α] : + ofList ([] : List (α × β)) = ∅ := + ext DHashMap.Raw.Const.ofList_nil + +@[simp] +theorem ofList_singleton [EquivBEq α] [LawfulHashable α] {k : α} {v : β} : + ofList [⟨k, v⟩] = (∅ : Raw α β).insert k v := + ext DHashMap.Raw.Const.ofList_singleton + +theorem ofList_cons [EquivBEq α] [LawfulHashable α] {k : α} {v : β} {tl : List (α × β)} : + ofList (⟨k, v⟩ :: tl) = insertMany ((∅ : Raw α β).insert k v) tl := + ext DHashMap.Raw.Const.ofList_cons + +@[simp] +theorem contains_ofList [EquivBEq α] [LawfulHashable α] + {l : List (α × β)} {k : α} : + (ofList l).contains k = (l.map Prod.fst).contains k := + DHashMap.Raw.Const.contains_ofList + +@[simp] +theorem mem_ofList [EquivBEq α] [LawfulHashable α] + {l : List (α × β)} {k : α} : + k ∈ (ofList l) ↔ (l.map Prod.fst).contains k := + DHashMap.Raw.Const.mem_ofList + +theorem get?_ofList_of_contains_eq_false [LawfulBEq α] + {l : List (α × β)} {k : α} + (contains_eq_false : (l.map Prod.fst).contains k = false) : + (ofList l)[k]? = none := + DHashMap.Raw.Const.get?_ofList_of_contains_eq_false contains_eq_false + +theorem get?_ofList_of_mem [LawfulBEq α] + {l : List (α × β)} {k k' : α} (k_beq : k == k') {v : β} + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) + (mem : ⟨k, v⟩ ∈ l) : + (ofList l)[k']? = some v := + DHashMap.Raw.Const.get?_ofList_of_mem k_beq distinct mem + +theorem get_ofList_of_mem [LawfulBEq α] + {l : List (α × β)} {k k' : α} (k_beq : k == k') {v : β} + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) + (mem : ⟨k, v⟩ ∈ l) + {h} : + (ofList l)[k'] = v := + DHashMap.Raw.Const.get_ofList_of_mem k_beq distinct mem (h:=h) + +theorem get!_ofList_of_contains_eq_false [LawfulBEq α] + {l : List (α × β)} {k : α} [Inhabited β] + (contains_eq_false : (l.map Prod.fst).contains k = false) : + (ofList l)[k]! = default := + DHashMap.Raw.Const.get!_ofList_of_contains_eq_false contains_eq_false + +theorem get!_ofList_of_mem [LawfulBEq α] + {l : List (α × β)} {k k' : α} (k_beq : k == k') {v : β} [Inhabited β] + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) + (mem : ⟨k, v⟩ ∈ l) : + (ofList l)[k']! = v := + DHashMap.Raw.Const.get!_ofList_of_mem k_beq distinct mem + +theorem getD_ofList_of_contains_eq_false [LawfulBEq α] + {l : List (α × β)} {k : α} {fallback : β} + (contains_eq_false : (l.map Prod.fst).contains k = false) : + getD (ofList l) k fallback = fallback := + DHashMap.Raw.Const.getD_ofList_of_contains_eq_false contains_eq_false + +theorem getD_ofList_of_mem [LawfulBEq α] + {l : List (α × β)} {k k' : α} (k_beq : k == k') {v : β} {fallback : β} + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) + (mem : ⟨k, v⟩ ∈ l) : + getD (ofList l) k' fallback = v := + DHashMap.Raw.Const.getD_ofList_of_mem k_beq distinct mem + +theorem getKey?_ofList_of_contains_eq_false [EquivBEq α] [LawfulHashable α] + {l : List (α × β)} {k : α} + (contains_eq_false : (l.map Prod.fst).contains k = false) : + (ofList l).getKey? k = none := + DHashMap.Raw.Const.getKey?_ofList_of_contains_eq_false contains_eq_false + +theorem getKey?_ofList_of_mem [EquivBEq α] [LawfulHashable α] + {l : List (α × β)} + {k k' : α} (k_beq : k == k') + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) + (mem : k ∈ l.map Prod.fst) : + (ofList l).getKey? k' = some k := + DHashMap.Raw.Const.getKey?_ofList_of_mem k_beq distinct mem + +theorem getKey_ofList_of_mem [EquivBEq α] [LawfulHashable α] + {l : List (α × β)} + {k k' : α} (k_beq : k == k') + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) + (mem : k ∈ l.map Prod.fst) + {h'} : + (ofList l).getKey k' h' = k := + DHashMap.Raw.Const.getKey_ofList_of_mem k_beq distinct mem + +theorem getKey!_ofList_of_contains_eq_false [EquivBEq α] [LawfulHashable α] [Inhabited α] + {l : List (α × β)} {k : α} + (contains_eq_false : (l.map Prod.fst).contains k = false) : + (ofList l).getKey! k = default := + DHashMap.Raw.Const.getKey!_ofList_of_contains_eq_false contains_eq_false + +theorem getKey!_ofList_of_mem [EquivBEq α] [LawfulHashable α] [Inhabited α] + {l : List (α × β)} + {k k' : α} (k_beq : k == k') + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) + (mem : k ∈ l.map Prod.fst) : + (ofList l).getKey! k' = k := + DHashMap.Raw.Const.getKey!_ofList_of_mem k_beq distinct mem + +theorem getKeyD_ofList_of_contains_eq_false [EquivBEq α] [LawfulHashable α] + {l : List (α × β)} {k fallback : α} + (contains_eq_false : (l.map Prod.fst).contains k = false) : + (ofList l).getKeyD k fallback = fallback := + DHashMap.Raw.Const.getKeyD_ofList_of_contains_eq_false contains_eq_false + +theorem getKeyD_ofList_of_mem [EquivBEq α] [LawfulHashable α] + {l : List (α × β)} + {k k' fallback : α} (k_beq : k == k') + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) + (mem : k ∈ l.map Prod.fst) : + (ofList l).getKeyD k' fallback = k := + DHashMap.Raw.Const.getKeyD_ofList_of_mem k_beq distinct mem + +theorem size_ofList [EquivBEq α] [LawfulHashable α] + {l : List (α × β)} (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) : + (ofList l).size = l.length := + DHashMap.Raw.Const.size_ofList distinct + +theorem size_ofList_le [EquivBEq α] [LawfulHashable α] + {l : List (α × β)} : + (ofList l).size ≤ l.length := + DHashMap.Raw.Const.size_ofList_le + +@[simp] +theorem isEmpty_ofList [EquivBEq α] [LawfulHashable α] + {l : List (α × β)} : + (ofList l).isEmpty = l.isEmpty := + DHashMap.Raw.Const.isEmpty_ofList + +@[simp] +theorem unitOfList_nil [EquivBEq α] [LawfulHashable α] : + unitOfList ([] : List α) = ∅ := + ext DHashMap.Raw.Const.unitOfList_nil + +@[simp] +theorem unitOfList_singleton [EquivBEq α] [LawfulHashable α] {k : α} : + unitOfList [k] = (∅ : Raw α Unit).insertIfNew k () := + ext DHashMap.Raw.Const.unitOfList_singleton + +theorem unitOfList_cons [EquivBEq α] [LawfulHashable α] {hd : α} {tl : List α} : + unitOfList (hd :: tl) = insertManyIfNewUnit ((∅ : Raw α Unit).insertIfNew hd ()) tl := + ext DHashMap.Raw.Const.unitOfList_cons + +@[simp] +theorem contains_unitOfList [EquivBEq α] [LawfulHashable α] + {l : List α} {k : α} : + (unitOfList l).contains k = l.contains k := + DHashMap.Raw.Const.contains_unitOfList + +@[simp] +theorem mem_unitOfList [EquivBEq α] [LawfulHashable α] + {l : List α} {k : α} : + k ∈ unitOfList l ↔ l.contains k := + DHashMap.Raw.Const.mem_unitOfList + +theorem getKey?_unitOfList_of_contains_eq_false [EquivBEq α] [LawfulHashable α] + {l : List α} {k : α} (contains_eq_false : l.contains k = false) : + getKey? (unitOfList l) k = none := + DHashMap.Raw.Const.getKey?_unitOfList_of_contains_eq_false contains_eq_false + +theorem getKey?_unitOfList_of_mem [EquivBEq α] [LawfulHashable α] + {l : List α} {k k' : α} (k_beq : k == k') + (distinct : l.Pairwise (fun a b => (a == b) = false)) (mem : k ∈ l) : + getKey? (unitOfList l) k' = some k := + DHashMap.Raw.Const.getKey?_unitOfList_of_mem k_beq distinct mem + +theorem getKey_unitOfList_of_mem [EquivBEq α] [LawfulHashable α] + {l : List α} + {k k' : α} (k_beq : k == k') + (distinct : l.Pairwise (fun a b => (a == b) = false)) + (mem : k ∈ l) {h'} : + getKey (unitOfList l) k' h' = k := + DHashMap.Raw.Const.getKey_unitOfList_of_mem k_beq distinct mem + +theorem getKey!_unitOfList_of_contains_eq_false [EquivBEq α] [LawfulHashable α] + [Inhabited α] {l : List α} {k : α} + (contains_eq_false : l.contains k = false) : + getKey! (unitOfList l) k = default := + DHashMap.Raw.Const.getKey!_unitOfList_of_contains_eq_false contains_eq_false + +theorem getKey!_unitOfList_of_mem [EquivBEq α] [LawfulHashable α] + [Inhabited α] {l : List α} {k k' : α} (k_beq : k == k') + (distinct : l.Pairwise (fun a b => (a == b) = false)) + (mem : k ∈ l) : + getKey! (unitOfList l) k' = k := + DHashMap.Raw.Const.getKey!_unitOfList_of_mem k_beq distinct mem + +theorem getKeyD_unitOfList_of_contains_eq_false [EquivBEq α] [LawfulHashable α] + {l : List α} {k fallback : α} + (contains_eq_false : l.contains k = false) : + getKeyD (unitOfList l) k fallback = fallback := + DHashMap.Raw.Const.getKeyD_unitOfList_of_contains_eq_false contains_eq_false + +theorem getKeyD_unitOfList_of_mem [EquivBEq α] [LawfulHashable α] + {l : List α} {k k' fallback : α} (k_beq : k == k') + (distinct : l.Pairwise (fun a b => (a == b) = false)) + (mem : k ∈ l) : + getKeyD (unitOfList l) k' fallback = k := + DHashMap.Raw.Const.getKeyD_unitOfList_of_mem k_beq distinct mem + +theorem size_unitOfList [EquivBEq α] [LawfulHashable α] + {l : List α} + (distinct : l.Pairwise (fun a b => (a == b) = false)) : + (unitOfList l).size = l.length := + DHashMap.Raw.Const.size_unitOfList distinct + +theorem size_unitOfList_le [EquivBEq α] [LawfulHashable α] + {l : List α} : + (unitOfList l).size ≤ l.length := + DHashMap.Raw.Const.size_unitOfList_le + +@[simp] +theorem isEmpty_unitOfList [EquivBEq α] [LawfulHashable α] + {l : List α} : + (unitOfList l).isEmpty = l.isEmpty := + DHashMap.Raw.Const.isEmpty_unitOfList + +@[simp] +theorem get?_unitOfList [EquivBEq α] [LawfulHashable α] + {l : List α} {k : α} : + (unitOfList l)[k]? = + if l.contains k then some () else none := + DHashMap.Raw.Const.get?_unitOfList + +@[simp] +theorem get_unitOfList + {l : List α} {k : α} {h} : + (unitOfList l)[k] = () := + DHashMap.Raw.Const.get_unitOfList (h:=h) + +@[simp] +theorem get!_unitOfList + {l : List α} {k : α} : + (unitOfList l)[k]! = () := + DHashMap.Raw.Const.get!_unitOfList + +@[simp] +theorem getD_unitOfList + {l : List α} {k : α} {fallback : Unit} : + getD (unitOfList l) k fallback = () := by + simp + end Raw end Std.HashMap From 564ce90a782292addc29518dd7dd2e58db073780 Mon Sep 17 00:00:00 2001 From: jt0202 Date: Thu, 9 Jan 2025 14:37:08 +0100 Subject: [PATCH 71/83] Replace symbols for get --- src/Std/Data/HashMap/Lemmas.lean | 200 +++++++++++++++---------------- 1 file changed, 100 insertions(+), 100 deletions(-) diff --git a/src/Std/Data/HashMap/Lemmas.lean b/src/Std/Data/HashMap/Lemmas.lean index 8fa3164915c7..88b42ccda87e 100644 --- a/src/Std/Data/HashMap/Lemmas.lean +++ b/src/Std/Data/HashMap/Lemmas.lean @@ -730,6 +730,55 @@ theorem contains_of_contains_insertMany_list [EquivBEq α] [LawfulHashable α] m.contains k := DHashMap.Const.contains_of_contains_insertMany_list contains contains_eq_false +theorem get?_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] + {l : List (α × β)} {k : α} + (contains_eq_false : (l.map Prod.fst).contains k = false) : + (insertMany m l)[k]? = m[k]? := + DHashMap.Const.get?_insertMany_list_of_contains_eq_false contains_eq_false + +theorem get?_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] + {l : List (α × β)} {k k' : α} (k_beq : k == k') {v : β} + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) (mem : ⟨k, v⟩ ∈ l) : + (insertMany m l)[k']? = some v := + DHashMap.Const.get?_insertMany_list_of_mem k_beq distinct mem + +theorem get_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] + {l : List (α × β)} {k : α} + (contains_eq_false : (l.map Prod.fst).contains k = false) + {h} : + (insertMany m l)[k] = m[k]'(contains_of_contains_insertMany_list h contains_eq_false) := + DHashMap.Const.get_insertMany_list_of_contains_eq_false contains_eq_false + +theorem get_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] + {l : List (α × β)} {k k' : α} (k_beq : k == k') {v : β} + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) (mem : ⟨k, v⟩ ∈ l) {h} : + (insertMany m l)[k'] = v := + DHashMap.Const.get_insertMany_list_of_mem k_beq distinct mem + +theorem get!_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] + [Inhabited β] {l : List (α × β)} {k : α} + (contains_eq_false : (l.map Prod.fst).contains k = false) : + (insertMany m l)[k]! = m[k]! := + DHashMap.Const.get!_insertMany_list_of_contains_eq_false contains_eq_false + +theorem get!_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] [Inhabited β] + {l : List (α × β)} {k k' : α} (k_beq : k == k') {v : β} + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) (mem : ⟨k, v⟩ ∈ l) : + (insertMany m l)[k']! = v := + DHashMap.Const.get!_insertMany_list_of_mem k_beq distinct mem + +theorem getD_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] + {l : List (α × β)} {k : α} {fallback : β} + (contains_eq_false : (l.map Prod.fst).contains k = false) : + getD (insertMany m l) k fallback = getD m k fallback := + DHashMap.Const.getD_insertMany_list_of_contains_eq_false contains_eq_false + +theorem getD_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] + {l : List (α × β)} {k k' : α} (k_beq : k == k') {v fallback : β} + (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) (mem : ⟨k, v⟩ ∈ l) : + getD (insertMany m l) k' fallback = v := + DHashMap.Const.getD_insertMany_list_of_mem k_beq distinct mem + theorem getKey?_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] {l : List (α × β)} {k : α} (contains_eq_false : (l.map Prod.fst).contains k = false) : @@ -807,55 +856,6 @@ theorem isEmpty_insertMany_list [EquivBEq α] [LawfulHashable α] (insertMany m l).isEmpty = (m.isEmpty && l.isEmpty) := DHashMap.Const.isEmpty_insertMany_list -theorem get?_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] - {l : List (α × β)} {k : α} - (contains_eq_false : (l.map Prod.fst).contains k = false) : - get? (insertMany m l) k = get? m k := - DHashMap.Const.get?_insertMany_list_of_contains_eq_false contains_eq_false - -theorem get?_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] - {l : List (α × β)} {k k' : α} (k_beq : k == k') {v : β} - (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) (mem : ⟨k, v⟩ ∈ l) : - get? (insertMany m l) k' = v := - DHashMap.Const.get?_insertMany_list_of_mem k_beq distinct mem - -theorem get_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] - {l : List (α × β)} {k : α} - (contains_eq_false : (l.map Prod.fst).contains k = false) - {h} : - get (insertMany m l) k h = get m k (contains_of_contains_insertMany_list h contains_eq_false) := - DHashMap.Const.get_insertMany_list_of_contains_eq_false contains_eq_false - -theorem get_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] - {l : List (α × β)} {k k' : α} (k_beq : k == k') {v : β} - (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) (mem : ⟨k, v⟩ ∈ l) {h} : - get (insertMany m l) k' h = v := - DHashMap.Const.get_insertMany_list_of_mem k_beq distinct mem - -theorem get!_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] - [Inhabited β] {l : List (α × β)} {k : α} - (contains_eq_false : (l.map Prod.fst).contains k = false) : - get! (insertMany m l) k = get! m k := - DHashMap.Const.get!_insertMany_list_of_contains_eq_false contains_eq_false - -theorem get!_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] [Inhabited β] - {l : List (α × β)} {k k' : α} (k_beq : k == k') {v : β} - (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) (mem : ⟨k, v⟩ ∈ l) : - get! (insertMany m l) k' = v := - DHashMap.Const.get!_insertMany_list_of_mem k_beq distinct mem - -theorem getD_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] - {l : List (α × β)} {k : α} {fallback : β} - (contains_eq_false : (l.map Prod.fst).contains k = false) : - getD (insertMany m l) k fallback = getD m k fallback := - DHashMap.Const.getD_insertMany_list_of_contains_eq_false contains_eq_false - -theorem getD_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] - {l : List (α × β)} {k k' : α} (k_beq : k == k') {v fallback : β} - (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) (mem : ⟨k, v⟩ ∈ l) : - getD (insertMany m l) k' fallback = v := - DHashMap.Const.getD_insertMany_list_of_mem k_beq distinct mem - variable {m : HashMap α Unit} @[simp] @@ -889,6 +889,27 @@ theorem contains_of_contains_insertManyIfNewUnit_list [EquivBEq α] [LawfulHasha (insertManyIfNewUnit m l).contains k → m.contains k := DHashMap.Const.contains_of_contains_insertManyIfNewUnit_list contains_eq_false +theorem get?_insertManyIfNewUnit_list [EquivBEq α] [LawfulHashable α] + {l : List α} {k : α} : + (insertManyIfNewUnit m l)[k]? = + if m.contains k ∨ l.contains k then some () else none := + DHashMap.Const.get?_insertManyIfNewUnit_list + +theorem get_insertManyIfNewUnit_list + {l : List α} {k : α} {h} : + (insertManyIfNewUnit m l)[k] = () := + DHashMap.Const.get_insertManyIfNewUnit_list + +theorem get!_insertManyIfNewUnit_list + {l : List α} {k : α} : + (insertManyIfNewUnit m l)[k]! = () := + DHashMap.Const.get!_insertManyIfNewUnit_list + +theorem getD_insertManyIfNewUnit_list + {l : List α} {k : α} {fallback : Unit} : + getD (insertManyIfNewUnit m l) k fallback = () := by + simp + theorem getKey?_insertManyIfNewUnit_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] {l : List α} {k : α} (contains_eq_false : l.contains k = false) : getKey? (insertManyIfNewUnit m l) k = getKey? m k := @@ -989,27 +1010,6 @@ theorem isEmpty_insertManyIfNewUnit_list [EquivBEq α] [LawfulHashable α] (insertManyIfNewUnit m l).isEmpty = (m.isEmpty && l.isEmpty) := DHashMap.Const.isEmpty_insertManyIfNewUnit_list -theorem get?_insertManyIfNewUnit_list [EquivBEq α] [LawfulHashable α] - {l : List α} {k : α} : - get? (insertManyIfNewUnit m l) k = - if m.contains k ∨ l.contains k then some () else none := - DHashMap.Const.get?_insertManyIfNewUnit_list - -theorem get_insertManyIfNewUnit_list - {l : List α} {k : α} {h} : - get (insertManyIfNewUnit m l) k h = () := - DHashMap.Const.get_insertManyIfNewUnit_list - -theorem get!_insertManyIfNewUnit_list - {l : List α} {k : α} : - get! (insertManyIfNewUnit m l) k = () := - DHashMap.Const.get!_insertManyIfNewUnit_list - -theorem getD_insertManyIfNewUnit_list - {l : List α} {k : α} {fallback : Unit} : - getD (insertManyIfNewUnit m l) k fallback = () := by - simp - end section @@ -1043,14 +1043,14 @@ theorem mem_ofList [EquivBEq α] [LawfulHashable α] theorem get?_ofList_of_contains_eq_false [LawfulBEq α] {l : List (α × β)} {k : α} (contains_eq_false : (l.map Prod.fst).contains k = false) : - get? (ofList l) k = none := + (ofList l)[k]? = none := DHashMap.Const.get?_ofList_of_contains_eq_false contains_eq_false theorem get?_ofList_of_mem [LawfulBEq α] {l : List (α × β)} {k k' : α} (k_beq : k == k') {v : β} (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) (mem : ⟨k, v⟩ ∈ l) : - get? (ofList l) k' = some v := + (ofList l)[k']? = some v := DHashMap.Const.get?_ofList_of_mem k_beq distinct mem theorem get_ofList_of_mem [LawfulBEq α] @@ -1058,20 +1058,20 @@ theorem get_ofList_of_mem [LawfulBEq α] (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) (mem : ⟨k, v⟩ ∈ l) {h} : - get (ofList l) k' h = v := + (ofList l)[k'] = v := DHashMap.Const.get_ofList_of_mem k_beq distinct mem theorem get!_ofList_of_contains_eq_false [LawfulBEq α] {l : List (α × β)} {k : α} [Inhabited β] (contains_eq_false : (l.map Prod.fst).contains k = false) : - get! (ofList l) k = (default : β) := + (ofList l)[k]! = (default : β) := DHashMap.Const.get!_ofList_of_contains_eq_false contains_eq_false theorem get!_ofList_of_mem [LawfulBEq α] {l : List (α × β)} {k k' : α} (k_beq : k == k') {v : β} [Inhabited β] (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) (mem : ⟨k, v⟩ ∈ l) : - get! (ofList l) k' = v := + (ofList l)[k']! = v := DHashMap.Const.get!_ofList_of_mem k_beq distinct mem theorem getD_ofList_of_contains_eq_false [LawfulBEq α] @@ -1181,6 +1181,31 @@ theorem mem_unitOfList [EquivBEq α] [LawfulHashable α] k ∈ unitOfList l ↔ l.contains k := DHashMap.Const.mem_unitOfList +@[simp] +theorem get?_unitOfList [EquivBEq α] [LawfulHashable α] + {l : List α} {k : α} : + (unitOfList l)[k]? = + if l.contains k then some () else none := + DHashMap.Const.get?_unitOfList + +@[simp] +theorem get_unitOfList + {l : List α} {k : α} {h} : + (unitOfList l)[k] = () := + DHashMap.Const.get_unitOfList + +@[simp] +theorem get!_unitOfList + {l : List α} {k : α} : + (unitOfList l)[k]! = () := + DHashMap.Const.get!_unitOfList + +@[simp] +theorem getD_unitOfList + {l : List α} {k : α} {fallback : Unit} : + getD (unitOfList l) k fallback = () := by + simp + theorem getKey?_unitOfList_of_contains_eq_false [EquivBEq α] [LawfulHashable α] {l : List α} {k : α} (contains_eq_false : l.contains k = false) : getKey? (unitOfList l) k = none := @@ -1243,31 +1268,6 @@ theorem isEmpty_unitOfList [EquivBEq α] [LawfulHashable α] (unitOfList l).isEmpty = l.isEmpty := DHashMap.Const.isEmpty_unitOfList -@[simp] -theorem get?_unitOfList [EquivBEq α] [LawfulHashable α] - {l : List α} {k : α} : - get? (unitOfList l) k = - if l.contains k then some () else none := - DHashMap.Const.get?_unitOfList - -@[simp] -theorem get_unitOfList - {l : List α} {k : α} {h} : - get (unitOfList l) k h = () := - DHashMap.Const.get_unitOfList - -@[simp] -theorem get!_unitOfList - {l : List α} {k : α} : - get! (unitOfList l) k = () := - DHashMap.Const.get!_unitOfList - -@[simp] -theorem getD_unitOfList - {l : List α} {k : α} {fallback : Unit} : - getD (unitOfList l) k fallback = () := by - simp - end end Std.HashMap From 26c9ec724fed7d2a4c060b3e7197f5923db76740 Mon Sep 17 00:00:00 2001 From: jt0202 Date: Thu, 9 Jan 2025 15:07:27 +0100 Subject: [PATCH 72/83] Result for HashSets --- src/Std/Data/DHashMap/RawLemmas.lean | 2 +- src/Std/Data/HashMap/RawLemmas.lean | 4 +- src/Std/Data/HashSet/Lemmas.lean | 220 +++++++++++++++++++++++++++ src/Std/Data/HashSet/RawLemmas.lean | 218 +++++++++++++++++++++++++- 4 files changed, 440 insertions(+), 4 deletions(-) diff --git a/src/Std/Data/DHashMap/RawLemmas.lean b/src/Std/Data/DHashMap/RawLemmas.lean index 13c014050738..67b85e20aa97 100644 --- a/src/Std/Data/DHashMap/RawLemmas.lean +++ b/src/Std/Data/DHashMap/RawLemmas.lean @@ -1399,7 +1399,7 @@ theorem insertManyIfNewUnit_list_singleton [EquivBEq α] [LawfulHashable α] {k rw [Raw₀.Const.insertManyIfNewUnit_list_singleton] theorem insertManyIfNewUnit_cons [EquivBEq α] [LawfulHashable α] (h : m.WF) {l : List α} {k : α} : - (insertManyIfNewUnit m (k :: l)) = (insertManyIfNewUnit (m.insertIfNew k ()) l) := by + insertManyIfNewUnit m (k :: l) = insertManyIfNewUnit (m.insertIfNew k ()) l := by simp_to_raw rw [Raw₀.Const.insertManyIfNewUnit_cons] diff --git a/src/Std/Data/HashMap/RawLemmas.lean b/src/Std/Data/HashMap/RawLemmas.lean index 051aa22143f5..4968d71ec930 100644 --- a/src/Std/Data/HashMap/RawLemmas.lean +++ b/src/Std/Data/HashMap/RawLemmas.lean @@ -878,7 +878,7 @@ theorem insertManyIfNewUnit_list_singleton [EquivBEq α] [LawfulHashable α] (h ext (DHashMap.Raw.Const.insertManyIfNewUnit_list_singleton h.out) theorem insertManyIfNewUnit_cons [EquivBEq α] [LawfulHashable α] (h : m.WF) {l : List α} {k : α} : - (insertManyIfNewUnit m (k :: l)) = (insertManyIfNewUnit (m.insertIfNew k ()) l) := + insertManyIfNewUnit m (k :: l) = insertManyIfNewUnit (m.insertIfNew k ()) l := ext (DHashMap.Raw.Const.insertManyIfNewUnit_cons h.out) @[simp] @@ -947,7 +947,7 @@ theorem getKey!_insertManyIfNewUnit_list_of_contains_eq_false [EquivBEq α] [Law DHashMap.Raw.Const.getKey!_insertManyIfNewUnit_list_of_contains_eq_false h.out contains_eq_false theorem getKey!_insertManyIfNewUnit_list_of_mem_of_contains_eq_false [EquivBEq α] [LawfulHashable α] - [Inhabited α] (h : m.WF) {l : List α} {k k' : α} (k_beq : k == k') + [Inhabited α] (h : m.WF) {l : List α} {k k' : α} (k_beq : k == k') (distinct : l.Pairwise (fun a b => (a == b) = false)) (mem : k ∈ l) : contains m k = false → getKey! (insertManyIfNewUnit m l) k' = k := diff --git a/src/Std/Data/HashSet/Lemmas.lean b/src/Std/Data/HashSet/Lemmas.lean index 232724f26c0d..66f2a7126724 100644 --- a/src/Std/Data/HashSet/Lemmas.lean +++ b/src/Std/Data/HashSet/Lemmas.lean @@ -373,6 +373,226 @@ theorem mem_toList [LawfulBEq α] [LawfulHashable α] {k : α} : theorem distinct_toList [EquivBEq α] [LawfulHashable α]: m.toList.Pairwise (fun a b => (a == b) = false) := HashMap.distinct_keys + +@[simp] +theorem insertMany_nil : + insertMany m [] = m := + ext HashMap.insertManyIfNewUnit_nil + +@[simp] +theorem insertMany_list_singleton [EquivBEq α] [LawfulHashable α] {k : α} : + insertMany m [k] = m.insert k := + ext HashMap.insertManyIfNewUnit_list_singleton + +theorem insertMany_cons {l : List α} {k : α} : + insertMany m (k :: l) = insertMany (m.insert k) l := + ext HashMap.insertManyIfNewUnit_cons + +@[simp] +theorem contains_insertMany_list [EquivBEq α] [LawfulHashable α] + {l : List α} {k : α} : + (insertMany m l).contains k = (m.contains k || l.contains k) := + HashMap.contains_insertManyIfNewUnit_list + +@[simp] +theorem mem_insertMany_list [EquivBEq α] [LawfulHashable α] + {l : List α} {k : α} : + k ∈ insertMany m l ↔ k ∈ m ∨ l.contains k := + HashMap.mem_insertManyIfNewUnit_list + +theorem contains_of_contains_insertMany_list [EquivBEq α] [LawfulHashable α] + {l : List α} {k : α} (contains_eq_false : l.contains k = false) : + (insertMany m l).contains k → m.contains k := + HashMap.contains_of_contains_insertManyIfNewUnit_list contains_eq_false + +theorem getKey?_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] + {l : List α} {k : α} (contains_eq_false : l.contains k = false) : + get? (insertMany m l) k = get? m k := + HashMap.getKey?_insertManyIfNewUnit_list_of_contains_eq_false contains_eq_false + +theorem getKey?_insertMany_list_of_mem_of_contains_eq_false [EquivBEq α] [LawfulHashable α] + {l : List α} {k k' : α} (k_beq : k == k') + (distinct : l.Pairwise (fun a b => (a == b) = false)) (mem : k ∈ l) + (contains_eq_false : m.contains k = false) : + get? (insertMany m l) k' = some k := + HashMap.getKey?_insertManyIfNewUnit_list_of_mem_of_contains_eq_false k_beq distinct mem contains_eq_false + +theorem getKey?_insertMany_list_of_contains_of_contains [EquivBEq α] [LawfulHashable α] + {l : List α} {k : α} + (distinct : l.Pairwise (fun a b => (a == b) = false)) + (h : l.contains k) + (h' : m.contains k) : + get? (insertMany m l) k = get? m k := + HashMap.getKey?_insertManyIfNewUnit_list_of_contains_of_contains distinct h h' + +theorem getKey_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] + {l : List α} {k : α} (contains_eq_false : l.contains k = false) {h} : + get (insertMany m l) k h = + get m k (contains_of_contains_insertMany_list contains_eq_false h) := + HashMap.getKey_insertManyIfNewUnit_list_of_contains_eq_false contains_eq_false + +theorem getKey_insertMany_list_of_mem_of_contains_eq_false [EquivBEq α] [LawfulHashable α] + {l : List α} + {k k' : α} (k_beq : k == k') + (distinct : l.Pairwise (fun a b => (a == b) = false)) + (mem : k ∈ l) {h} + (contains_eq_false : m.contains k = false) : + get (insertMany m l) k' h = k := + HashMap.getKey_insertManyIfNewUnit_list_of_mem_of_contains_eq_false k_beq distinct mem contains_eq_false + +theorem getKey_insertMany_list_mem_of_contains_of_contains [EquivBEq α] [LawfulHashable α] + {l : List α} {k : α} + (distinct : l.Pairwise (fun a b => (a == b) = false)) + (h₁ : l.contains k) {h₂ : m.contains k} {h} : + get (insertMany m l) k h = get m k h₂ := + HashMap.getKey_insertManyIfNewUnit_list_mem_of_contains_of_contains distinct h₁ + +theorem getKey!_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] + [Inhabited α] {l : List α} {k : α} + (contains_eq_false : l.contains k = false) : + get! (insertMany m l) k = get! m k := + HashMap.getKey!_insertManyIfNewUnit_list_of_contains_eq_false contains_eq_false + +theorem getKey!_insertMany_list_of_mem_of_contains_eq_false [EquivBEq α] [LawfulHashable α] + [Inhabited α] {l : List α} {k k' : α} (k_beq : k == k') + (distinct : l.Pairwise (fun a b => (a == b) = false)) + (mem : k ∈ l) : + contains m k = false → get! (insertMany m l) k' = k := + HashMap.getKey!_insertManyIfNewUnit_list_of_mem_of_contains_eq_false k_beq distinct mem + +theorem getKey!_insertMany_list_mem_of_contains_of_contains [EquivBEq α] [LawfulHashable α] + [Inhabited α] {l : List α} {k : α} + (distinct : l.Pairwise (fun a b => (a == b) = false)) + (h : l.contains k) : + m.contains k → get! (insertMany m l) k = get! m k := + HashMap.getKey!_insertManyIfNewUnit_list_mem_of_contains_of_contains distinct h + +theorem getKeyD_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] + {l : List α} {k fallback : α} + (contains_eq_false : l.contains k = false) : + getD (insertMany m l) k fallback = getD m k fallback := + HashMap.getKeyD_insertManyIfNewUnit_list_of_contains_eq_false contains_eq_false + +theorem getKeyD_insertMany_list_of_mem_of_contains_eq_false [EquivBEq α] [LawfulHashable α] + {l : List α} {k k' fallback : α} (k_beq : k == k') + (distinct : l.Pairwise (fun a b => (a == b) = false)) + (mem : k ∈ l ) (contains_eq_false : m.contains k = false) : + getD (insertMany m l) k' fallback = k := + HashMap.getKeyD_insertManyIfNewUnit_list_of_mem_of_contains_eq_false k_beq distinct mem contains_eq_false + +theorem getKeyD_insertMany_list_of_contains_of_contains [EquivBEq α] [LawfulHashable α] + {l : List α} {k fallback : α} + (distinct : l.Pairwise (fun a b => (a == b) = false)) + (h : l.contains k) : + m.contains k → getD (insertMany m l) k fallback = getD m k fallback := + HashMap.getKeyD_insertManyIfNewUnit_list_of_contains_of_contains distinct h + +theorem size_insertMany_list [EquivBEq α] [LawfulHashable α] + {l : List α} + (distinct : l.Pairwise (fun a b => (a == b) = false)) : + (∀ (a : α), m.contains a → l.contains a = false) → + (insertMany m l).size = m.size + l.length := + HashMap.size_insertManyIfNewUnit_list distinct + +theorem size_insertMany_list_le [EquivBEq α] [LawfulHashable α] + {l : List α} : + (insertMany m l).size ≤ m.size + l.length := + HashMap.size_insertManyIfNewUnit_list_le + +@[simp] +theorem isEmpty_insertMany_list [EquivBEq α] [LawfulHashable α] + {l : List α} : + (insertMany m l).isEmpty = (m.isEmpty && l.isEmpty) := + HashMap.isEmpty_insertManyIfNewUnit_list + +end + +section + +@[simp] +theorem ofList_nil [EquivBEq α] [LawfulHashable α] : + ofList ([] : List α) = ∅ := + ext HashMap.unitOfList_nil + +@[simp] +theorem ofList_singleton [EquivBEq α] [LawfulHashable α] {k : α} : + ofList [k] = (∅ : HashSet α).insert k := + ext HashMap.unitOfList_singleton + +theorem ofList_cons [EquivBEq α] [LawfulHashable α] {hd : α} {tl : List α} : + ofList (hd :: tl) = + insertMany ((∅ : HashSet α).insert hd) tl := + ext HashMap.unitOfList_cons + +@[simp] +theorem contains_ofList [EquivBEq α] [LawfulHashable α] + {l : List α} {k : α} : + (ofList l).contains k = l.contains k := + HashMap.contains_unitOfList + +theorem get?_ofList_of_contains_eq_false [EquivBEq α] [LawfulHashable α] + {l : List α} {k : α} (contains_eq_false : l.contains k = false) : + get? (ofList l) k = none := + HashMap.getKey?_unitOfList_of_contains_eq_false contains_eq_false + +theorem get?_ofList_of_mem [EquivBEq α] [LawfulHashable α] + {l : List α} {k k' : α} (k_beq : k == k') + (distinct : l.Pairwise (fun a b => (a == b) = false)) (mem : k ∈ l) : + get? (ofList l) k' = some k := + HashMap.getKey?_unitOfList_of_mem k_beq distinct mem + +theorem getKey_ofList_of_mem [EquivBEq α] [LawfulHashable α] + {l : List α} + {k k' : α} (k_beq : k == k') + (distinct : l.Pairwise (fun a b => (a == b) = false)) + (mem : k ∈ l) {h} : + get (ofList l) k' h = k := + HashMap.getKey_unitOfList_of_mem k_beq distinct mem + +theorem getKey!_ofList_of_contains_eq_false [EquivBEq α] [LawfulHashable α] + [Inhabited α] {l : List α} {k : α} + (contains_eq_false : l.contains k = false) : + get! (ofList l) k = default := + HashMap.getKey!_unitOfList_of_contains_eq_false contains_eq_false + +theorem getKey!_ofList_of_mem [EquivBEq α] [LawfulHashable α] + [Inhabited α] {l : List α} {k k' : α} (k_beq : k == k') + (distinct : l.Pairwise (fun a b => (a == b) = false)) + (mem : k ∈ l) : + get! (ofList l) k' = k := + HashMap.getKey!_unitOfList_of_mem k_beq distinct mem + +theorem getKeyD_ofList_of_contains_eq_false [EquivBEq α] [LawfulHashable α] + {l : List α} {k fallback : α} + (contains_eq_false : l.contains k = false) : + getD (ofList l) k fallback = fallback := + HashMap.getKeyD_unitOfList_of_contains_eq_false contains_eq_false + +theorem getKeyD_ofList_of_mem [EquivBEq α] [LawfulHashable α] + {l : List α} {k k' fallback : α} (k_beq : k == k') + (distinct : l.Pairwise (fun a b => (a == b) = false)) + (mem : k ∈ l) : + getD (ofList l) k' fallback = k := + HashMap.getKeyD_unitOfList_of_mem k_beq distinct mem + +theorem size_ofList [EquivBEq α] [LawfulHashable α] + {l : List α} + (distinct : l.Pairwise (fun a b => (a == b) = false)) : + (ofList l).size = l.length := + HashMap.size_unitOfList distinct + +theorem size_ofList_le [EquivBEq α] [LawfulHashable α] + {l : List α} : + (ofList l).size ≤ l.length := + HashMap.size_unitOfList_le + +@[simp] +theorem isEmpty_ofList [EquivBEq α] [LawfulHashable α] + {l : List α} : + (ofList l).isEmpty = l.isEmpty := + HashMap.isEmpty_unitOfList + + end end Std.HashSet diff --git a/src/Std/Data/HashSet/RawLemmas.lean b/src/Std/Data/HashSet/RawLemmas.lean index 77dcaa5bcfd8..0a82d686144f 100644 --- a/src/Std/Data/HashSet/RawLemmas.lean +++ b/src/Std/Data/HashSet/RawLemmas.lean @@ -367,7 +367,7 @@ theorem containsThenInsert_snd (h : m.WF) {k : α} : (m.containsThenInsert k).2 ext (HashMap.Raw.containsThenInsertIfNew_snd h.out) @[simp] -theorem length_toList [EquivBEq α] [LawfulHashable α] (h : m.WF): +theorem length_toList [EquivBEq α] [LawfulHashable α] (h : m.WF): m.toList.length = m.size := HashMap.Raw.length_keys h.1 @@ -390,6 +390,222 @@ theorem distinct_toList [EquivBEq α] [LawfulHashable α] (h : m.WF) : m.toList.Pairwise (fun a b => (a == b) = false) := HashMap.Raw.distinct_keys h.1 +@[simp] +theorem insertMany_nil [EquivBEq α] [LawfulHashable α] (h : m.WF) : + insertMany m [] = m := + ext (HashMap.Raw.insertManyIfNewUnit_nil h.1) + +@[simp] +theorem insertMany_list_singleton [EquivBEq α] [LawfulHashable α] (h : m.WF) {k : α} : + insertMany m [k] = m.insert k := + ext (HashMap.Raw.insertManyIfNewUnit_list_singleton h.1) + +theorem insertMany_cons [EquivBEq α] [LawfulHashable α] (h : m.WF) {l : List α} {k : α} : + insertMany m (k :: l) = insertMany (m.insert k) l := + ext (HashMap.Raw.insertManyIfNewUnit_cons h.1) + +@[simp] +theorem contains_insertMany_list [EquivBEq α] [LawfulHashable α] (h : m.WF) + {l : List α} {k : α} : + (insertMany m l).contains k = (m.contains k || l.contains k) := + HashMap.Raw.contains_insertManyIfNewUnit_list h.1 + +@[simp] +theorem mem_insertMany_list [EquivBEq α] [LawfulHashable α] (h : m.WF) + {l : List α} {k : α} : + k ∈ insertMany m l ↔ k ∈ m ∨ l.contains k := + HashMap.Raw.mem_insertManyIfNewUnit_list h.1 + +theorem contains_of_contains_insertMany_list [EquivBEq α] [LawfulHashable α] (h : m.WF) + {l : List α} {k : α} (contains_eq_false : l.contains k = false) : + (insertMany m l).contains k → m.contains k := + HashMap.Raw.contains_of_contains_insertManyIfNewUnit_list h.1 contains_eq_false + +theorem get?_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] + (h : m.WF) {l : List α} {k : α} (contains_eq_false : l.contains k = false) : + get? (insertMany m l) k = get? m k := + HashMap.Raw.getKey?_insertManyIfNewUnit_list_of_contains_eq_false h.1 contains_eq_false + +theorem get?_insertMany_list_of_mem_of_contains_eq_false [EquivBEq α] [LawfulHashable α] + (h : m.WF) {l : List α} {k k' : α} (k_beq : k == k') + (distinct : l.Pairwise (fun a b => (a == b) = false)) (mem : k ∈ l) : + m.contains k = false → get? (insertMany m l) k' = some k := + HashMap.Raw.getKey?_insertManyIfNewUnit_list_of_mem_of_contains_eq_false h.1 k_beq distinct mem + +theorem get?_insertMany_list_of_contains_of_contains [EquivBEq α] [LawfulHashable α] + (h : m.WF) {l : List α} {k : α} + (distinct : l.Pairwise (fun a b => (a == b) = false)) + (h' : l.contains k) : + m.contains k → get? (insertMany m l) k = get? m k := + HashMap.Raw.getKey?_insertManyIfNewUnit_list_of_contains_of_contains h.1 distinct h' + +theorem get_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] + (h : m.WF) {l : List α} {k : α} (contains_eq_false : l.contains k = false) {h'} : + get (insertMany m l) k h' = + get m k (contains_of_contains_insertMany_list h contains_eq_false h') := + HashMap.Raw.getKey_insertManyIfNewUnit_list_of_contains_eq_false h.1 contains_eq_false + +theorem get_insertMany_list_of_mem_of_contains_eq_false [EquivBEq α] [LawfulHashable α] + (h : m.WF) {l : List α} + {k k' : α} (k_beq : k == k') + (distinct : l.Pairwise (fun a b => (a == b) = false)) + (mem : k ∈ l) {h'} : + m.contains k = false → get (insertMany m l) k' h' = k := + HashMap.Raw.getKey_insertManyIfNewUnit_list_of_mem_of_contains_eq_false h.1 k_beq distinct mem + +theorem get_insertMany_list_mem_of_contains_of_contains [EquivBEq α] [LawfulHashable α] + (h : m.WF) {l : List α} {k : α} + (distinct : l.Pairwise (fun a b => (a == b) = false)) + (h₁ : l.contains k) (h₂ : m.contains k) {h₃} : + get (insertMany m l) k h₃ = get m k h₂ := + HashMap.Raw.getKey_insertManyIfNewUnit_list_mem_of_contains_of_contains h.1 distinct h₁ h₂ + +theorem get!_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] + [Inhabited α] (h : m.WF) {l : List α} {k : α} + (contains_eq_false : l.contains k = false) : + get! (insertMany m l) k = get! m k := + HashMap.Raw.getKey!_insertManyIfNewUnit_list_of_contains_eq_false h.1 contains_eq_false + +theorem get!_insertMany_list_of_mem_of_contains_eq_false [EquivBEq α] [LawfulHashable α] + [Inhabited α] (h : m.WF) {l : List α} {k k' : α} (k_beq : k == k') + (distinct : l.Pairwise (fun a b => (a == b) = false)) + (mem : k ∈ l) : + contains m k = false → get! (insertMany m l) k' = k := + HashMap.Raw.getKey!_insertManyIfNewUnit_list_of_mem_of_contains_eq_false h.1 k_beq distinct mem + +theorem get!_insertMany_list_mem_of_contains_of_contains [EquivBEq α] [LawfulHashable α] + [Inhabited α] (h : m.WF) {l : List α} {k : α} + (distinct : l.Pairwise (fun a b => (a == b) = false)) + (h' : l.contains k) : + m.contains k → get! (insertMany m l) k = get! m k := + HashMap.Raw.getKey!_insertManyIfNewUnit_list_mem_of_contains_of_contains h.1 distinct h' + +theorem getD_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] + (h : m.WF) {l : List α} {k fallback : α} + (contains_eq_false : l.contains k = false) : + getD (insertMany m l) k fallback = getD m k fallback := + HashMap.Raw.getKeyD_insertManyIfNewUnit_list_of_contains_eq_false h.1 contains_eq_false + +theorem getD_insertMany_list_of_mem_of_contains_eq_false [EquivBEq α] [LawfulHashable α] + (h : m.WF) {l : List α} {k k' fallback : α} (k_beq : k == k') + (distinct : l.Pairwise (fun a b => (a == b) = false)) + (mem : k ∈ l) : + m.contains k = false → getD (insertMany m l) k' fallback = k := + HashMap.Raw.getKeyD_insertManyIfNewUnit_list_of_mem_of_contains_eq_false h.1 k_beq distinct mem + +theorem getD_insertMany_list_of_contains_of_contains [EquivBEq α] [LawfulHashable α] + (h : m.WF) {l : List α} {k fallback : α} + (distinct : l.Pairwise (fun a b => (a == b) = false)) + (h' : l.contains k) : + m.contains k → getD (insertMany m l) k fallback = getD m k fallback := + HashMap.Raw.getKeyD_insertManyIfNewUnit_list_of_contains_of_contains h.1 distinct h' + +theorem size_insertMany_list [EquivBEq α] [LawfulHashable α] (h : m.WF) + {l : List α} + (distinct : l.Pairwise (fun a b => (a == b) = false)) : + (∀ (a : α), m.contains a → l.contains a = false) → + (insertMany m l).size = m.size + l.length := + HashMap.Raw.size_insertManyIfNewUnit_list h.1 distinct + +theorem size_insertMany_list_le [EquivBEq α] [LawfulHashable α] (h : m.WF) + {l : List α} : + (insertMany m l).size ≤ m.size + l.length := + HashMap.Raw.size_insertManyIfNewUnit_list_le h.1 + +@[simp] +theorem isEmpty_insertMany_list [EquivBEq α] [LawfulHashable α] (h : m.WF) + {l : List α} : + (insertMany m l).isEmpty = (m.isEmpty && l.isEmpty) := + HashMap.Raw.isEmpty_insertManyIfNewUnit_list h.1 + +@[simp] +theorem ofList_nil [EquivBEq α] [LawfulHashable α] : + ofList ([] : List α) = ∅ := + ext HashMap.Raw.unitOfList_nil + +@[simp] +theorem ofList_singleton [EquivBEq α] [LawfulHashable α] {k : α} : + ofList [k] = (∅ : Raw α).insert k := + ext HashMap.Raw.unitOfList_singleton + +theorem ofList_cons [EquivBEq α] [LawfulHashable α] {hd : α} {tl : List α} : + ofList (hd :: tl) = insertMany ((∅ : Raw α).insert hd) tl := + ext HashMap.Raw.unitOfList_cons + +@[simp] +theorem contains_ofList [EquivBEq α] [LawfulHashable α] + {l : List α} {k : α} : + (ofList l).contains k = l.contains k := + HashMap.Raw.contains_unitOfList + +@[simp] +theorem mem_ofList [EquivBEq α] [LawfulHashable α] + {l : List α} {k : α} : + k ∈ ofList l ↔ l.contains k := + HashMap.Raw.mem_unitOfList + +theorem get?_ofList_of_contains_eq_false [EquivBEq α] [LawfulHashable α] + {l : List α} {k : α} (contains_eq_false : l.contains k = false) : + get? (ofList l) k = none := + HashMap.Raw.getKey?_unitOfList_of_contains_eq_false contains_eq_false + +theorem get?_ofList_of_mem [EquivBEq α] [LawfulHashable α] + {l : List α} {k k' : α} (k_beq : k == k') + (distinct : l.Pairwise (fun a b => (a == b) = false)) (mem : k ∈ l) : + get? (ofList l) k' = some k := + HashMap.Raw.getKey?_unitOfList_of_mem k_beq distinct mem + +theorem get_ofList_of_mem [EquivBEq α] [LawfulHashable α] + {l : List α} + {k k' : α} (k_beq : k == k') + (distinct : l.Pairwise (fun a b => (a == b) = false)) + (mem : k ∈ l) {h'} : + get (ofList l) k' h' = k := + HashMap.Raw.getKey_unitOfList_of_mem k_beq distinct mem + +theorem get!_ofList_of_contains_eq_false [EquivBEq α] [LawfulHashable α] + [Inhabited α] {l : List α} {k : α} + (contains_eq_false : l.contains k = false) : + get! (ofList l) k = default := + HashMap.Raw.getKey!_unitOfList_of_contains_eq_false contains_eq_false + +theorem get!_ofList_of_mem [EquivBEq α] [LawfulHashable α] + [Inhabited α] {l : List α} {k k' : α} (k_beq : k == k') + (distinct : l.Pairwise (fun a b => (a == b) = false)) + (mem : k ∈ l) : + get! (ofList l) k' = k := + HashMap.Raw.getKey!_unitOfList_of_mem k_beq distinct mem + +theorem getD_ofList_of_contains_eq_false [EquivBEq α] [LawfulHashable α] + {l : List α} {k fallback : α} + (contains_eq_false : l.contains k = false) : + getD (ofList l) k fallback = fallback := + HashMap.Raw.getKeyD_unitOfList_of_contains_eq_false contains_eq_false + +theorem getD_ofList_of_mem [EquivBEq α] [LawfulHashable α] + {l : List α} {k k' fallback : α} (k_beq : k == k') + (distinct : l.Pairwise (fun a b => (a == b) = false)) + (mem : k ∈ l) : + getD (ofList l) k' fallback = k := + HashMap.Raw.getKeyD_unitOfList_of_mem k_beq distinct mem + +theorem size_ofList [EquivBEq α] [LawfulHashable α] + {l : List α} + (distinct : l.Pairwise (fun a b => (a == b) = false)) : + (ofList l).size = l.length := + HashMap.Raw.size_unitOfList distinct + +theorem size_ofList_le [EquivBEq α] [LawfulHashable α] + {l : List α} : + (ofList l).size ≤ l.length := + HashMap.Raw.size_unitOfList_le + +@[simp] +theorem isEmpty_ofList [EquivBEq α] [LawfulHashable α] + {l : List α} : + (ofList l).isEmpty = l.isEmpty := + HashMap.Raw.isEmpty_unitOfList + end Raw end Std.HashSet From d4edd979960805d6d8f9a2de9613487f7c307834 Mon Sep 17 00:00:00 2001 From: jt0202 Date: Thu, 9 Jan 2025 15:08:28 +0100 Subject: [PATCH 73/83] Fix names --- src/Std/Data/HashSet/Lemmas.lean | 34 ++++++++++++++++---------------- 1 file changed, 17 insertions(+), 17 deletions(-) diff --git a/src/Std/Data/HashSet/Lemmas.lean b/src/Std/Data/HashSet/Lemmas.lean index 66f2a7126724..6900e62c7283 100644 --- a/src/Std/Data/HashSet/Lemmas.lean +++ b/src/Std/Data/HashSet/Lemmas.lean @@ -405,19 +405,19 @@ theorem contains_of_contains_insertMany_list [EquivBEq α] [LawfulHashable α] (insertMany m l).contains k → m.contains k := HashMap.contains_of_contains_insertManyIfNewUnit_list contains_eq_false -theorem getKey?_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] +theorem get?_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] {l : List α} {k : α} (contains_eq_false : l.contains k = false) : get? (insertMany m l) k = get? m k := HashMap.getKey?_insertManyIfNewUnit_list_of_contains_eq_false contains_eq_false -theorem getKey?_insertMany_list_of_mem_of_contains_eq_false [EquivBEq α] [LawfulHashable α] +theorem get?_insertMany_list_of_mem_of_contains_eq_false [EquivBEq α] [LawfulHashable α] {l : List α} {k k' : α} (k_beq : k == k') (distinct : l.Pairwise (fun a b => (a == b) = false)) (mem : k ∈ l) (contains_eq_false : m.contains k = false) : get? (insertMany m l) k' = some k := HashMap.getKey?_insertManyIfNewUnit_list_of_mem_of_contains_eq_false k_beq distinct mem contains_eq_false -theorem getKey?_insertMany_list_of_contains_of_contains [EquivBEq α] [LawfulHashable α] +theorem get?_insertMany_list_of_contains_of_contains [EquivBEq α] [LawfulHashable α] {l : List α} {k : α} (distinct : l.Pairwise (fun a b => (a == b) = false)) (h : l.contains k) @@ -425,13 +425,13 @@ theorem getKey?_insertMany_list_of_contains_of_contains [EquivBEq α] [LawfulHas get? (insertMany m l) k = get? m k := HashMap.getKey?_insertManyIfNewUnit_list_of_contains_of_contains distinct h h' -theorem getKey_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] +theorem get_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] {l : List α} {k : α} (contains_eq_false : l.contains k = false) {h} : get (insertMany m l) k h = get m k (contains_of_contains_insertMany_list contains_eq_false h) := HashMap.getKey_insertManyIfNewUnit_list_of_contains_eq_false contains_eq_false -theorem getKey_insertMany_list_of_mem_of_contains_eq_false [EquivBEq α] [LawfulHashable α] +theorem get_insertMany_list_of_mem_of_contains_eq_false [EquivBEq α] [LawfulHashable α] {l : List α} {k k' : α} (k_beq : k == k') (distinct : l.Pairwise (fun a b => (a == b) = false)) @@ -440,47 +440,47 @@ theorem getKey_insertMany_list_of_mem_of_contains_eq_false [EquivBEq α] [Lawful get (insertMany m l) k' h = k := HashMap.getKey_insertManyIfNewUnit_list_of_mem_of_contains_eq_false k_beq distinct mem contains_eq_false -theorem getKey_insertMany_list_mem_of_contains_of_contains [EquivBEq α] [LawfulHashable α] +theorem get_insertMany_list_mem_of_contains_of_contains [EquivBEq α] [LawfulHashable α] {l : List α} {k : α} (distinct : l.Pairwise (fun a b => (a == b) = false)) (h₁ : l.contains k) {h₂ : m.contains k} {h} : get (insertMany m l) k h = get m k h₂ := HashMap.getKey_insertManyIfNewUnit_list_mem_of_contains_of_contains distinct h₁ -theorem getKey!_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] +theorem get!_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] [Inhabited α] {l : List α} {k : α} (contains_eq_false : l.contains k = false) : get! (insertMany m l) k = get! m k := HashMap.getKey!_insertManyIfNewUnit_list_of_contains_eq_false contains_eq_false -theorem getKey!_insertMany_list_of_mem_of_contains_eq_false [EquivBEq α] [LawfulHashable α] +theorem get!_insertMany_list_of_mem_of_contains_eq_false [EquivBEq α] [LawfulHashable α] [Inhabited α] {l : List α} {k k' : α} (k_beq : k == k') (distinct : l.Pairwise (fun a b => (a == b) = false)) (mem : k ∈ l) : contains m k = false → get! (insertMany m l) k' = k := HashMap.getKey!_insertManyIfNewUnit_list_of_mem_of_contains_eq_false k_beq distinct mem -theorem getKey!_insertMany_list_mem_of_contains_of_contains [EquivBEq α] [LawfulHashable α] +theorem get!_insertMany_list_mem_of_contains_of_contains [EquivBEq α] [LawfulHashable α] [Inhabited α] {l : List α} {k : α} (distinct : l.Pairwise (fun a b => (a == b) = false)) (h : l.contains k) : m.contains k → get! (insertMany m l) k = get! m k := HashMap.getKey!_insertManyIfNewUnit_list_mem_of_contains_of_contains distinct h -theorem getKeyD_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] +theorem getD_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] {l : List α} {k fallback : α} (contains_eq_false : l.contains k = false) : getD (insertMany m l) k fallback = getD m k fallback := HashMap.getKeyD_insertManyIfNewUnit_list_of_contains_eq_false contains_eq_false -theorem getKeyD_insertMany_list_of_mem_of_contains_eq_false [EquivBEq α] [LawfulHashable α] +theorem getD_insertMany_list_of_mem_of_contains_eq_false [EquivBEq α] [LawfulHashable α] {l : List α} {k k' fallback : α} (k_beq : k == k') (distinct : l.Pairwise (fun a b => (a == b) = false)) (mem : k ∈ l ) (contains_eq_false : m.contains k = false) : getD (insertMany m l) k' fallback = k := HashMap.getKeyD_insertManyIfNewUnit_list_of_mem_of_contains_eq_false k_beq distinct mem contains_eq_false -theorem getKeyD_insertMany_list_of_contains_of_contains [EquivBEq α] [LawfulHashable α] +theorem getD_insertMany_list_of_contains_of_contains [EquivBEq α] [LawfulHashable α] {l : List α} {k fallback : α} (distinct : l.Pairwise (fun a b => (a == b) = false)) (h : l.contains k) : @@ -541,7 +541,7 @@ theorem get?_ofList_of_mem [EquivBEq α] [LawfulHashable α] get? (ofList l) k' = some k := HashMap.getKey?_unitOfList_of_mem k_beq distinct mem -theorem getKey_ofList_of_mem [EquivBEq α] [LawfulHashable α] +theorem get_ofList_of_mem [EquivBEq α] [LawfulHashable α] {l : List α} {k k' : α} (k_beq : k == k') (distinct : l.Pairwise (fun a b => (a == b) = false)) @@ -549,26 +549,26 @@ theorem getKey_ofList_of_mem [EquivBEq α] [LawfulHashable α] get (ofList l) k' h = k := HashMap.getKey_unitOfList_of_mem k_beq distinct mem -theorem getKey!_ofList_of_contains_eq_false [EquivBEq α] [LawfulHashable α] +theorem get!_ofList_of_contains_eq_false [EquivBEq α] [LawfulHashable α] [Inhabited α] {l : List α} {k : α} (contains_eq_false : l.contains k = false) : get! (ofList l) k = default := HashMap.getKey!_unitOfList_of_contains_eq_false contains_eq_false -theorem getKey!_ofList_of_mem [EquivBEq α] [LawfulHashable α] +theorem get!_ofList_of_mem [EquivBEq α] [LawfulHashable α] [Inhabited α] {l : List α} {k k' : α} (k_beq : k == k') (distinct : l.Pairwise (fun a b => (a == b) = false)) (mem : k ∈ l) : get! (ofList l) k' = k := HashMap.getKey!_unitOfList_of_mem k_beq distinct mem -theorem getKeyD_ofList_of_contains_eq_false [EquivBEq α] [LawfulHashable α] +theorem getD_ofList_of_contains_eq_false [EquivBEq α] [LawfulHashable α] {l : List α} {k fallback : α} (contains_eq_false : l.contains k = false) : getD (ofList l) k fallback = fallback := HashMap.getKeyD_unitOfList_of_contains_eq_false contains_eq_false -theorem getKeyD_ofList_of_mem [EquivBEq α] [LawfulHashable α] +theorem getD_ofList_of_mem [EquivBEq α] [LawfulHashable α] {l : List α} {k k' fallback : α} (k_beq : k == k') (distinct : l.Pairwise (fun a b => (a == b) = false)) (mem : k ∈ l) : From d36c5fd0477dab7bad6aca568323bb7e79479532 Mon Sep 17 00:00:00 2001 From: jt0202 Date: Fri, 10 Jan 2025 11:27:36 +0100 Subject: [PATCH 74/83] Style fixes --- src/Std/Data/DHashMap/Internal/Model.lean | 4 +- src/Std/Data/DHashMap/Lemmas.lean | 30 ++++++------- src/Std/Data/DHashMap/RawLemmas.lean | 26 ++++++----- src/Std/Data/HashMap/Lemmas.lean | 54 +++++++++++------------ src/Std/Data/HashMap/RawLemmas.lean | 52 +++++++++++----------- src/Std/Data/HashSet/Lemmas.lean | 8 ++-- src/Std/Data/HashSet/RawLemmas.lean | 8 ++-- 7 files changed, 93 insertions(+), 89 deletions(-) diff --git a/src/Std/Data/DHashMap/Internal/Model.lean b/src/Std/Data/DHashMap/Internal/Model.lean index 3ef7cbd46ce4..4a341c61159f 100644 --- a/src/Std/Data/DHashMap/Internal/Model.lean +++ b/src/Std/Data/DHashMap/Internal/Model.lean @@ -349,14 +349,14 @@ def Const.get!ₘ [BEq α] [Hashable α] [Inhabited β] (m : Raw₀ α (fun _ => (Const.get?ₘ m a).get! /-- Internal implementation detail of the hash map -/ -def Const.insertListₘ [BEq α] [Hashable α](m : Raw₀ α (fun _ => β)) (l: List (α × β)) : +def Const.insertListₘ [BEq α] [Hashable α] (m : Raw₀ α (fun _ => β)) (l: List (α × β)) : Raw₀ α (fun _ => β) := match l with | .nil => m | .cons hd tl => insertListₘ (m.insert hd.1 hd.2) tl /-- Internal implementation detail of the hash map -/ -def Const.insertListIfNewUnitₘ [BEq α] [Hashable α](m : Raw₀ α (fun _ => Unit)) (l: List α) : +def Const.insertListIfNewUnitₘ [BEq α] [Hashable α] (m : Raw₀ α (fun _ => Unit)) (l: List α) : Raw₀ α (fun _ => Unit) := match l with | .nil => m diff --git a/src/Std/Data/DHashMap/Lemmas.lean b/src/Std/Data/DHashMap/Lemmas.lean index 5ca7feee8fb9..6dac15cc9b82 100644 --- a/src/Std/Data/DHashMap/Lemmas.lean +++ b/src/Std/Data/DHashMap/Lemmas.lean @@ -989,11 +989,11 @@ theorem mem_insertMany_list [EquivBEq α] [LawfulHashable α] k ∈ m.insertMany l ↔ k ∈ m ∨ (l.map Sigma.fst).contains k := by simp [mem_iff_contains] -theorem contains_of_contains_insertMany_list [EquivBEq α] [LawfulHashable α] - {l : List ((a : α) × β a)} {k : α} (contains : (m.insertMany l).contains k) +theorem mem_of_mem_insertMany_list [EquivBEq α] [LawfulHashable α] + {l : List ((a : α) × β a)} {k : α} (mem : k ∈ (m.insertMany l)) (contains_eq_false : (l.map Sigma.fst).contains k = false) : - m.contains k := - Raw₀.contains_of_contains_insertMany_list ⟨m.1, _⟩ m.2 contains contains_eq_false + k ∈ m := + Raw₀.contains_of_contains_insertMany_list ⟨m.1, _⟩ m.2 mem contains_eq_false theorem get?_insertMany_list_of_contains_eq_false [LawfulBEq α] {l : List ((a : α) × β a)} {k : α} @@ -1013,7 +1013,7 @@ theorem get_insertMany_list_of_contains_eq_false [LawfulBEq α] (contains_eq_false : (l.map Sigma.fst).contains k = false) {h} : (m.insertMany l).get k h = - m.get k (contains_of_contains_insertMany_list h contains_eq_false) := + m.get k (mem_of_mem_insertMany_list h contains_eq_false) := Raw₀.get_insertMany_list_of_contains_eq_false ⟨m.1, _⟩ m.2 contains_eq_false theorem get_insertMany_list_of_mem [LawfulBEq α] @@ -1069,7 +1069,7 @@ theorem getKey_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashabl (contains_eq_false : (l.map Sigma.fst).contains k = false) {h} : (m.insertMany l).getKey k h = - m.getKey k (contains_of_contains_insertMany_list h contains_eq_false) := + m.getKey k (mem_of_mem_insertMany_list h contains_eq_false) := Raw₀.getKey_insertMany_list_of_contains_eq_false ⟨m.1, _⟩ m.2 contains_eq_false theorem getKey_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] @@ -1157,11 +1157,11 @@ theorem mem_insertMany_list [EquivBEq α] [LawfulHashable α] k ∈ insertMany m l ↔ k ∈ m ∨ (l.map Prod.fst).contains k := by simp [mem_iff_contains] -theorem contains_of_contains_insertMany_list [EquivBEq α] [LawfulHashable α] - {l : List ( α × β )} {k : α} (contains : (insertMany m l).contains k) +theorem mem_of_mem_insertMany_list [EquivBEq α] [LawfulHashable α] + {l : List ( α × β )} {k : α} (mem : k ∈ (insertMany m l)) (contains_eq_false : (l.map Prod.fst).contains k = false) : - m.contains k := - Raw₀.Const.contains_of_contains_insertMany_list ⟨m.1, _⟩ m.2 contains contains_eq_false + k ∈ m := + Raw₀.Const.contains_of_contains_insertMany_list ⟨m.1, _⟩ m.2 mem contains_eq_false theorem getKey?_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] {l : List (α × β)} {k : α} @@ -1182,7 +1182,7 @@ theorem getKey_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashabl (contains_eq_false : (l.map Prod.fst).contains k = false) {h} : (insertMany m l).getKey k h = - m.getKey k (contains_of_contains_insertMany_list h contains_eq_false) := + m.getKey k (mem_of_mem_insertMany_list h contains_eq_false) := Raw₀.Const.getKey_insertMany_list_of_contains_eq_false ⟨m.1, _⟩ m.2 contains_eq_false theorem getKey_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] @@ -1256,7 +1256,7 @@ theorem get_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable {l : List (α × β)} {k : α} (contains_eq_false : (l.map Prod.fst).contains k = false) {h} : - get (insertMany m l) k h = get m k (contains_of_contains_insertMany_list h contains_eq_false) := + get (insertMany m l) k h = get m k (mem_of_mem_insertMany_list h contains_eq_false) := Raw₀.Const.get_insertMany_list_of_contains_eq_false ⟨m.1, _⟩ m.2 contains_eq_false theorem get_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] @@ -1320,9 +1320,9 @@ theorem mem_insertManyIfNewUnit_list [EquivBEq α] [LawfulHashable α] k ∈ insertManyIfNewUnit m l ↔ k ∈ m ∨ l.contains k := by simp [mem_iff_contains] -theorem contains_of_contains_insertManyIfNewUnit_list [EquivBEq α] [LawfulHashable α] +theorem mem_of_mem_insertManyIfNewUnit_list [EquivBEq α] [LawfulHashable α] {l : List α} {k : α} (contains_eq_false : l.contains k = false) : - (insertManyIfNewUnit m l).contains k → m.contains k := + k ∈ (insertManyIfNewUnit m l) → k ∈ m := Raw₀.Const.contains_of_contains_insertManyIfNewUnit_list ⟨m.1, _⟩ m.2 contains_eq_false theorem getKey?_insertManyIfNewUnit_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] @@ -1349,7 +1349,7 @@ theorem getKey?_insertManyIfNewUnit_list_of_contains_of_contains [EquivBEq α] [ theorem getKey_insertManyIfNewUnit_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] {l : List α} {k : α} (contains_eq_false : l.contains k = false) {h} : getKey (insertManyIfNewUnit m l) k h = - getKey m k (contains_of_contains_insertManyIfNewUnit_list contains_eq_false h) := + getKey m k (mem_of_mem_insertManyIfNewUnit_list contains_eq_false h) := Raw₀.Const.getKey_insertManyIfNewUnit_list_of_contains_eq_false ⟨m.1, _⟩ m.2 contains_eq_false theorem getKey_insertManyIfNewUnit_list_of_mem_of_contains_eq_false [EquivBEq α] [LawfulHashable α] diff --git a/src/Std/Data/DHashMap/RawLemmas.lean b/src/Std/Data/DHashMap/RawLemmas.lean index 67b85e20aa97..c409eaf006c1 100644 --- a/src/Std/Data/DHashMap/RawLemmas.lean +++ b/src/Std/Data/DHashMap/RawLemmas.lean @@ -1081,9 +1081,10 @@ theorem mem_insertMany_list [EquivBEq α] [LawfulHashable α] (h : m.WF) k ∈ (m.insertMany l) ↔ k ∈ m ∨ (l.map Sigma.fst).contains k := by simp [mem_iff_contains, contains_insertMany_list h] -theorem contains_of_contains_insertMany_list [EquivBEq α] [LawfulHashable α] (h : m.WF) +theorem mem_of_mem_insertMany_list [EquivBEq α] [LawfulHashable α] (h : m.WF) {l : List ((a : α) × β a)} {k : α} : - (m.insertMany l).contains k → (l.map Sigma.fst).contains k = false → m.contains k := by + k ∈ (m.insertMany l) → (l.map Sigma.fst).contains k = false → k ∈ m := by + simp only [mem_iff_contains] simp_to_raw using Raw₀.contains_of_contains_insertMany_list theorem get?_insertMany_list_of_contains_eq_false [LawfulBEq α] (h : m.WF) @@ -1104,7 +1105,7 @@ theorem get_insertMany_list_of_contains_eq_false [LawfulBEq α] (h : m.WF) (contains_eq_false : (l.map Sigma.fst).contains k = false) {h'} : (m.insertMany l).get k h' = - m.get k (contains_of_contains_insertMany_list h h' contains_eq_false) := by + m.get k (mem_of_mem_insertMany_list h h' contains_eq_false) := by simp_to_raw using Raw₀.get_insertMany_list_of_contains_eq_false theorem get_insertMany_list_of_mem [LawfulBEq α] (h : m.WF) @@ -1160,7 +1161,7 @@ theorem getKey_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashabl (contains_eq_false : (l.map Sigma.fst).contains k = false) {h'} : (m.insertMany l).getKey k h' = - m.getKey k (contains_of_contains_insertMany_list h h' contains_eq_false) := by + m.getKey k (mem_of_mem_insertMany_list h h' contains_eq_false) := by simp_to_raw using Raw₀.getKey_insertMany_list_of_contains_eq_false theorem getKey_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] (h : m.WF) @@ -1252,9 +1253,10 @@ theorem mem_insertMany_list [EquivBEq α] [LawfulHashable α] (h : m.WF) k ∈ (insertMany m l) ↔ k ∈ m ∨ (l.map Prod.fst).contains k := by simp [mem_iff_contains, contains_insertMany_list h] -theorem contains_of_contains_insertMany_list [EquivBEq α] [LawfulHashable α] (h : m.WF) +theorem mem_of_mem_insertMany_list [EquivBEq α] [LawfulHashable α] (h : m.WF) {l : List ( α × β )} {k : α} : - (insertMany m l).contains k → (l.map Prod.fst).contains k = false → m.contains k := by + k ∈ (insertMany m l) → (l.map Prod.fst).contains k = false → k ∈ m := by + simp only [mem_iff_contains] simp_to_raw using Raw₀.Const.contains_of_contains_insertMany_list theorem getKey?_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] (h : m.WF) @@ -1276,7 +1278,7 @@ theorem getKey_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashabl (contains_eq_false : (l.map Prod.fst).contains k = false) {h'} : (insertMany m l).getKey k h' = - m.getKey k (contains_of_contains_insertMany_list h h' contains_eq_false) := by + m.getKey k (mem_of_mem_insertMany_list h h' contains_eq_false) := by simp_to_raw using Raw₀.Const.getKey_insertMany_list_of_contains_eq_false theorem getKey_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] (h : m.WF) @@ -1351,7 +1353,7 @@ theorem get_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable (contains_eq_false : (l.map Prod.fst).contains k = false) {h'} : get (insertMany m l) k h' = - get m k (contains_of_contains_insertMany_list h h' contains_eq_false) := by + get m k (mem_of_mem_insertMany_list h h' contains_eq_false) := by simp_to_raw using Raw₀.Const.get_insertMany_list_of_contains_eq_false theorem get_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] (h : m.WF) @@ -1415,9 +1417,10 @@ theorem mem_insertManyIfNewUnit_list [EquivBEq α] [LawfulHashable α] (h : m.WF k ∈ insertManyIfNewUnit m l ↔ k ∈ m ∨ l.contains k := by simp [mem_iff_contains, contains_insertManyIfNewUnit_list h] -theorem contains_of_contains_insertManyIfNewUnit_list [EquivBEq α] [LawfulHashable α] (h : m.WF) +theorem mem_of_mem_insertManyIfNewUnit_list [EquivBEq α] [LawfulHashable α] (h : m.WF) {l : List α} {k : α} (contains_eq_false : l.contains k = false) : - (insertManyIfNewUnit m l).contains k → m.contains k := by + k ∈ (insertManyIfNewUnit m l) → k ∈ m := by + simp only [mem_iff_contains] simp_to_raw using Raw₀.Const.contains_of_contains_insertManyIfNewUnit_list theorem getKey?_insertManyIfNewUnit_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] @@ -1441,7 +1444,7 @@ theorem getKey?_insertManyIfNewUnit_list_of_contains_of_contains [EquivBEq α] [ theorem getKey_insertManyIfNewUnit_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] (h : m.WF) {l : List α} {k : α} (contains_eq_false : l.contains k = false) {h'} : getKey (insertManyIfNewUnit m l) k h' = - getKey m k (contains_of_contains_insertManyIfNewUnit_list h contains_eq_false h') := by + getKey m k (mem_of_mem_insertManyIfNewUnit_list h contains_eq_false h') := by simp_to_raw using Raw₀.Const.getKey_insertManyIfNewUnit_list_of_contains_eq_false theorem getKey_insertManyIfNewUnit_list_of_mem_of_contains_eq_false [EquivBEq α] [LawfulHashable α] @@ -1551,6 +1554,7 @@ namespace Raw variable [BEq α] [Hashable α] open Internal.Raw Internal.Raw₀ + @[simp] theorem ofList_nil [EquivBEq α] [LawfulHashable α] : ofList ([] : List ((a : α) × (β a))) = ∅ := by diff --git a/src/Std/Data/HashMap/Lemmas.lean b/src/Std/Data/HashMap/Lemmas.lean index 88b42ccda87e..83a6faa720fa 100644 --- a/src/Std/Data/HashMap/Lemmas.lean +++ b/src/Std/Data/HashMap/Lemmas.lean @@ -724,44 +724,44 @@ theorem mem_insertMany_list [EquivBEq α] [LawfulHashable α] k ∈ insertMany m l ↔ k ∈ m ∨ (l.map Prod.fst).contains k := DHashMap.Const.mem_insertMany_list -theorem contains_of_contains_insertMany_list [EquivBEq α] [LawfulHashable α] - {l : List ( α × β )} {k : α} (contains : (insertMany m l).contains k) +theorem mem_of_mem_insertMany_list [EquivBEq α] [LawfulHashable α] + {l : List (α × β)} {k : α} (mem : k ∈ (insertMany m l)) (contains_eq_false : (l.map Prod.fst).contains k = false) : - m.contains k := - DHashMap.Const.contains_of_contains_insertMany_list contains contains_eq_false + k ∈ m := + DHashMap.Const.mem_of_mem_insertMany_list mem contains_eq_false -theorem get?_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] +theorem getElem?_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] {l : List (α × β)} {k : α} (contains_eq_false : (l.map Prod.fst).contains k = false) : (insertMany m l)[k]? = m[k]? := DHashMap.Const.get?_insertMany_list_of_contains_eq_false contains_eq_false -theorem get?_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] +theorem getElem?_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] {l : List (α × β)} {k k' : α} (k_beq : k == k') {v : β} (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) (mem : ⟨k, v⟩ ∈ l) : (insertMany m l)[k']? = some v := DHashMap.Const.get?_insertMany_list_of_mem k_beq distinct mem -theorem get_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] +theorem getElem_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] {l : List (α × β)} {k : α} (contains_eq_false : (l.map Prod.fst).contains k = false) {h} : - (insertMany m l)[k] = m[k]'(contains_of_contains_insertMany_list h contains_eq_false) := + (insertMany m l)[k] = m[k]'(mem_of_mem_insertMany_list h contains_eq_false) := DHashMap.Const.get_insertMany_list_of_contains_eq_false contains_eq_false -theorem get_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] +theorem getElem_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] {l : List (α × β)} {k k' : α} (k_beq : k == k') {v : β} (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) (mem : ⟨k, v⟩ ∈ l) {h} : (insertMany m l)[k'] = v := DHashMap.Const.get_insertMany_list_of_mem k_beq distinct mem -theorem get!_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] +theorem getElem!_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] [Inhabited β] {l : List (α × β)} {k : α} (contains_eq_false : (l.map Prod.fst).contains k = false) : (insertMany m l)[k]! = m[k]! := DHashMap.Const.get!_insertMany_list_of_contains_eq_false contains_eq_false -theorem get!_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] [Inhabited β] +theorem getElem!_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] [Inhabited β] {l : List (α × β)} {k k' : α} (k_beq : k == k') {v : β} (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) (mem : ⟨k, v⟩ ∈ l) : (insertMany m l)[k']! = v := @@ -798,7 +798,7 @@ theorem getKey_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashabl (contains_eq_false : (l.map Prod.fst).contains k = false) {h} : (insertMany m l).getKey k h = - m.getKey k (contains_of_contains_insertMany_list h contains_eq_false) := + m.getKey k (mem_of_mem_insertMany_list h contains_eq_false) := DHashMap.Const.getKey_insertMany_list_of_contains_eq_false contains_eq_false theorem getKey_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] @@ -884,23 +884,23 @@ theorem mem_insertManyIfNewUnit_list [EquivBEq α] [LawfulHashable α] k ∈ insertManyIfNewUnit m l ↔ k ∈ m ∨ l.contains k := DHashMap.Const.mem_insertManyIfNewUnit_list -theorem contains_of_contains_insertManyIfNewUnit_list [EquivBEq α] [LawfulHashable α] +theorem mem_of_mem_insertManyIfNewUnit_list [EquivBEq α] [LawfulHashable α] {l : List α} {k : α} (contains_eq_false : l.contains k = false) : - (insertManyIfNewUnit m l).contains k → m.contains k := - DHashMap.Const.contains_of_contains_insertManyIfNewUnit_list contains_eq_false + k ∈ (insertManyIfNewUnit m l) → k ∈ m := + DHashMap.Const.mem_of_mem_insertManyIfNewUnit_list contains_eq_false -theorem get?_insertManyIfNewUnit_list [EquivBEq α] [LawfulHashable α] +theorem getElem?_insertManyIfNewUnit_list [EquivBEq α] [LawfulHashable α] {l : List α} {k : α} : (insertManyIfNewUnit m l)[k]? = if m.contains k ∨ l.contains k then some () else none := DHashMap.Const.get?_insertManyIfNewUnit_list -theorem get_insertManyIfNewUnit_list +theorem getElem_insertManyIfNewUnit_list {l : List α} {k : α} {h} : (insertManyIfNewUnit m l)[k] = () := DHashMap.Const.get_insertManyIfNewUnit_list -theorem get!_insertManyIfNewUnit_list +theorem getElem!_insertManyIfNewUnit_list {l : List α} {k : α} : (insertManyIfNewUnit m l)[k]! = () := DHashMap.Const.get!_insertManyIfNewUnit_list @@ -933,7 +933,7 @@ theorem getKey?_insertManyIfNewUnit_list_of_contains_of_contains [EquivBEq α] [ theorem getKey_insertManyIfNewUnit_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] {l : List α} {k : α} (contains_eq_false : l.contains k = false) {h} : getKey (insertManyIfNewUnit m l) k h = - getKey m k (contains_of_contains_insertManyIfNewUnit_list contains_eq_false h) := + getKey m k (mem_of_mem_insertManyIfNewUnit_list contains_eq_false h) := DHashMap.Const.getKey_insertManyIfNewUnit_list_of_contains_eq_false contains_eq_false theorem getKey_insertManyIfNewUnit_list_of_mem_of_contains_eq_false [EquivBEq α] [LawfulHashable α] @@ -1040,20 +1040,20 @@ theorem mem_ofList [EquivBEq α] [LawfulHashable α] k ∈ ofList l ↔ (l.map Prod.fst).contains k := DHashMap.Const.mem_ofList -theorem get?_ofList_of_contains_eq_false [LawfulBEq α] +theorem getElem?_ofList_of_contains_eq_false [LawfulBEq α] {l : List (α × β)} {k : α} (contains_eq_false : (l.map Prod.fst).contains k = false) : (ofList l)[k]? = none := DHashMap.Const.get?_ofList_of_contains_eq_false contains_eq_false -theorem get?_ofList_of_mem [LawfulBEq α] +theorem getElem?_ofList_of_mem [LawfulBEq α] {l : List (α × β)} {k k' : α} (k_beq : k == k') {v : β} (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) (mem : ⟨k, v⟩ ∈ l) : (ofList l)[k']? = some v := DHashMap.Const.get?_ofList_of_mem k_beq distinct mem -theorem get_ofList_of_mem [LawfulBEq α] +theorem getElem_ofList_of_mem [LawfulBEq α] {l : List (α × β)} {k k' : α} (k_beq : k == k') {v : β} (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) (mem : ⟨k, v⟩ ∈ l) @@ -1061,13 +1061,13 @@ theorem get_ofList_of_mem [LawfulBEq α] (ofList l)[k'] = v := DHashMap.Const.get_ofList_of_mem k_beq distinct mem -theorem get!_ofList_of_contains_eq_false [LawfulBEq α] +theorem getElem!_ofList_of_contains_eq_false [LawfulBEq α] {l : List (α × β)} {k : α} [Inhabited β] (contains_eq_false : (l.map Prod.fst).contains k = false) : (ofList l)[k]! = (default : β) := DHashMap.Const.get!_ofList_of_contains_eq_false contains_eq_false -theorem get!_ofList_of_mem [LawfulBEq α] +theorem getElem!_ofList_of_mem [LawfulBEq α] {l : List (α × β)} {k k' : α} (k_beq : k == k') {v : β} [Inhabited β] (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) (mem : ⟨k, v⟩ ∈ l) : @@ -1182,20 +1182,20 @@ theorem mem_unitOfList [EquivBEq α] [LawfulHashable α] DHashMap.Const.mem_unitOfList @[simp] -theorem get?_unitOfList [EquivBEq α] [LawfulHashable α] +theorem getElem?_unitOfList [EquivBEq α] [LawfulHashable α] {l : List α} {k : α} : (unitOfList l)[k]? = if l.contains k then some () else none := DHashMap.Const.get?_unitOfList @[simp] -theorem get_unitOfList +theorem getElem_unitOfList {l : List α} {k : α} {h} : (unitOfList l)[k] = () := DHashMap.Const.get_unitOfList @[simp] -theorem get!_unitOfList +theorem getElem!_unitOfList {l : List α} {k : α} : (unitOfList l)[k]! = () := DHashMap.Const.get!_unitOfList diff --git a/src/Std/Data/HashMap/RawLemmas.lean b/src/Std/Data/HashMap/RawLemmas.lean index 4968d71ec930..d2f5c5b3672e 100644 --- a/src/Std/Data/HashMap/RawLemmas.lean +++ b/src/Std/Data/HashMap/RawLemmas.lean @@ -733,10 +733,10 @@ theorem mem_insertMany_list [EquivBEq α] [LawfulHashable α] (h : m.WF) k ∈ (insertMany m l) ↔ k ∈ m ∨ (l.map Prod.fst).contains k := DHashMap.Raw.Const.mem_insertMany_list h.out -theorem contains_of_contains_insertMany_list [EquivBEq α] [LawfulHashable α] (h : m.WF) +theorem mem_of_mem_insertMany_list [EquivBEq α] [LawfulHashable α] (h : m.WF) {l : List ( α × β )} {k : α} : - (insertMany m l).contains k → (l.map Prod.fst).contains k = false → m.contains k := - DHashMap.Raw.Const.contains_of_contains_insertMany_list h.out + k ∈ (insertMany m l) → (l.map Prod.fst).contains k = false → k ∈ m := + DHashMap.Raw.Const.mem_of_mem_insertMany_list h.out theorem getKey?_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] (h : m.WF) {l : List (α × β)} {k : α} @@ -757,7 +757,7 @@ theorem getKey_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashabl (contains_eq_false : (l.map Prod.fst).contains k = false) {h'} : (insertMany m l).getKey k h' = - m.getKey k (contains_of_contains_insertMany_list h h' contains_eq_false) := + m.getKey k (mem_of_mem_insertMany_list h h' contains_eq_false) := DHashMap.Raw.Const.getKey_insertMany_list_of_contains_eq_false h.out contains_eq_false theorem getKey_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] @@ -815,39 +815,39 @@ theorem isEmpty_insertMany_list [EquivBEq α] [LawfulHashable α] (insertMany m l).isEmpty = (m.isEmpty && l.isEmpty) := DHashMap.Raw.Const.isEmpty_insertMany_list h.out -theorem get?_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] +theorem getElem?_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] (h : m.WF) {l : List (α × β)} {k : α} (contains_eq_false : (l.map Prod.fst).contains k = false) : (insertMany m l)[k]? = m[k]? := DHashMap.Raw.Const.get?_insertMany_list_of_contains_eq_false h.out contains_eq_false -theorem get?_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] +theorem getElem?_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] (h : m.WF) {l : List (α × β)} {k k' : α} (k_beq : k == k') {v : β} (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) (mem : ⟨k, v⟩ ∈ l) : (insertMany m l)[k']? = v := DHashMap.Raw.Const.get?_insertMany_list_of_mem h.out k_beq distinct mem -theorem get_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] +theorem getElem_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] (h : m.WF) {l : List (α × β)} {k : α} (contains_eq_false : (l.map Prod.fst).contains k = false) {h'} : (insertMany m l)[k] = - m[k]'(contains_of_contains_insertMany_list h h' contains_eq_false) := + m[k]'(mem_of_mem_insertMany_list h h' contains_eq_false) := DHashMap.Raw.Const.get_insertMany_list_of_contains_eq_false h.out contains_eq_false (h':= h') -theorem get_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] +theorem getElem_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] (h : m.WF) {l : List (α × β)} {k k' : α} (k_beq : k == k') {v : β} (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) (mem : ⟨k, v⟩ ∈ l) {h'} : (insertMany m l)[k'] = v := DHashMap.Raw.Const.get_insertMany_list_of_mem h.out k_beq distinct mem (h' := h') -theorem get!_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] +theorem getElem!_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] [Inhabited β] (h : m.WF) {l : List (α × β)} {k : α} (contains_eq_false : (l.map Prod.fst).contains k = false) : (insertMany m l)[k]! = m[k]! := DHashMap.Raw.Const.get!_insertMany_list_of_contains_eq_false h.out contains_eq_false -theorem get!_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] [Inhabited β] +theorem getElem!_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] [Inhabited β] (h : m.WF) {l : List (α × β)} {k k' : α} (k_beq : k == k') {v : β} (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) (mem : ⟨k, v⟩ ∈ l) : (insertMany m l)[k']! = v := @@ -893,10 +893,10 @@ theorem mem_insertManyIfNewUnit_list [EquivBEq α] [LawfulHashable α] (h : m.WF k ∈ insertManyIfNewUnit m l ↔ k ∈ m ∨ l.contains k := DHashMap.Raw.Const.mem_insertManyIfNewUnit_list h.out -theorem contains_of_contains_insertManyIfNewUnit_list [EquivBEq α] [LawfulHashable α] (h : m.WF) +theorem mem_of_mem_insertManyIfNewUnit_list [EquivBEq α] [LawfulHashable α] (h : m.WF) {l : List α} {k : α} (contains_eq_false : l.contains k = false) : - (insertManyIfNewUnit m l).contains k → m.contains k := - DHashMap.Raw.Const.contains_of_contains_insertManyIfNewUnit_list h.out contains_eq_false + k ∈ (insertManyIfNewUnit m l) → k ∈ m := + DHashMap.Raw.Const.mem_of_mem_insertManyIfNewUnit_list h.out contains_eq_false theorem getKey?_insertManyIfNewUnit_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] (h : m.WF) {l : List α} {k : α} (contains_eq_false : l.contains k = false) : @@ -920,7 +920,7 @@ theorem getKey?_insertManyIfNewUnit_list_of_contains_of_contains [EquivBEq α] [ theorem getKey_insertManyIfNewUnit_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] (h : m.WF) {l : List α} {k : α} (contains_eq_false : l.contains k = false) {h'} : getKey (insertManyIfNewUnit m l) k h' = - getKey m k (contains_of_contains_insertManyIfNewUnit_list h contains_eq_false h') := + getKey m k (mem_of_mem_insertManyIfNewUnit_list h contains_eq_false h') := DHashMap.Raw.Const.getKey_insertManyIfNewUnit_list_of_contains_eq_false h.out contains_eq_false theorem getKey_insertManyIfNewUnit_list_of_mem_of_contains_eq_false [EquivBEq α] [LawfulHashable α] @@ -1001,20 +1001,20 @@ theorem isEmpty_insertManyIfNewUnit_list [EquivBEq α] [LawfulHashable α] (h : DHashMap.Raw.Const.isEmpty_insertManyIfNewUnit_list h.out @[simp] -theorem get?_insertManyIfNewUnit_list [EquivBEq α] [LawfulHashable α] (h : m.WF) +theorem getElem?_insertManyIfNewUnit_list [EquivBEq α] [LawfulHashable α] (h : m.WF) {l : List α} {k : α} : (insertManyIfNewUnit m l)[k]? = if m.contains k ∨ l.contains k then some () else none := DHashMap.Raw.Const.get?_insertManyIfNewUnit_list h.out @[simp] -theorem get_insertManyIfNewUnit_list +theorem ggetElem_insertManyIfNewUnit_list {l : List α} {k : α} {h} : (insertManyIfNewUnit m l)[k] = () := DHashMap.Raw.Const.get_insertManyIfNewUnit_list (h:=h) @[simp] -theorem get!_insertManyIfNewUnit_list +theorem getElem!_insertManyIfNewUnit_list {l : List α} {k : α} : (insertManyIfNewUnit m l)[k]! = () := DHashMap.Raw.Const.get!_insertManyIfNewUnit_list @@ -1057,20 +1057,20 @@ theorem mem_ofList [EquivBEq α] [LawfulHashable α] k ∈ (ofList l) ↔ (l.map Prod.fst).contains k := DHashMap.Raw.Const.mem_ofList -theorem get?_ofList_of_contains_eq_false [LawfulBEq α] +theorem getElem?_ofList_of_contains_eq_false [LawfulBEq α] {l : List (α × β)} {k : α} (contains_eq_false : (l.map Prod.fst).contains k = false) : (ofList l)[k]? = none := DHashMap.Raw.Const.get?_ofList_of_contains_eq_false contains_eq_false -theorem get?_ofList_of_mem [LawfulBEq α] +theorem getElem?_ofList_of_mem [LawfulBEq α] {l : List (α × β)} {k k' : α} (k_beq : k == k') {v : β} (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) (mem : ⟨k, v⟩ ∈ l) : (ofList l)[k']? = some v := DHashMap.Raw.Const.get?_ofList_of_mem k_beq distinct mem -theorem get_ofList_of_mem [LawfulBEq α] +theorem getElem_ofList_of_mem [LawfulBEq α] {l : List (α × β)} {k k' : α} (k_beq : k == k') {v : β} (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) (mem : ⟨k, v⟩ ∈ l) @@ -1078,13 +1078,13 @@ theorem get_ofList_of_mem [LawfulBEq α] (ofList l)[k'] = v := DHashMap.Raw.Const.get_ofList_of_mem k_beq distinct mem (h:=h) -theorem get!_ofList_of_contains_eq_false [LawfulBEq α] +theorem getElem!_ofList_of_contains_eq_false [LawfulBEq α] {l : List (α × β)} {k : α} [Inhabited β] (contains_eq_false : (l.map Prod.fst).contains k = false) : (ofList l)[k]! = default := DHashMap.Raw.Const.get!_ofList_of_contains_eq_false contains_eq_false -theorem get!_ofList_of_mem [LawfulBEq α] +theorem getElem!_ofList_of_mem [LawfulBEq α] {l : List (α × β)} {k k' : α} (k_beq : k == k') {v : β} [Inhabited β] (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) (mem : ⟨k, v⟩ ∈ l) : @@ -1260,20 +1260,20 @@ theorem isEmpty_unitOfList [EquivBEq α] [LawfulHashable α] DHashMap.Raw.Const.isEmpty_unitOfList @[simp] -theorem get?_unitOfList [EquivBEq α] [LawfulHashable α] +theorem getElem?_unitOfList [EquivBEq α] [LawfulHashable α] {l : List α} {k : α} : (unitOfList l)[k]? = if l.contains k then some () else none := DHashMap.Raw.Const.get?_unitOfList @[simp] -theorem get_unitOfList +theorem getElem_unitOfList {l : List α} {k : α} {h} : (unitOfList l)[k] = () := DHashMap.Raw.Const.get_unitOfList (h:=h) @[simp] -theorem get!_unitOfList +theorem getElem!_unitOfList {l : List α} {k : α} : (unitOfList l)[k]! = () := DHashMap.Raw.Const.get!_unitOfList diff --git a/src/Std/Data/HashSet/Lemmas.lean b/src/Std/Data/HashSet/Lemmas.lean index 6900e62c7283..340689731f10 100644 --- a/src/Std/Data/HashSet/Lemmas.lean +++ b/src/Std/Data/HashSet/Lemmas.lean @@ -400,10 +400,10 @@ theorem mem_insertMany_list [EquivBEq α] [LawfulHashable α] k ∈ insertMany m l ↔ k ∈ m ∨ l.contains k := HashMap.mem_insertManyIfNewUnit_list -theorem contains_of_contains_insertMany_list [EquivBEq α] [LawfulHashable α] +theorem mem_of_mem_insertMany_list [EquivBEq α] [LawfulHashable α] {l : List α} {k : α} (contains_eq_false : l.contains k = false) : - (insertMany m l).contains k → m.contains k := - HashMap.contains_of_contains_insertManyIfNewUnit_list contains_eq_false + k ∈ (insertMany m l) → k ∈ m := + HashMap.mem_of_mem_insertManyIfNewUnit_list contains_eq_false theorem get?_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] {l : List α} {k : α} (contains_eq_false : l.contains k = false) : @@ -428,7 +428,7 @@ theorem get?_insertMany_list_of_contains_of_contains [EquivBEq α] [LawfulHashab theorem get_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] {l : List α} {k : α} (contains_eq_false : l.contains k = false) {h} : get (insertMany m l) k h = - get m k (contains_of_contains_insertMany_list contains_eq_false h) := + get m k (mem_of_mem_insertMany_list contains_eq_false h) := HashMap.getKey_insertManyIfNewUnit_list_of_contains_eq_false contains_eq_false theorem get_insertMany_list_of_mem_of_contains_eq_false [EquivBEq α] [LawfulHashable α] diff --git a/src/Std/Data/HashSet/RawLemmas.lean b/src/Std/Data/HashSet/RawLemmas.lean index 0a82d686144f..f86a52f78a3d 100644 --- a/src/Std/Data/HashSet/RawLemmas.lean +++ b/src/Std/Data/HashSet/RawLemmas.lean @@ -416,10 +416,10 @@ theorem mem_insertMany_list [EquivBEq α] [LawfulHashable α] (h : m.WF) k ∈ insertMany m l ↔ k ∈ m ∨ l.contains k := HashMap.Raw.mem_insertManyIfNewUnit_list h.1 -theorem contains_of_contains_insertMany_list [EquivBEq α] [LawfulHashable α] (h : m.WF) +theorem mem_of_mem_insertMany_list [EquivBEq α] [LawfulHashable α] (h : m.WF) {l : List α} {k : α} (contains_eq_false : l.contains k = false) : - (insertMany m l).contains k → m.contains k := - HashMap.Raw.contains_of_contains_insertManyIfNewUnit_list h.1 contains_eq_false + k ∈ (insertMany m l) → k ∈ m := + HashMap.Raw.mem_of_mem_insertManyIfNewUnit_list h.1 contains_eq_false theorem get?_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] (h : m.WF) {l : List α} {k : α} (contains_eq_false : l.contains k = false) : @@ -442,7 +442,7 @@ theorem get?_insertMany_list_of_contains_of_contains [EquivBEq α] [LawfulHashab theorem get_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] (h : m.WF) {l : List α} {k : α} (contains_eq_false : l.contains k = false) {h'} : get (insertMany m l) k h' = - get m k (contains_of_contains_insertMany_list h contains_eq_false h') := + get m k (mem_of_mem_insertMany_list h contains_eq_false h') := HashMap.Raw.getKey_insertManyIfNewUnit_list_of_contains_eq_false h.1 contains_eq_false theorem get_insertMany_list_of_mem_of_contains_eq_false [EquivBEq α] [LawfulHashable α] From 1360be3d6fb3c18278ccbda3496ca8fb5cee506d Mon Sep 17 00:00:00 2001 From: jt0202 Date: Fri, 10 Jan 2025 14:04:47 +0100 Subject: [PATCH 75/83] Golf associative --- .../DHashMap/Internal/List/Associative.lean | 477 ++++++------------ 1 file changed, 141 insertions(+), 336 deletions(-) diff --git a/src/Std/Data/DHashMap/Internal/List/Associative.lean b/src/Std/Data/DHashMap/Internal/List/Associative.lean index a948706d35e0..91dd005125fe 100644 --- a/src/Std/Data/DHashMap/Internal/List/Associative.lean +++ b/src/Std/Data/DHashMap/Internal/List/Associative.lean @@ -1890,11 +1890,7 @@ theorem insertList_perm_of_perm_first [BEq α] [EquivBEq α] {l1 l2 toInsert : L | nil => simp [insertList, h] | cons hd tl ih => simp only [insertList] - apply ih - apply insertEntry_of_perm - exact distinct - exact h - apply DistinctKeys.insertEntry distinct + apply ih (insertEntry_of_perm distinct h) (DistinctKeys.insertEntry distinct) theorem insertList_cons_perm [BEq α] [EquivBEq α] {l₁ l₂ : List ((a : α) × β a)} {p : (a : α) × β a} (hl₁ : DistinctKeys l₁) (hl₂ : DistinctKeys (p :: l₂)) : @@ -1985,7 +1981,7 @@ theorem getValueCast?_insertList_of_contains_eq_false [BEq α] [LawfulBEq α] theorem getValueCast?_insertList_of_mem [BEq α] [LawfulBEq α] {l toInsert : List ((a : α) × β a)} - {k k' : α} (k_beq : k == k') (v : β k) + {k k' : α} (k_beq : k == k') {v : β k} (distinct_l : DistinctKeys l) (distinct_toInsert : toInsert.Pairwise (fun a b => (a.1 == b.1) = false)) (mem : ⟨k, v⟩ ∈ toInsert) : @@ -2007,10 +2003,8 @@ theorem getValueCast_insertList_of_contains_eq_false [BEq α] [LawfulBEq α] {h} : getValueCast k (insertList l toInsert) h = getValueCast k l (containsKey_of_containsKey_insertList h not_contains) := by - rw [← Option.some_inj] - rw [← getValueCast?_eq_some_getValueCast] - rw [← getValueCast?_eq_some_getValueCast] - rw [getValueCast?_insertList_of_contains_eq_false] + rw [← Option.some_inj, ← getValueCast?_eq_some_getValueCast, ← getValueCast?_eq_some_getValueCast, + getValueCast?_insertList_of_contains_eq_false] exact not_contains theorem getValueCast_insertList_of_mem [BEq α] [LawfulBEq α] @@ -2022,21 +2016,15 @@ theorem getValueCast_insertList_of_mem [BEq α] [LawfulBEq α] {h} : getValueCast k' (insertList l toInsert) h = cast (by congr; exact LawfulBEq.eq_of_beq k_beq) v := by - rw [← Option.some_inj] - rw [← getValueCast?_eq_some_getValueCast] - rw [getValueCast?_insertList_of_mem] - . exact k_beq - . exact distinct_l - . exact distinct_toInsert - . exact mem + rw [← Option.some_inj, ← getValueCast?_eq_some_getValueCast, + getValueCast?_insertList_of_mem k_beq distinct_l distinct_toInsert mem] theorem getValueCast!_insertList_of_contains_eq_false [BEq α] [LawfulBEq α] {l toInsert : List ((a : α) × β a)} {k : α} [Inhabited (β k)] (not_contains : (toInsert.map Sigma.fst).contains k = false) : getValueCast! k (insertList l toInsert) = getValueCast! k l := by rw [getValueCast!_eq_getValueCast?, getValueCast!_eq_getValueCast?, - getValueCast?_insertList_of_contains_eq_false] - exact not_contains + getValueCast?_insertList_of_contains_eq_false not_contains] theorem getValueCast!_insertList_of_mem [BEq α] [LawfulBEq α] {l toInsert : List ((a : α) × β a)} @@ -2046,19 +2034,15 @@ theorem getValueCast!_insertList_of_mem [BEq α] [LawfulBEq α] (mem : ⟨k, v⟩ ∈ toInsert) : getValueCast! k' (insertList l toInsert) = cast (by congr; exact LawfulBEq.eq_of_beq k_beq) v := by - rw [getValueCast!_eq_getValueCast?, getValueCast?_insertList_of_mem] - . rw [Option.get!_some]; exact k_beq - . exact distinct_l - . exact distinct_toInsert - . exact mem + rw [getValueCast!_eq_getValueCast?, + getValueCast?_insertList_of_mem k_beq distinct_l distinct_toInsert mem, Option.get!_some] theorem getValueCastD_insertList_of_contains_eq_false [BEq α] [LawfulBEq α] {l toInsert : List ((a : α) × β a)} {k : α} {fallback : β k} (not_mem : (toInsert.map Sigma.fst).contains k = false) : getValueCastD k (insertList l toInsert) fallback = getValueCastD k l fallback := by rw [getValueCastD_eq_getValueCast?, getValueCastD_eq_getValueCast?, - getValueCast?_insertList_of_contains_eq_false] - exact not_mem + getValueCast?_insertList_of_contains_eq_false not_mem] theorem getValueCastD_insertList_of_mem [BEq α] [LawfulBEq α] {l toInsert : List ((a : α) × β a)} @@ -2068,19 +2052,16 @@ theorem getValueCastD_insertList_of_mem [BEq α] [LawfulBEq α] (mem : ⟨k, v⟩ ∈ toInsert) : getValueCastD k' (insertList l toInsert) fallback = cast (by congr; exact LawfulBEq.eq_of_beq k_beq) v := by - rw [getValueCastD_eq_getValueCast?, getValueCast?_insertList_of_mem] - . rw [Option.getD_some]; exact k_beq - . exact distinct_l - . exact distinct_toInsert - . exact mem + rw [getValueCastD_eq_getValueCast?, + getValueCast?_insertList_of_mem k_beq distinct_l distinct_toInsert mem, Option.getD_some] theorem getKey?_insertList_of_contains_eq_false [BEq α] [EquivBEq α] {l toInsert : List ((a : α) × β a)} {k : α} (not_contains : (toInsert.map Sigma.fst).contains k = false) : getKey? k (insertList l toInsert) = getKey? k l := by rw [← containsKey_eq_contains_map_fst] at not_contains - rw [getKey?_eq_getEntry?, getKey?_eq_getEntry?] - rw [getEntry?_insertList_of_contains_eq_false not_contains] + rw [getKey?_eq_getEntry?, getKey?_eq_getEntry?, + getEntry?_insertList_of_contains_eq_false not_contains] theorem getKey?_insertList_of_mem [BEq α] [EquivBEq α] {l toInsert : List ((a : α) × β a)} @@ -2100,11 +2081,8 @@ theorem getKey_insertList_of_contains_eq_false [BEq α] [EquivBEq α] {h} : getKey k (insertList l toInsert) h = getKey k l (containsKey_of_containsKey_insertList h not_contains) := by - rw [← Option.some_inj] - rw [← getKey?_eq_some_getKey] - rw [getKey?_insertList_of_contains_eq_false] - . rw [getKey?_eq_some_getKey] - . exact not_contains + rw [← Option.some_inj, ← getKey?_eq_some_getKey, + getKey?_insertList_of_contains_eq_false not_contains, getKey?_eq_some_getKey] theorem getKey_insertList_of_mem [BEq α] [EquivBEq α] {l toInsert : List ((a : α) × β a)} @@ -2114,22 +2092,14 @@ theorem getKey_insertList_of_mem [BEq α] [EquivBEq α] (mem : k ∈ toInsert.map Sigma.fst) {h} : getKey k' (insertList l toInsert) h = k := by - rw [← Option.some_inj] - rw [← getKey?_eq_some_getKey] - rw [getKey?_insertList_of_mem] - . exact k_beq - . exact distinct_l - . exact distinct_toInsert - . exact mem + rw [← Option.some_inj, ← getKey?_eq_some_getKey, + getKey?_insertList_of_mem k_beq distinct_l distinct_toInsert mem] theorem getKey!_insertList_of_contains_eq_false [BEq α] [EquivBEq α] [Inhabited α] {l toInsert : List ((a : α) × β a)} {k : α} (contains_false : (toInsert.map Sigma.fst).contains k = false) : getKey! k (insertList l toInsert) = getKey! k l := by - rw [getKey!_eq_getKey?] - rw [getKey?_insertList_of_contains_eq_false] - . rw [getKey!_eq_getKey?] - . exact contains_false + rw [getKey!_eq_getKey?, getKey?_insertList_of_contains_eq_false contains_false, getKey!_eq_getKey?] theorem getKey!_insertList_of_mem [BEq α] [EquivBEq α] [Inhabited α] {l toInsert : List ((a : α) × β a)} @@ -2138,22 +2108,14 @@ theorem getKey!_insertList_of_mem [BEq α] [EquivBEq α] [Inhabited α] (distinct_toInsert : toInsert.Pairwise (fun a b => (a.1 == b.1) = false)) (mem : k ∈ toInsert.map Sigma.fst) : getKey! k' (insertList l toInsert) = k := by - rw [getKey!_eq_getKey?] - rw [getKey?_insertList_of_mem] - . rw [Option.get!_some] - . exact k_beq - . exact distinct_l - . exact distinct_toInsert - . exact mem + rw [getKey!_eq_getKey?, getKey?_insertList_of_mem k_beq distinct_l distinct_toInsert mem, + Option.get!_some] theorem getKeyD_insertList_of_contains_eq_false [BEq α] [EquivBEq α] {l toInsert : List ((a : α) × β a)} {k fallback : α} (not_contains : (toInsert.map Sigma.fst).contains k = false) : getKeyD k (insertList l toInsert) fallback = getKeyD k l fallback := by - rw [getKeyD_eq_getKey?] - rw [getKey?_insertList_of_contains_eq_false] - . rw [getKeyD_eq_getKey?] - . exact not_contains + rw [getKeyD_eq_getKey?, getKey?_insertList_of_contains_eq_false not_contains, getKeyD_eq_getKey?] theorem getKeyD_insertList_of_mem [BEq α] [EquivBEq α] {l toInsert : List ((a : α) × β a)} @@ -2162,13 +2124,8 @@ theorem getKeyD_insertList_of_mem [BEq α] [EquivBEq α] (distinct_toInsert : toInsert.Pairwise (fun a b => (a.1 == b.1) = false)) (mem : k ∈ toInsert.map Sigma.fst) : getKeyD k' (insertList l toInsert) fallback = k := by - rw [getKeyD_eq_getKey?] - rw [getKey?_insertList_of_mem] - . rw [Option.getD_some] - . exact k_beq - . exact distinct_l - . exact distinct_toInsert - . exact mem + rw [getKeyD_eq_getKey?, getKey?_insertList_of_mem k_beq distinct_l distinct_toInsert mem, + Option.getD_some] theorem perm_insertList [BEq α] [EquivBEq α] {l toInsert : List ((a : α) × β a)} @@ -2180,14 +2137,14 @@ theorem perm_insertList [BEq α] [EquivBEq α] induction toInsert generalizing l with | nil => simp only [insertList, List.append_nil, Perm.refl] | cons hd tl ih => - simp only [List.map_cons, List.contains_cons, Bool.or_eq_false_iff] at distinct_both - refine (insertList_cons_perm distinct_l distinct_toInsert).trans ?_ - rw [insertEntry_of_containsKey_eq_false] - · refine (Perm.cons _ (ih distinct_l (distinct_toInsert).tail ?_)).trans List.perm_middle.symm - exact fun a ha => (distinct_both a ha).2 - · simp only [containsKey_insertList, Bool.or_eq_false_iff, ← containsKey_eq_contains_map_fst] - refine ⟨Bool.not_eq_true _ ▸ fun h => ?_, distinct_toInsert.containsKey_eq_false⟩ - simpa using (distinct_both _ h).1 + simp only [List.map_cons, List.contains_cons, Bool.or_eq_false_iff] at distinct_both + refine (insertList_cons_perm distinct_l distinct_toInsert).trans ?_ + rw [insertEntry_of_containsKey_eq_false] + · refine (Perm.cons _ (ih distinct_l (distinct_toInsert).tail ?_)).trans List.perm_middle.symm + exact fun a ha => (distinct_both a ha).2 + · simp only [containsKey_insertList, Bool.or_eq_false_iff, ← containsKey_eq_contains_map_fst] + refine ⟨Bool.not_eq_true _ ▸ fun h => ?_, distinct_toInsert.containsKey_eq_false⟩ + simpa using (distinct_both _ h).1 theorem length_insertList [BEq α] [EquivBEq α] {l toInsert : List ((a : α) × β a)} @@ -2206,8 +2163,7 @@ theorem length_insertList_le [BEq α] simp only [insertList, List.length_cons] apply Nat.le_trans ih rw [Nat.add_comm tl.length 1, ← Nat.add_assoc] - apply Nat.add_le_add _ (Nat.le_refl _) - apply length_insertEntry_le + apply Nat.add_le_add length_insertEntry_le (Nat.le_refl _) theorem isEmpty_insertList [BEq α] {l toInsert : List ((a : α) × β a)} : @@ -2267,9 +2223,7 @@ theorem getKey?_insertListConst_of_mem [BEq α] [EquivBEq α] (mem : k ∈ toInsert.map Prod.fst) : getKey? k' (insertListConst l toInsert) = some k := by unfold insertListConst - apply getKey?_insertList_of_mem - · exact k_beq - · exact distinct_l + apply getKey?_insertList_of_mem k_beq distinct_l · simpa [List.pairwise_map] · simp only [List.map_map, Prod.fst_comp_toSigma] exact mem @@ -2280,11 +2234,8 @@ theorem getKey_insertListConst_of_contains_eq_false [BEq α] [EquivBEq α] {h} : getKey k (insertListConst l toInsert) h = getKey k l (containsKey_of_containsKey_insertListConst h not_contains) := by - rw [← Option.some_inj] - rw [← getKey?_eq_some_getKey] - rw [getKey?_insertListConst_of_contains_eq_false] - . rw [getKey?_eq_some_getKey] - . exact not_contains + rw [← Option.some_inj, ← getKey?_eq_some_getKey, + getKey?_insertListConst_of_contains_eq_false not_contains, getKey?_eq_some_getKey] theorem getKey_insertListConst_of_mem [BEq α] [EquivBEq α] {l : List ((_ : α) × β)} {toInsert : List (α × β)} @@ -2294,22 +2245,15 @@ theorem getKey_insertListConst_of_mem [BEq α] [EquivBEq α] (mem : k ∈ toInsert.map Prod.fst) {h} : getKey k' (insertListConst l toInsert) h = k := by - rw [← Option.some_inj] - rw [← getKey?_eq_some_getKey] - rw [getKey?_insertListConst_of_mem] - . exact k_beq - . exact distinct_l - . exact distinct_toInsert - . exact mem + rw [← Option.some_inj, ← getKey?_eq_some_getKey, + getKey?_insertListConst_of_mem k_beq distinct_l distinct_toInsert mem] theorem getKey!_insertListConst_of_contains_eq_false [BEq α] [EquivBEq α] [Inhabited α] {l : List ((_ : α) × β)} {toInsert : List (α × β)} {k : α} (not_contains : (toInsert.map Prod.fst).contains k = false) : getKey! k (insertListConst l toInsert) = getKey! k l := by - rw [getKey!_eq_getKey?] - rw [getKey?_insertListConst_of_contains_eq_false] - . rw [getKey!_eq_getKey?] - . exact not_contains + rw [getKey!_eq_getKey?, getKey?_insertListConst_of_contains_eq_false not_contains, + getKey!_eq_getKey?] theorem getKey!_insertListConst_of_mem [BEq α] [EquivBEq α] [Inhabited α] {l : List ((_ : α) × β)} {toInsert : List (α × β)} @@ -2318,22 +2262,15 @@ theorem getKey!_insertListConst_of_mem [BEq α] [EquivBEq α] [Inhabited α] (distinct_toInsert : toInsert.Pairwise (fun a b => (a.1 == b.1) = false)) (mem : k ∈ toInsert.map Prod.fst) : getKey! k' (insertListConst l toInsert) = k := by - rw [getKey!_eq_getKey?] - rw [getKey?_insertListConst_of_mem] - . rw [Option.get!_some] - . exact k_beq - . exact distinct_l - . exact distinct_toInsert - . exact mem + rw [getKey!_eq_getKey?, getKey?_insertListConst_of_mem k_beq distinct_l distinct_toInsert mem, + Option.get!_some] theorem getKeyD_insertListConst_of_contains_eq_false [BEq α] [EquivBEq α] {l : List ((_ : α) × β)} {toInsert : List (α × β)} {k fallback : α} (not_contains : (toInsert.map Prod.fst).contains k = false) : getKeyD k (insertListConst l toInsert) fallback = getKeyD k l fallback := by - rw [getKeyD_eq_getKey?] - rw [getKey?_insertListConst_of_contains_eq_false] - . rw [getKeyD_eq_getKey?] - . exact not_contains + rw [getKeyD_eq_getKey?, getKey?_insertListConst_of_contains_eq_false not_contains, + getKeyD_eq_getKey?] theorem getKeyD_insertListConst_of_mem [BEq α] [EquivBEq α] {l : List ((_ : α) × β)} {toInsert : List (α × β)} @@ -2342,13 +2279,8 @@ theorem getKeyD_insertListConst_of_mem [BEq α] [EquivBEq α] (distinct_toInsert : toInsert.Pairwise (fun a b => (a.1 == b.1) = false)) (mem : k ∈ toInsert.map Prod.fst) : getKeyD k' (insertListConst l toInsert) fallback = k := by - rw [getKeyD_eq_getKey?] - rw [getKey?_insertListConst_of_mem] - . rw [Option.getD_some] - . exact k_beq - . exact distinct_l - . exact distinct_toInsert - . exact mem + rw [getKeyD_eq_getKey?, getKey?_insertListConst_of_mem k_beq distinct_l distinct_toInsert mem, + Option.getD_some] theorem length_insertListConst [BEq α] [EquivBEq α] {l : List ((_ : α) × β)} {toInsert : List (α × β)} @@ -2357,9 +2289,8 @@ theorem length_insertListConst [BEq α] [EquivBEq α] (distinct_both : ∀ (a : α), containsKey a l → (toInsert.map Prod.fst).contains a = false) : (insertListConst l toInsert).length = l.length + toInsert.length := by unfold insertListConst - rw [length_insertList] + rw [length_insertList distinct_l] · simp - · exact distinct_l · simpa [List.pairwise_map] · simpa using distinct_both @@ -2374,8 +2305,7 @@ theorem isEmpty_insertListConst [BEq α] {l : List ((_ : α) × β)} {toInsert : List (α × β)} : (List.insertListConst l toInsert).isEmpty = (l.isEmpty && toInsert.isEmpty) := by unfold insertListConst - rw [isEmpty_insertList] - rw [Bool.eq_iff_iff] + rw [isEmpty_insertList, Bool.eq_iff_iff] simp theorem getValue?_insertListConst_of_contains_eq_false [BEq α] [PartialEquivBEq α] @@ -2383,8 +2313,7 @@ theorem getValue?_insertListConst_of_contains_eq_false [BEq α] [PartialEquivBEq (not_contains : (toInsert.map Prod.fst).contains k = false) : getValue? k (insertListConst l toInsert) = getValue? k l := by unfold insertListConst - rw [getValue?_eq_getEntry?, getValue?_eq_getEntry?] - rw [getEntry?_insertList_of_contains_eq_false] + rw [getValue?_eq_getEntry?, getValue?_eq_getEntry?, getEntry?_insertList_of_contains_eq_false] rw [containsKey_eq_contains_map_fst] simpa @@ -2398,7 +2327,7 @@ theorem getValue?_insertListConst_of_mem [BEq α] [EquivBEq α] unfold insertListConst rw [getValue?_eq_getEntry?] have : getEntry? k' (insertList l (toInsert.map Prod.toSigma)) = - getEntry? k' (toInsert.map Prod.toSigma) := by + getEntry? k' (toInsert.map Prod.toSigma) := by apply getEntry?_insertList_of_contains_eq_true distinct_l (by simpa [List.pairwise_map]) apply containsKey_of_beq _ k_beq rw [containsKey_eq_contains_map_fst, List.map_map, Prod.fst_comp_toSigma, @@ -2421,11 +2350,8 @@ theorem getValue_insertListConst_of_contains_eq_false [BEq α] [PartialEquivBEq {h} : getValue k (insertListConst l toInsert) h = getValue k l (containsKey_of_containsKey_insertListConst h not_contains) := by - rw [← Option.some_inj] - rw [← getValue?_eq_some_getValue] - rw [← getValue?_eq_some_getValue] - rw [getValue?_insertListConst_of_contains_eq_false] - exact not_contains + rw [← Option.some_inj, ← getValue?_eq_some_getValue, ← getValue?_eq_some_getValue, + getValue?_insertListConst_of_contains_eq_false not_contains] theorem getValue_insertListConst_of_mem [BEq α] [EquivBEq α] {l : List ((_ : α) × β)} {toInsert : List (α × β)} @@ -2435,20 +2361,15 @@ theorem getValue_insertListConst_of_mem [BEq α] [EquivBEq α] (mem : ⟨k, v⟩ ∈ toInsert) {h} : getValue k' (insertListConst l toInsert) h = v := by - rw [← Option.some_inj] - rw [← getValue?_eq_some_getValue] - rw [getValue?_insertListConst_of_mem (mem:=mem)] - . exact k_beq - . exact distinct_l - . exact distinct_toInsert + rw [← Option.some_inj, ← getValue?_eq_some_getValue, + getValue?_insertListConst_of_mem k_beq distinct_l distinct_toInsert mem] theorem getValue!_insertListConst_of_contains_eq_false [BEq α] [PartialEquivBEq α] [Inhabited β] {l : List ((_ : α) × β)} {toInsert : List (α × β)} {k : α} (not_contains : (toInsert.map Prod.fst).contains k = false) : getValue! k (insertListConst l toInsert) = getValue! k l := by simp only [getValue!_eq_getValue?] - rw [getValue?_insertListConst_of_contains_eq_false] - exact not_contains + rw [getValue?_insertListConst_of_contains_eq_false not_contains] theorem getValue!_insertListConst_of_mem [BEq α] [EquivBEq α] [Inhabited β] {l : List ((_ : α) × β)} {toInsert : List (α × β)} {k k' : α} {v: β} (k_beq : k == k') @@ -2456,21 +2377,15 @@ theorem getValue!_insertListConst_of_mem [BEq α] [EquivBEq α] [Inhabited β] (distinct_toInsert : toInsert.Pairwise (fun a b => (a.1 == b.1) = false)) (mem : ⟨k, v⟩ ∈ toInsert): getValue! k' (insertListConst l toInsert) = v := by - simp only [getValue!_eq_getValue?] - rw [getValue?_insertListConst_of_mem] - · rw [Option.get!_some] - · exact k_beq - · exact distinct_l - . exact distinct_toInsert - . exact mem + rw [getValue!_eq_getValue?, + getValue?_insertListConst_of_mem k_beq distinct_l distinct_toInsert mem, Option.get!_some] theorem getValueD_insertListConst_of_contains_eq_false [BEq α] [PartialEquivBEq α] {l : List ((_ : α) × β)} {toInsert : List (α × β)} {k : α} {fallback : β} (not_contains : (toInsert.map Prod.fst).contains k = false) : getValueD k (insertListConst l toInsert) fallback = getValueD k l fallback := by simp only [getValueD_eq_getValue?] - rw [getValue?_insertListConst_of_contains_eq_false] - exact not_contains + rw [getValue?_insertListConst_of_contains_eq_false not_contains] theorem getValueD_insertListConst_of_mem [BEq α] [EquivBEq α] {l : List ((_ : α) × β)} {toInsert : List (α × β)} {k k' : α} {v fallback: β} (k_beq : k == k') @@ -2479,12 +2394,7 @@ theorem getValueD_insertListConst_of_mem [BEq α] [EquivBEq α] (mem : ⟨k, v⟩ ∈ toInsert): getValueD k' (insertListConst l toInsert) fallback= v := by simp only [getValueD_eq_getValue?] - rw [getValue?_insertListConst_of_mem] - · rw [Option.getD_some] - · exact k_beq - · exact distinct_l - . exact distinct_toInsert - . exact mem + rw [getValue?_insertListConst_of_mem k_beq distinct_l distinct_toInsert mem, Option.getD_some] /-- Internal implementation detail of the hash map -/ def insertListIfNewUnit [BEq α] (l: List ((_ : α) × Unit)) (toInsert: List α) : @@ -2511,71 +2421,82 @@ theorem insertListIfNewUnit_perm_of_perm_first [BEq α] [EquivBEq α] {l1 l2 : L exact h · apply DistinctKeys.insertEntryIfNew distinct -theorem DistinctKeys.insertListIfNewUnit [BEq α] [PartialEquivBEq α] (l : List ((_ : α) × Unit)) - (toInsert : List α) (distinct: DistinctKeys l): +theorem DistinctKeys.insertListIfNewUnit [BEq α] [PartialEquivBEq α] {l : List ((_ : α) × Unit)} + {toInsert : List α} (distinct: DistinctKeys l): DistinctKeys (insertListIfNewUnit l toInsert) := by induction toInsert generalizing l with | nil => simp [List.insertListIfNewUnit, distinct] | cons hd tl ih => simp only [List.insertListIfNewUnit] - apply ih - apply insertEntryIfNew - exact distinct + apply ih (insertEntryIfNew distinct) -theorem containsKey_insertListIfNewUnit [BEq α] [PartialEquivBEq α] {l : List ((_ : α) × Unit)} +theorem getEntry?_insertListIfNewUnit [BEq α] [PartialEquivBEq α] {l : List ((_ : α) × Unit)} {toInsert : List α} {k : α} : - containsKey k (insertListIfNewUnit l toInsert) = (containsKey k l || toInsert.contains k) := by + getEntry? k (insertListIfNewUnit l toInsert) = + (getEntry? k l).or (getEntry? k (toInsert.map (⟨·, ()⟩))) := by induction toInsert generalizing l with | nil => simp [insertListIfNewUnit] | cons hd tl ih => - simp only [insertListIfNewUnit, List.contains_cons] - rw [ih, containsKey_insertEntryIfNew] - rw [Bool.or_comm (hd == k), Bool.or_assoc, BEq.comm (a:=hd)] + simp only [insertListIfNewUnit, ih, getEntry?_insertEntryIfNew, List.map_cons, + getEntry?_cons] + cases hhd : hd == k + · simp + · cases hc : containsKey hd l + · simp only [Bool.not_false, Bool.and_self, ↓reduceIte, Option.some_or, cond_true, + Option.or_some', Option.some.injEq] + rw [getEntry?_eq_none.2, Option.getD_none] + rwa [← containsKey_congr hhd] + · simp only [Bool.not_true, Bool.and_false, Bool.false_eq_true, ↓reduceIte, cond_true, + Option.or_some', getEntry?_eq_none] + rw [containsKey_congr hhd, containsKey_eq_isSome_getEntry?] at hc + obtain ⟨v, hv⟩ := Option.isSome_iff_exists.1 hc + simp [hv] + +theorem DistinctKeys.mapUnit [BEq α] + {l : List α} (distinct: l.Pairwise (fun a b => (a == b) = false)) : + DistinctKeys (l.map (⟨·, ()⟩)) := by + rw [DistinctKeys.def] + refine List.Pairwise.map ?_ ?_ distinct + simp -theorem containsKey_insertListIfNewUnit_iff [BEq α] [PartialEquivBEq α] {l : List ((_ : α) × Unit)} +theorem getEntry?_insertListIfNewUnit_of_contains_eq_false [BEq α] [PartialEquivBEq α] + {l : List ((_ : α) × Unit)} {toInsert : List α } {k : α} + (not_contains : toInsert.contains k = false) : + getEntry? k (insertListIfNewUnit l toInsert) = getEntry? k l := by + induction toInsert generalizing l with + | nil => simp [insertListIfNewUnit] + | cons h t ih => + unfold insertListIfNewUnit + simp only [List.contains_cons, Bool.or_eq_false_iff] at not_contains + rw [ih not_contains.right, getEntry?_insertEntryIfNew] + simp only [Bool.and_eq_true, Bool.not_eq_eq_eq_not, Bool.not_true, ite_eq_right_iff, and_imp] + intro h' + rw [BEq.comm, And.left not_contains] at h' + simp at h' + +theorem containsKey_insertListIfNewUnit [BEq α] [PartialEquivBEq α] {l : List ((_ : α) × Unit)} {toInsert : List α} {k : α} : - containsKey k (insertListIfNewUnit l toInsert) ↔ containsKey k l ∨ toInsert.contains k := by + containsKey k (insertListIfNewUnit l toInsert) = (containsKey k l || toInsert.contains k) := by induction toInsert generalizing l with | nil => simp [insertListIfNewUnit] | cons hd tl ih => simp only [insertListIfNewUnit, List.contains_cons] rw [ih, containsKey_insertEntryIfNew] - simp only [Bool.or_eq_true] - rw [or_comm (a:= (hd == k)= true), or_assoc, BEq.comm] + rw [Bool.or_comm (hd == k), Bool.or_assoc, BEq.comm (a:=hd)] theorem containsKey_of_containsKey_insertListIfNewUnit [BEq α] [PartialEquivBEq α] {l : List ((_ : α) × Unit)} {toInsert : List α} {k : α} (h₂ : toInsert.contains k = false) : containsKey k (insertListIfNewUnit l toInsert) → containsKey k l := by intro h₁ - rw [containsKey_insertListIfNewUnit] at h₁ - simp only [Bool.or_eq_true] at h₁ - cases h₁ with - | inl h => exact h - | inr h => - simp only [h, Bool.true_eq_false] at h₂ + rw [containsKey_insertListIfNewUnit, h₂] at h₁; simp at h₁; exact h₁ theorem getKey?_insertListIfNewUnit_of_contains_eq_false [BEq α] [EquivBEq α] {l : List ((_ : α) × Unit)} {toInsert : List α} {k : α} (not_contains : toInsert.contains k = false) : getKey? k (insertListIfNewUnit l toInsert) = getKey? k l := by - induction toInsert generalizing l with - | nil => simp [insertListIfNewUnit] - | cons hd tl ih => - simp only [insertListIfNewUnit] - rw [ih] - · rw [getKey?_insertEntryIfNew] - split - · rename_i hd_k - simp at not_contains - exfalso - rcases hd_k with ⟨h,_⟩ - rcases not_contains with ⟨h',_⟩ - rw [BEq.symm h] at h' - simp at h' - · rfl - · simp only [List.contains_cons, Bool.or_eq_false_iff] at not_contains - apply And.right not_contains + rw [getKey?_eq_getEntry?, getKey?_eq_getEntry?, + getEntry?_insertListIfNewUnit_of_contains_eq_false not_contains] theorem getKey?_insertListIfNewUnit_of_mem_of_contains_eq_false [BEq α] [EquivBEq α] {l : List ((_ : α) × Unit)} {toInsert : List α} @@ -2583,77 +2504,25 @@ theorem getKey?_insertListIfNewUnit_of_mem_of_contains_eq_false [BEq α] [EquivB (distinct : toInsert.Pairwise (fun a b => (a == b) = false)) (mem : k ∈ toInsert) (mem' : containsKey k l = false) : getKey? k' (insertListIfNewUnit l toInsert) = some k := by - induction toInsert generalizing l with - | nil => simp at mem - | cons hd tl ih => - simp only [insertListIfNewUnit] - simp only [List.mem_cons] at mem - cases mem with - | inl mem => - rw [getKey?_insertListIfNewUnit_of_contains_eq_false, ← mem] - · simp only [insertEntryIfNew, mem', cond_eq_if, Bool.false_eq_true, ↓reduceIte, - getKey?_cons, ite_eq_left_iff, Bool.not_eq_true] - intro h - rw [h] at k_beq - simp only [Bool.false_eq_true] at k_beq - · simp only [pairwise_cons] at distinct - simp only [List.contains_eq_any_beq, List.any_eq_false, Bool.not_eq_true] - intro a h - apply BEq.neq_of_beq_of_neq - apply BEq.symm k_beq - rw [mem] - apply (And.left distinct) a h - | inr mem => - apply ih - · simp only [pairwise_cons] at distinct - apply And.right distinct - · exact mem - · rw [containsKey_insertEntryIfNew] - simp only [Bool.or_eq_false_iff] - constructor - · simp only [pairwise_cons] at distinct - apply (And.left distinct) k mem - · exact mem' + simp only [getKey?_eq_getEntry?, getEntry?_insertListIfNewUnit, Option.map_eq_some', + Option.or_eq_some, getEntry?_eq_none] + exists ⟨k, ()⟩ + simp only [and_true] + right + constructor + · rw [containsKey_congr k_beq] at mem' + exact mem' + · apply getEntry?_of_mem (DistinctKeys.mapUnit distinct) k_beq + simp only [List.mem_map] + exists k theorem getKey?_insertListIfNewUnit_of_contains_of_contains [BEq α] [EquivBEq α] {l : List ((_ : α) × Unit)} {toInsert : List α} {k : α} - (distinct : toInsert.Pairwise (fun a b => (a == b) = false)) - (mem : toInsert.contains k) (mem' : containsKey k l = true): + (mem' : containsKey k l = true): getKey? k (insertListIfNewUnit l toInsert) = getKey? k l := by - induction toInsert generalizing l with - | nil => simp at mem - | cons hd tl ih => - simp at mem - cases mem with - | inl mem => - simp only [insertListIfNewUnit] - rw [getKey?_insertListIfNewUnit_of_contains_eq_false, getKey?_insertEntryIfNew] - · simp only [BEq.symm mem, true_and, ite_eq_right_iff] - have h':= containsKey_of_beq mem' mem - intro h - rw [h] at h' - simp only [Bool.false_eq_true] at h' - · simp only [pairwise_cons] at distinct - simp only [List.contains_eq_any_beq, List.any_eq_false, Bool.not_eq_true] - intro x x_mem - apply BEq.neq_of_beq_of_neq mem - apply And.left distinct x x_mem - | inr mem => - simp only [insertListIfNewUnit] - simp only [pairwise_cons] at distinct - rw [ih (distinct := And.right distinct) (mem := mem)] - · rw [getKey?_insertEntryIfNew] - split - · rename_i h - have h':= containsKey_of_beq mem' (BEq.symm (And.left h)) - rw [And.right h] at h' - simp only [Bool.false_eq_true] at h' - · rfl - · rw [containsKey_insertEntryIfNew] - simp only [Bool.or_eq_true] - right - apply mem' + rw [containsKey_eq_isSome_getEntry?] at mem' + simp [getKey?_eq_getEntry?, getEntry?_insertListIfNewUnit, Option.or_of_isSome mem'] theorem getKey_insertListIfNewUnit_of_contains_eq_false [BEq α] [EquivBEq α] {l : List ((_ : α) × Unit)} {toInsert : List α} {k : α} @@ -2661,11 +2530,7 @@ theorem getKey_insertListIfNewUnit_of_contains_eq_false [BEq α] [EquivBEq α] {h} : getKey k (insertListIfNewUnit l toInsert) h = getKey k l (containsKey_of_containsKey_insertListIfNewUnit not_contains h) := by - rw [← Option.some_inj] - rw [← getKey?_eq_some_getKey] - rw [getKey?_insertListIfNewUnit_of_contains_eq_false] - . rw [getKey?_eq_some_getKey] - . exact not_contains + rw [← Option.some_inj, ← getKey?_eq_some_getKey, getKey?_insertListIfNewUnit_of_contains_eq_false not_contains, getKey?_eq_some_getKey] theorem getKey_insertListIfNewUnit_of_mem_of_contains_eq_false [BEq α] [EquivBEq α] {l : List ((_ : α) × Unit)} {toInsert : List α} @@ -2675,36 +2540,23 @@ theorem getKey_insertListIfNewUnit_of_mem_of_contains_eq_false [BEq α] [EquivBE {h} : containsKey k l = false → getKey k' (insertListIfNewUnit l toInsert) h = k := by intro mem' - rw [← Option.some_inj] - rw [← getKey?_eq_some_getKey] - rw [getKey?_insertListIfNewUnit_of_mem_of_contains_eq_false] - . exact k_beq - . exact distinct - . exact mem - . exact mem' + rw [← Option.some_inj, ← getKey?_eq_some_getKey, + getKey?_insertListIfNewUnit_of_mem_of_contains_eq_false k_beq distinct mem mem'] theorem getKey_insertListIfNewUnit_of_contains_of_contains [BEq α] [EquivBEq α] {l : List ((_ : α) × Unit)} {toInsert : List α} {k : α} - (distinct : toInsert.Pairwise (fun a b => (a == b) = false)) - (mem : toInsert.contains k) (mem' : containsKey k l = true) {h}: + (mem' : containsKey k l = true) {h}: getKey k (insertListIfNewUnit l toInsert) h = getKey k l mem' := by - rw [← Option.some_inj] - rw [← getKey?_eq_some_getKey] - rw [← getKey?_eq_some_getKey] - rw [getKey?_insertListIfNewUnit_of_contains_of_contains] - · exact distinct - · exact mem - · exact mem' + rw [← Option.some_inj, ← getKey?_eq_some_getKey, ← getKey?_eq_some_getKey, + getKey?_insertListIfNewUnit_of_contains_of_contains mem'] theorem getKey!_insertListIfNewUnit_of_contains_eq_false [BEq α] [EquivBEq α] [Inhabited α] {l : List ((_ : α) × Unit)} {toInsert : List α} {k : α} (not_contains : toInsert.contains k = false) : getKey! k (insertListIfNewUnit l toInsert) = getKey! k l := by - rw [getKey!_eq_getKey?] - rw [getKey?_insertListIfNewUnit_of_contains_eq_false] - . rw [getKey!_eq_getKey?] - . exact not_contains + rw [getKey!_eq_getKey?, getKey?_insertListIfNewUnit_of_contains_eq_false not_contains, + getKey!_eq_getKey?] theorem getKey!_insertListIfNewUnit_of_mem_of_contains_eq_false [BEq α] [EquivBEq α] [Inhabited α] {l : List ((_ : α) × Unit)} {toInsert : List α} @@ -2712,35 +2564,22 @@ theorem getKey!_insertListIfNewUnit_of_mem_of_contains_eq_false [BEq α] [EquivB (distinct : toInsert.Pairwise (fun a b => (a == b) = false)) (mem : k ∈ toInsert) (h : containsKey k l = false) : getKey! k' (insertListIfNewUnit l toInsert) = k := by - rw [getKey!_eq_getKey?] - rw [getKey?_insertListIfNewUnit_of_mem_of_contains_eq_false] - . rw [Option.get!_some] - . exact k_beq - . exact distinct - . exact mem - . exact h + rw [getKey!_eq_getKey?, + getKey?_insertListIfNewUnit_of_mem_of_contains_eq_false k_beq distinct mem h, Option.get!_some] theorem getKey!_insertListIfNewUnit_of_contains_of_contains [BEq α] [EquivBEq α] [Inhabited α] {l : List ((_ : α) × Unit)} {toInsert : List α} {k : α} - (distinct : toInsert.Pairwise (fun a b => (a == b) = false)) - (mem : toInsert.contains k) (h : containsKey k l = true) : + (h : containsKey k l = true) : getKey! k (insertListIfNewUnit l toInsert) = getKey! k l := by - rw [getKey!_eq_getKey?] - rw [getKey?_insertListIfNewUnit_of_contains_of_contains] - . rw [getKey!_eq_getKey?] - · exact distinct - · exact mem - · exact h + rw [getKey!_eq_getKey?, getKey?_insertListIfNewUnit_of_contains_of_contains h, getKey!_eq_getKey?] theorem getKeyD_insertListIfNewUnit_of_contains_eq_false [BEq α] [EquivBEq α] {l : List ((_ : α) × Unit)} {toInsert : List α} {k fallback : α} (not_contains : toInsert.contains k = false) : getKeyD k (insertListIfNewUnit l toInsert) fallback = getKeyD k l fallback := by - rw [getKeyD_eq_getKey?] - rw [getKey?_insertListIfNewUnit_of_contains_eq_false] - . rw [getKeyD_eq_getKey?] - . exact not_contains + rw [getKeyD_eq_getKey?, getKey?_insertListIfNewUnit_of_contains_eq_false not_contains, + getKeyD_eq_getKey?] theorem getKeyD_insertListIfNewUnit_of_mem_of_contains_eq_false [BEq α] [EquivBEq α] {l : List ((_ : α) × Unit)} {toInsert : List α} @@ -2748,26 +2587,16 @@ theorem getKeyD_insertListIfNewUnit_of_mem_of_contains_eq_false [BEq α] [EquivB (distinct : toInsert.Pairwise (fun a b => (a == b) = false)) {mem : k ∈ toInsert } {h : containsKey k l = false} : getKeyD k' (insertListIfNewUnit l toInsert) fallback = k := by - rw [getKeyD_eq_getKey?] - rw [getKey?_insertListIfNewUnit_of_mem_of_contains_eq_false] - . rw [Option.getD_some] - . exact k_beq - . exact distinct - . exact mem - . exact h + rw [getKeyD_eq_getKey?, getKey?_insertListIfNewUnit_of_mem_of_contains_eq_false k_beq distinct + mem h, Option.getD_some] theorem getKeyD_insertListIfNewUnit_of_contains_of_contains [BEq α] [EquivBEq α] {l : List ((_ : α) × Unit)} {toInsert : List α} {k fallback : α} - (distinct : toInsert.Pairwise (fun a b => (a == b) = false)) - (mem : toInsert.contains k) (mem' : containsKey k l = true) : + (mem' : containsKey k l = true) : getKeyD k (insertListIfNewUnit l toInsert) fallback = getKeyD k l fallback := by - rw [getKeyD_eq_getKey?] - rw [getKey?_insertListIfNewUnit_of_contains_of_contains] - . rw [getKeyD_eq_getKey?] - · exact distinct - · exact mem - · exact mem' + rw [getKeyD_eq_getKey?, + getKey?_insertListIfNewUnit_of_contains_of_contains mem', getKeyD_eq_getKey?] theorem length_insertListIfNewUnit [BEq α] [EquivBEq α] {l : List ((_ : α) × Unit)} {toInsert : List α} @@ -2845,35 +2674,11 @@ theorem getValue?_list_unit [BEq α] {l : List ((_ : α) × Unit)} {k : α}: · simp [hd_k] · simp [hd_k, ih] -theorem getValue_list_unit [BEq α] {l : List ((_ : α) × Unit)} {k : α} {h}: - getValue k l h = () := by simp - -theorem getValue!_list_unit [BEq α] {l : List ((_ : α) × Unit)} {k : α} : - getValue! k l = () := by simp - -theorem getValueD_list_unit [BEq α] {l : List ((_ : α) × Unit)} {k : α} {fallback : Unit} : - getValueD k l fallback = () := by simp - theorem getValue?_insertListIfNewUnit [BEq α] [PartialEquivBEq α] {l : List ((_ : α) × Unit)} {toInsert : List α} {k : α}: getValue? k (insertListIfNewUnit l toInsert) = if containsKey k l ∨ toInsert.contains k then some () else none := by - simp [← containsKey_insertListIfNewUnit_iff, getValue?_list_unit] - -theorem getValue_insertListIfNewUnit [BEq α] [PartialEquivBEq α] - {l : List ((_ : α) × Unit)} {toInsert : List α} {k : α} {h}: - getValue k (insertListIfNewUnit l toInsert) h = () := by - rw [getValue_list_unit] - -theorem getValue!_insertListIfNewUnit [BEq α] [PartialEquivBEq α] - {l : List ((_ : α) × Unit)} {toInsert : List α} {k : α} : - getValue! k (insertListIfNewUnit l toInsert) = () := by - rw [getValue!_list_unit] - -theorem getValueD_insertListIfNewUnit [BEq α] [PartialEquivBEq α] - {l : List ((_ : α) × Unit)} {toInsert : List α} {k : α} {fallback : Unit}: - getValueD k (insertListIfNewUnit l toInsert) fallback = () := by - rw [getValueD_list_unit] + simp [containsKey_insertListIfNewUnit, getValue?_list_unit] end From 83d7b4a86ee8fc638464bfd9870060d8082bccc2 Mon Sep 17 00:00:00 2001 From: jt0202 Date: Fri, 10 Jan 2025 21:52:37 +0100 Subject: [PATCH 76/83] Rewrote getKey functions for insertManyIfNewUnit_list --- .../DHashMap/Internal/List/Associative.lean | 90 ++++++------- src/Std/Data/DHashMap/Internal/RawLemmas.lean | 106 +++++++-------- src/Std/Data/DHashMap/Lemmas.lean | 112 +++++++--------- src/Std/Data/DHashMap/RawLemmas.lean | 92 ++++++------- src/Std/Data/HashMap/Lemmas.lean | 123 ++++++++--------- src/Std/Data/HashMap/RawLemmas.lean | 126 ++++++++---------- src/Std/Data/HashSet/Lemmas.lean | 116 ++++++++-------- src/Std/Data/HashSet/RawLemmas.lean | 121 ++++++++--------- 8 files changed, 409 insertions(+), 477 deletions(-) diff --git a/src/Std/Data/DHashMap/Internal/List/Associative.lean b/src/Std/Data/DHashMap/Internal/List/Associative.lean index 91dd005125fe..26b849ad39a4 100644 --- a/src/Std/Data/DHashMap/Internal/List/Associative.lean +++ b/src/Std/Data/DHashMap/Internal/List/Associative.lean @@ -2491,14 +2491,15 @@ theorem containsKey_of_containsKey_insertListIfNewUnit [BEq α] [PartialEquivBEq intro h₁ rw [containsKey_insertListIfNewUnit, h₂] at h₁; simp at h₁; exact h₁ -theorem getKey?_insertListIfNewUnit_of_contains_eq_false [BEq α] [EquivBEq α] +theorem getKey?_insertListIfNewUnit_of_contains_eq_false_of_contains_eq_false [BEq α] [EquivBEq α] {l : List ((_ : α) × Unit)} {toInsert : List α} {k : α} - (not_contains : toInsert.contains k = false) : - getKey? k (insertListIfNewUnit l toInsert) = getKey? k l := by - rw [getKey?_eq_getEntry?, getKey?_eq_getEntry?, - getEntry?_insertListIfNewUnit_of_contains_eq_false not_contains] + (h : toInsert.contains k = false) (h': containsKey k l = false) : + getKey? k (insertListIfNewUnit l toInsert) = none := by + rw [getKey?_eq_getEntry?, + getEntry?_insertListIfNewUnit_of_contains_eq_false h, Option.map_eq_none', getEntry?_eq_none] + exact h' -theorem getKey?_insertListIfNewUnit_of_mem_of_contains_eq_false [BEq α] [EquivBEq α] +theorem getKey?_insertListIfNewUnit_of_contains_eq_false_of_mem [BEq α] [EquivBEq α] {l : List ((_ : α) × Unit)} {toInsert : List α} {k k' : α} (k_beq : k == k') (distinct : toInsert.Pairwise (fun a b => (a == b) = false)) @@ -2516,87 +2517,80 @@ theorem getKey?_insertListIfNewUnit_of_mem_of_contains_eq_false [BEq α] [EquivB simp only [List.mem_map] exists k -theorem getKey?_insertListIfNewUnit_of_contains_of_contains [BEq α] [EquivBEq α] +theorem getKey?_insertListIfNewUnit_of_contains [BEq α] [EquivBEq α] {l : List ((_ : α) × Unit)} {toInsert : List α} {k : α} - (mem' : containsKey k l = true): + (h : containsKey k l = true): getKey? k (insertListIfNewUnit l toInsert) = getKey? k l := by - rw [containsKey_eq_isSome_getEntry?] at mem' - simp [getKey?_eq_getEntry?, getEntry?_insertListIfNewUnit, Option.or_of_isSome mem'] - -theorem getKey_insertListIfNewUnit_of_contains_eq_false [BEq α] [EquivBEq α] - {l : List ((_ : α) × Unit)} {toInsert : List α} {k : α} - {not_contains : toInsert.contains k = false} - {h} : - getKey k (insertListIfNewUnit l toInsert) h = - getKey k l (containsKey_of_containsKey_insertListIfNewUnit not_contains h) := by - rw [← Option.some_inj, ← getKey?_eq_some_getKey, getKey?_insertListIfNewUnit_of_contains_eq_false not_contains, getKey?_eq_some_getKey] + rw [containsKey_eq_isSome_getEntry?] at h + simp [getKey?_eq_getEntry?, getEntry?_insertListIfNewUnit, Option.or_of_isSome h] -theorem getKey_insertListIfNewUnit_of_mem_of_contains_eq_false [BEq α] [EquivBEq α] +theorem getKey_insertListIfNewUnit_of_contains_eq_false_of_mem [BEq α] [EquivBEq α] {l : List ((_ : α) × Unit)} {toInsert : List α} {k k' : α} (k_beq : k == k') (distinct : toInsert.Pairwise (fun a b => (a == b) = false)) (mem : k ∈ toInsert) - {h} : containsKey k l = false → + {h} (contains_eq_false : containsKey k l = false) : getKey k' (insertListIfNewUnit l toInsert) h = k := by - intro mem' rw [← Option.some_inj, ← getKey?_eq_some_getKey, - getKey?_insertListIfNewUnit_of_mem_of_contains_eq_false k_beq distinct mem mem'] + getKey?_insertListIfNewUnit_of_contains_eq_false_of_mem k_beq distinct mem contains_eq_false] -theorem getKey_insertListIfNewUnit_of_contains_of_contains [BEq α] [EquivBEq α] +theorem getKey_insertListIfNewUnit_of_contains [BEq α] [EquivBEq α] {l : List ((_ : α) × Unit)} {toInsert : List α} {k : α} - (mem' : containsKey k l = true) {h}: - getKey k (insertListIfNewUnit l toInsert) h = getKey k l mem' := by + (contains : containsKey k l = true) {h}: + getKey k (insertListIfNewUnit l toInsert) h = getKey k l contains := by rw [← Option.some_inj, ← getKey?_eq_some_getKey, ← getKey?_eq_some_getKey, - getKey?_insertListIfNewUnit_of_contains_of_contains mem'] - -theorem getKey!_insertListIfNewUnit_of_contains_eq_false [BEq α] [EquivBEq α] [Inhabited α] - {l : List ((_ : α) × Unit)} {toInsert : List α} {k : α} - (not_contains : toInsert.contains k = false) : - getKey! k (insertListIfNewUnit l toInsert) = getKey! k l := by - rw [getKey!_eq_getKey?, getKey?_insertListIfNewUnit_of_contains_eq_false not_contains, - getKey!_eq_getKey?] + getKey?_insertListIfNewUnit_of_contains contains] + +theorem getKey!_insertListIfNewUnit_of_contains_eq_false_of_contains_eq_false [BEq α] [EquivBEq α] + [Inhabited α] {l : List ((_ : α) × Unit)} {toInsert : List α} {k : α} + (contains_eq_false : toInsert.contains k = false) (contains_eq_false' : containsKey k l = false): + getKey! k (insertListIfNewUnit l toInsert) = default := by + rw [getKey!_eq_getKey?, getKey?_insertListIfNewUnit_of_contains_eq_false_of_contains_eq_false + contains_eq_false contains_eq_false'] + simp -theorem getKey!_insertListIfNewUnit_of_mem_of_contains_eq_false [BEq α] [EquivBEq α] [Inhabited α] +theorem getKey!_insertListIfNewUnit_of_contains_eq_false_of_mem [BEq α] [EquivBEq α] [Inhabited α] {l : List ((_ : α) × Unit)} {toInsert : List α} {k k' : α} (k_beq : k == k') (distinct : toInsert.Pairwise (fun a b => (a == b) = false)) (mem : k ∈ toInsert) (h : containsKey k l = false) : getKey! k' (insertListIfNewUnit l toInsert) = k := by rw [getKey!_eq_getKey?, - getKey?_insertListIfNewUnit_of_mem_of_contains_eq_false k_beq distinct mem h, Option.get!_some] + getKey?_insertListIfNewUnit_of_contains_eq_false_of_mem k_beq distinct mem h, Option.get!_some] -theorem getKey!_insertListIfNewUnit_of_contains_of_contains [BEq α] [EquivBEq α] [Inhabited α] +theorem getKey!_insertListIfNewUnit_of_contains [BEq α] [EquivBEq α] [Inhabited α] {l : List ((_ : α) × Unit)} {toInsert : List α} {k : α} (h : containsKey k l = true) : getKey! k (insertListIfNewUnit l toInsert) = getKey! k l := by - rw [getKey!_eq_getKey?, getKey?_insertListIfNewUnit_of_contains_of_contains h, getKey!_eq_getKey?] + rw [getKey!_eq_getKey?, getKey?_insertListIfNewUnit_of_contains h, getKey!_eq_getKey?] -theorem getKeyD_insertListIfNewUnit_of_contains_eq_false [BEq α] [EquivBEq α] +theorem getKeyD_insertListIfNewUnit_of_contains_eq_false_of_contains_eq_false [BEq α] [EquivBEq α] {l : List ((_ : α) × Unit)} {toInsert : List α} {k fallback : α} - (not_contains : toInsert.contains k = false) : - getKeyD k (insertListIfNewUnit l toInsert) fallback = getKeyD k l fallback := by - rw [getKeyD_eq_getKey?, getKey?_insertListIfNewUnit_of_contains_eq_false not_contains, - getKeyD_eq_getKey?] + (contains_eq_false : toInsert.contains k = false) (contains_eq_false' : containsKey k l = false): + getKeyD k (insertListIfNewUnit l toInsert) fallback = fallback := by + rw [getKeyD_eq_getKey?, getKey?_insertListIfNewUnit_of_contains_eq_false_of_contains_eq_false + contains_eq_false contains_eq_false'] + simp -theorem getKeyD_insertListIfNewUnit_of_mem_of_contains_eq_false [BEq α] [EquivBEq α] +theorem getKeyD_insertListIfNewUnit_of_contains_eq_false_of_mem [BEq α] [EquivBEq α] {l : List ((_ : α) × Unit)} {toInsert : List α} {k k' fallback : α} (k_beq : k == k') (distinct : toInsert.Pairwise (fun a b => (a == b) = false)) - {mem : k ∈ toInsert } {h : containsKey k l = false} : + (mem : k ∈ toInsert) (h : containsKey k l = false) : getKeyD k' (insertListIfNewUnit l toInsert) fallback = k := by - rw [getKeyD_eq_getKey?, getKey?_insertListIfNewUnit_of_mem_of_contains_eq_false k_beq distinct + rw [getKeyD_eq_getKey?, getKey?_insertListIfNewUnit_of_contains_eq_false_of_mem k_beq distinct mem h, Option.getD_some] -theorem getKeyD_insertListIfNewUnit_of_contains_of_contains [BEq α] [EquivBEq α] +theorem getKeyD_insertListIfNewUnit_of_contains [BEq α] [EquivBEq α] {l : List ((_ : α) × Unit)} {toInsert : List α} {k fallback : α} - (mem' : containsKey k l = true) : + (contains : containsKey k l = true) : getKeyD k (insertListIfNewUnit l toInsert) fallback = getKeyD k l fallback := by rw [getKeyD_eq_getKey?, - getKey?_insertListIfNewUnit_of_contains_of_contains mem', getKeyD_eq_getKey?] + getKey?_insertListIfNewUnit_of_contains contains, getKeyD_eq_getKey?] theorem length_insertListIfNewUnit [BEq α] [EquivBEq α] {l : List ((_ : α) × Unit)} {toInsert : List α} diff --git a/src/Std/Data/DHashMap/Internal/RawLemmas.lean b/src/Std/Data/DHashMap/Internal/RawLemmas.lean index a545538cc0b4..ce9ee38cd51b 100644 --- a/src/Std/Data/DHashMap/Internal/RawLemmas.lean +++ b/src/Std/Data/DHashMap/Internal/RawLemmas.lean @@ -1189,84 +1189,76 @@ theorem contains_of_contains_insertManyIfNewUnit_list [EquivBEq α] [LawfulHasha (insertManyIfNewUnit m l).1.contains k → m.contains k := by simp_to_model using containsKey_of_containsKey_insertListIfNewUnit -theorem getKey?_insertManyIfNewUnit_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] +theorem getKey?_insertManyIfNewUnit_list_of_contains_eq_false_of_contains_eq_false + [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l : List α} {k : α} (h' : l.contains k = false) : - getKey? (insertManyIfNewUnit m l).1 k = getKey? m k := by - simp_to_model using getKey?_insertListIfNewUnit_of_contains_eq_false + m.contains k = false → getKey? (insertManyIfNewUnit m l).1 k = none := by + simp_to_model using getKey?_insertListIfNewUnit_of_contains_eq_false_of_contains_eq_false -theorem getKey?_insertManyIfNewUnit_list_of_mem_of_contains_eq_false [EquivBEq α] [LawfulHashable α] +theorem getKey?_insertManyIfNewUnit_list_of_contains_eq_false_of_mem [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l : List α} {k k' : α} (k_beq : k == k') (distinct : l.Pairwise (fun a b => (a == b) = false)) (mem : k ∈ l) : m.contains k = false → getKey? (insertManyIfNewUnit m l).1 k' = some k := by - simp_to_model using getKey?_insertListIfNewUnit_of_mem_of_contains_eq_false + simp_to_model using getKey?_insertListIfNewUnit_of_contains_eq_false_of_mem -theorem getKey?_insertManyIfNewUnit_list_of_contains_of_contains [EquivBEq α] [LawfulHashable α] - (h : m.1.WF) {l : List α} {k : α} - (distinct : l.Pairwise (fun a b => (a == b) = false)) - (h' : l.contains k) : +theorem getKey?_insertManyIfNewUnit_list_of_contains [EquivBEq α] [LawfulHashable α] + (h : m.1.WF) {l : List α} {k : α} : m.contains k → getKey? (insertManyIfNewUnit m l).1 k = getKey? m k := by - simp_to_model using getKey?_insertListIfNewUnit_of_contains_of_contains + simp_to_model using getKey?_insertListIfNewUnit_of_contains -theorem getKey_insertManyIfNewUnit_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] - (h : m.1.WF) {l : List α} {k : α} (h₁ : l.contains k = false) {h'} : - getKey (insertManyIfNewUnit m l).1 k h' = - getKey m k (contains_of_contains_insertManyIfNewUnit_list m h h₁ h') := by - simp_to_model using getKey_insertListIfNewUnit_of_contains_eq_false +theorem getKey_insertManyIfNewUnit_list_of_contains [EquivBEq α] [LawfulHashable α] + (h : m.1.WF) {l : List α} {k : α} {h'} (contains : m.contains k): + getKey (insertManyIfNewUnit m l).1 k h' = getKey m k contains := by + simp_to_model using getKey_insertListIfNewUnit_of_contains -theorem getKey_insertManyIfNewUnit_list_of_mem_of_contains_eq_false [EquivBEq α] [LawfulHashable α] +theorem getKey_insertManyIfNewUnit_list_of_contains_eq_false_of_mem [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l : List α} {k k' : α} (k_beq : k == k') (distinct : l.Pairwise (fun a b => (a == b) = false)) (mem : k ∈ l) {h'} : m.contains k = false → getKey (insertManyIfNewUnit m l).1 k' h' = k := by - simp_to_model using getKey_insertListIfNewUnit_of_mem_of_contains_eq_false + simp_to_model using getKey_insertListIfNewUnit_of_contains_eq_false_of_mem -theorem getKey_insertManyIfNewUnit_list_mem_of_contains_of_contains [EquivBEq α] [LawfulHashable α] - (h : m.1.WF) {l : List α} {k : α} - (distinct : l.Pairwise (fun a b => (a == b) = false)) - (h₁ : l.contains k) {mem' : m.contains k} {h'} : - getKey (insertManyIfNewUnit m l).1 k h' = getKey m k mem' := by - simp_to_model using getKey_insertListIfNewUnit_of_contains_of_contains +theorem getKey_insertManyIfNewUnit_list_mem_of_contains [EquivBEq α] [LawfulHashable α] + (h : m.1.WF) {l : List α} {k : α} (contains : m.contains k) {h'} : + getKey (insertManyIfNewUnit m l).1 k h' = getKey m k contains := by + simp_to_model using getKey_insertListIfNewUnit_of_contains -theorem getKey!_insertManyIfNewUnit_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] - [Inhabited α] (h : m.1.WF) {l : List α} {k : α} - (h' : l.contains k = false) : - getKey! (insertManyIfNewUnit m l).1 k = getKey! m k := by - simp_to_model using getKey!_insertListIfNewUnit_of_contains_eq_false +theorem getKey!_insertManyIfNewUnit_list_of_contains_eq_false_of_contains_eq_false + [EquivBEq α] [LawfulHashable α] [Inhabited α] (h : m.1.WF) {l : List α} {k : α} + (contains_eq_false : l.contains k = false) : + m.contains k = false → getKey! (insertManyIfNewUnit m l).1 k = default := by + simp_to_model using getKey!_insertListIfNewUnit_of_contains_eq_false_of_contains_eq_false -theorem getKey!_insertManyIfNewUnit_list_of_mem_of_contains_eq_false [EquivBEq α] [LawfulHashable α] +theorem getKey!_insertManyIfNewUnit_list_of_contains_eq_false_of_mem [EquivBEq α] [LawfulHashable α] [Inhabited α] (h : m.1.WF) {l : List α} {k k' : α} (k_beq : k == k') (distinct : l.Pairwise (fun a b => (a == b) = false)) (mem : k ∈ l) : contains m k = false → getKey! (insertManyIfNewUnit m l).1 k' = k := by - simp_to_model using getKey!_insertListIfNewUnit_of_mem_of_contains_eq_false + simp_to_model using getKey!_insertListIfNewUnit_of_contains_eq_false_of_mem -theorem getKey!_insertManyIfNewUnit_list_mem_of_contains_of_contains [EquivBEq α] [LawfulHashable α] - [Inhabited α] (h : m.1.WF) {l : List α} {k : α} - (distinct : l.Pairwise (fun a b => (a == b) = false)) - (h' : l.contains k) : +theorem getKey!_insertManyIfNewUnit_list_of_contains [EquivBEq α] [LawfulHashable α] + [Inhabited α] (h : m.1.WF) {l : List α} {k : α} : m.contains k → getKey! (insertManyIfNewUnit m l).1 k = getKey! m k := by - simp_to_model using getKey!_insertListIfNewUnit_of_contains_of_contains + simp_to_model using getKey!_insertListIfNewUnit_of_contains -theorem getKeyD_insertManyIfNewUnit_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] - (h : m.1.WF) {l : List α} {k fallback : α} +theorem getKeyD_insertManyIfNewUnit_list_of_contains_eq_false_of_contains_eq_false + [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l : List α} {k fallback : α} (h' : l.contains k = false) : - getKeyD (insertManyIfNewUnit m l).1 k fallback = getKeyD m k fallback := by - simp_to_model using getKeyD_insertListIfNewUnit_of_contains_eq_false + m.contains k = false → getKeyD (insertManyIfNewUnit m l).1 k fallback = fallback := by + simp_to_model using getKeyD_insertListIfNewUnit_of_contains_eq_false_of_contains_eq_false -theorem getKeyD_insertManyIfNewUnit_list_of_mem_of_contains_eq_false [EquivBEq α] [LawfulHashable α] +theorem getKeyD_insertManyIfNewUnit_list_of_contains_eq_false_of_mem [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l : List α} {k k' fallback : α} (k_beq : k == k') (distinct : l.Pairwise (fun a b => (a == b) = false)) - (mem : k ∈ l ) : + (mem : k ∈ l) : m.contains k = false → getKeyD (insertManyIfNewUnit m l).1 k' fallback = k := by - simp_to_model using getKeyD_insertListIfNewUnit_of_mem_of_contains_eq_false + simp_to_model using getKeyD_insertListIfNewUnit_of_contains_eq_false_of_mem -theorem getKeyD_insertManyIfNewUnit_list_of_contains_of_contains [EquivBEq α] [LawfulHashable α] - (h : m.1.WF) {l : List α} {k fallback : α} - (distinct : l.Pairwise (fun a b => (a == b) = false)) - (h' : l.contains k) : +theorem getKeyD_insertManyIfNewUnit_list_of_contains [EquivBEq α] [LawfulHashable α] + (h : m.1.WF) {l : List α} {k fallback : α} : m.contains k → getKeyD (insertManyIfNewUnit m l).1 k fallback = getKeyD m k fallback := by - simp_to_model using getKeyD_insertListIfNewUnit_of_contains_of_contains + simp_to_model using getKeyD_insertListIfNewUnit_of_contains theorem size_insertManyIfNewUnit_list [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l : List α} @@ -1626,14 +1618,15 @@ theorem contains_insertManyIfNewUnit_empty_list [EquivBEq α] [LawfulHashable α theorem getKey?_insertManyIfNewUnit_empty_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] {l : List α} {k : α} (h' : l.contains k = false) : getKey? (insertManyIfNewUnit (empty : Raw₀ α (fun _ => Unit)) l).1 k = none := by - simp [getKey?_insertManyIfNewUnit_list_of_contains_eq_false _ Raw.WF.empty₀ h'] + simp [getKey?_insertManyIfNewUnit_list_of_contains_eq_false_of_contains_eq_false + _ Raw.WF.empty₀ h'] theorem getKey?_insertManyIfNewUnit_empty_list_of_mem [EquivBEq α] [LawfulHashable α] {l : List α} {k k' : α} (k_beq : k == k') (distinct : l.Pairwise (fun a b => (a == b) = false)) (mem : k ∈ l) : getKey? (insertManyIfNewUnit (empty : Raw₀ α (fun _ => Unit)) l).1 k' = some k := by - simp [getKey?_insertManyIfNewUnit_list_of_mem_of_contains_eq_false _ Raw.WF.empty₀ k_beq - distinct mem contains_empty] + simp [getKey?_insertManyIfNewUnit_list_of_contains_eq_false_of_mem _ Raw.WF.empty₀ k_beq + distinct mem contains_empty] theorem getKey_insertManyIfNewUnit_empty_list_of_mem [EquivBEq α] [LawfulHashable α] {l : List α} @@ -1641,37 +1634,36 @@ theorem getKey_insertManyIfNewUnit_empty_list_of_mem [EquivBEq α] [LawfulHashab (distinct : l.Pairwise (fun a b => (a == b) = false)) (mem : k ∈ l) {h'} : getKey (insertManyIfNewUnit (empty : Raw₀ α (fun _ => Unit)) l).1 k' h' = k := by - simp [getKey_insertManyIfNewUnit_list_of_mem_of_contains_eq_false _ Raw.WF.empty₀ k_beq + simp [getKey_insertManyIfNewUnit_list_of_contains_eq_false_of_mem _ Raw.WF.empty₀ k_beq distinct mem contains_empty] theorem getKey!_insertManyIfNewUnit_empty_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] [Inhabited α] {l : List α} {k : α} (h' : l.contains k = false) : getKey! (insertManyIfNewUnit (empty : Raw₀ α (fun _ => Unit)) l).1 k = default := by - simp [getKey!_insertManyIfNewUnit_list_of_contains_eq_false _ Raw.WF.empty₀ h'] - apply getKey!_empty + simp [getKey!_insertManyIfNewUnit_list_of_contains_eq_false_of_contains_eq_false _ Raw.WF.empty₀ h'] theorem getKey!_insertManyIfNewUnit_empty_list_of_mem [EquivBEq α] [LawfulHashable α] [Inhabited α] {l : List α} {k k' : α} (k_beq : k == k') (distinct : l.Pairwise (fun a b => (a == b) = false)) (mem : k ∈ l) : getKey! (insertManyIfNewUnit (empty : Raw₀ α (fun _ => Unit)) l).1 k' = k := by - simp [getKey!_insertManyIfNewUnit_list_of_mem_of_contains_eq_false _ Raw.WF.empty₀ k_beq + simp [getKey!_insertManyIfNewUnit_list_of_contains_eq_false_of_mem _ Raw.WF.empty₀ k_beq distinct mem contains_empty] theorem getKeyD_insertManyIfNewUnit_empty_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] {l : List α} {k fallback : α} (h' : l.contains k = false) : getKeyD (insertManyIfNewUnit (empty : Raw₀ α (fun _ => Unit)) l).1 k fallback = fallback := by - simp [getKeyD_insertManyIfNewUnit_list_of_contains_eq_false _ Raw.WF.empty₀ h'] - apply getKeyD_empty + simp [getKeyD_insertManyIfNewUnit_list_of_contains_eq_false_of_contains_eq_false + _ Raw.WF.empty₀ h'] theorem getKeyD_insertManyIfNewUnit_empty_list_of_mem [EquivBEq α] [LawfulHashable α] {l : List α} {k k' fallback : α} (k_beq : k == k') (distinct : l.Pairwise (fun a b => (a == b) = false)) (mem : k ∈ l) : getKeyD (insertManyIfNewUnit (empty : Raw₀ α (fun _ => Unit)) l).1 k' fallback = k := by - simp [getKeyD_insertManyIfNewUnit_list_of_mem_of_contains_eq_false _ Raw.WF.empty₀ k_beq + simp [getKeyD_insertManyIfNewUnit_list_of_contains_eq_false_of_mem _ Raw.WF.empty₀ k_beq distinct mem contains_empty] theorem size_insertManyIfNewUnit_empty_list [EquivBEq α] [LawfulHashable α] diff --git a/src/Std/Data/DHashMap/Lemmas.lean b/src/Std/Data/DHashMap/Lemmas.lean index 6dac15cc9b82..747815e1e3f1 100644 --- a/src/Std/Data/DHashMap/Lemmas.lean +++ b/src/Std/Data/DHashMap/Lemmas.lean @@ -1325,91 +1325,79 @@ theorem mem_of_mem_insertManyIfNewUnit_list [EquivBEq α] [LawfulHashable α] k ∈ (insertManyIfNewUnit m l) → k ∈ m := Raw₀.Const.contains_of_contains_insertManyIfNewUnit_list ⟨m.1, _⟩ m.2 contains_eq_false -theorem getKey?_insertManyIfNewUnit_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] - {l : List α} {k : α} (contains_eq_false : l.contains k = false) : - getKey? (insertManyIfNewUnit m l) k = getKey? m k := - Raw₀.Const.getKey?_insertManyIfNewUnit_list_of_contains_eq_false ⟨m.1, _⟩ m.2 contains_eq_false - -theorem getKey?_insertManyIfNewUnit_list_of_mem_of_contains_eq_false [EquivBEq α] [LawfulHashable α] +theorem getKey?_insertManyIfNewUnit_list_of_contains_eq_false_of_contains_eq_false + [EquivBEq α] [LawfulHashable α] {l : List α} {k : α} + (contains_eq_false : m.contains k = false) (contains_eq_false' : l.contains k = false): + getKey? (insertManyIfNewUnit m l) k = none := + Raw₀.Const.getKey?_insertManyIfNewUnit_list_of_contains_eq_false_of_contains_eq_false + ⟨m.1, _⟩ m.2 contains_eq_false' contains_eq_false + +theorem getKey?_insertManyIfNewUnit_list_of_contains_eq_false_of_mem [EquivBEq α] [LawfulHashable α] {l : List α} {k k' : α} (k_beq : k == k') - (distinct : l.Pairwise (fun a b => (a == b) = false)) (mem : k ∈ l) - (contains_eq_false : m.contains k = false) : + (distinct : l.Pairwise (fun a b => (a == b) = false)) + (contains_eq_false : m.contains k = false) (mem : k ∈ l) : getKey? (insertManyIfNewUnit m l) k' = some k := - Raw₀.Const.getKey?_insertManyIfNewUnit_list_of_mem_of_contains_eq_false ⟨m.1, _⟩ + Raw₀.Const.getKey?_insertManyIfNewUnit_list_of_contains_eq_false_of_mem ⟨m.1, _⟩ m.2 k_beq distinct mem contains_eq_false -theorem getKey?_insertManyIfNewUnit_list_of_contains_of_contains [EquivBEq α] [LawfulHashable α] - {l : List α} {k : α} - (distinct : l.Pairwise (fun a b => (a == b) = false)) - (h : l.contains k) - (h' : m.contains k) : +theorem getKey?_insertManyIfNewUnit_list_of_mem [EquivBEq α] [LawfulHashable α] + {l : List α} {k : α} (h' : k ∈ m) : getKey? (insertManyIfNewUnit m l) k = getKey? m k := - Raw₀.Const.getKey?_insertManyIfNewUnit_list_of_contains_of_contains ⟨m.1, _⟩ m.2 distinct h h' - -theorem getKey_insertManyIfNewUnit_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] - {l : List α} {k : α} (contains_eq_false : l.contains k = false) {h} : - getKey (insertManyIfNewUnit m l) k h = - getKey m k (mem_of_mem_insertManyIfNewUnit_list contains_eq_false h) := - Raw₀.Const.getKey_insertManyIfNewUnit_list_of_contains_eq_false ⟨m.1, _⟩ m.2 contains_eq_false + Raw₀.Const.getKey?_insertManyIfNewUnit_list_of_contains ⟨m.1, _⟩ m.2 h' -theorem getKey_insertManyIfNewUnit_list_of_mem_of_contains_eq_false [EquivBEq α] [LawfulHashable α] +theorem getKey_insertManyIfNewUnit_list_of_contains_eq_false_of_mem [EquivBEq α] [LawfulHashable α] {l : List α} {k k' : α} (k_beq : k == k') - (distinct : l.Pairwise (fun a b => (a == b) = false)) - (mem : k ∈ l) {h} - (contains_eq_false : m.contains k = false) : + (distinct : l.Pairwise (fun a b => (a == b) = false)) (contains_eq_false : m.contains k = false) + (mem : k ∈ l) {h} : getKey (insertManyIfNewUnit m l) k' h = k := - Raw₀.Const.getKey_insertManyIfNewUnit_list_of_mem_of_contains_eq_false ⟨m.1, _⟩ m.2 k_beq + Raw₀.Const.getKey_insertManyIfNewUnit_list_of_contains_eq_false_of_mem ⟨m.1, _⟩ m.2 k_beq distinct mem contains_eq_false -theorem getKey_insertManyIfNewUnit_list_mem_of_contains_of_contains [EquivBEq α] [LawfulHashable α] - {l : List α} {k : α} - (distinct : l.Pairwise (fun a b => (a == b) = false)) - (h₁ : l.contains k) {h₂ : m.contains k} {h} : - getKey (insertManyIfNewUnit m l) k h = getKey m k h₂ := - Raw₀.Const.getKey_insertManyIfNewUnit_list_mem_of_contains_of_contains ⟨m.1, _⟩ m.2 distinct h₁ +theorem getKey_insertManyIfNewUnit_list_of_mem [EquivBEq α] [LawfulHashable α] + {l : List α} {k : α} (mem : k ∈ m) {h} : + getKey (insertManyIfNewUnit m l) k h = getKey m k mem := + Raw₀.Const.getKey_insertManyIfNewUnit_list_of_contains ⟨m.1, _⟩ m.2 _ -theorem getKey!_insertManyIfNewUnit_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] - [Inhabited α] {l : List α} {k : α} - (contains_eq_false : l.contains k = false) : - getKey! (insertManyIfNewUnit m l) k = getKey! m k := - Raw₀.Const.getKey!_insertManyIfNewUnit_list_of_contains_eq_false ⟨m.1, _⟩ m.2 contains_eq_false +theorem getKey!_insertManyIfNewUnit_list_of_contains_eq_false_of_contains_eq_false + [EquivBEq α] [LawfulHashable α] [Inhabited α] {l : List α} {k : α} + (contains_eq_false : m.contains k = false) (contains_eq_false' : l.contains k = false) : + getKey! (insertManyIfNewUnit m l) k = default := + Raw₀.Const.getKey!_insertManyIfNewUnit_list_of_contains_eq_false_of_contains_eq_false + ⟨m.1, _⟩ m.2 contains_eq_false' contains_eq_false -theorem getKey!_insertManyIfNewUnit_list_of_mem_of_contains_eq_false [EquivBEq α] [LawfulHashable α] +theorem getKey!_insertManyIfNewUnit_list_of_contains_eq_false_of_mem [EquivBEq α] [LawfulHashable α] [Inhabited α] {l : List α} {k k' : α} (k_beq : k == k') - (distinct : l.Pairwise (fun a b => (a == b) = false)) + (distinct : l.Pairwise (fun a b => (a == b) = false)) (contains_eq_false: m.contains k = false) (mem : k ∈ l) : - contains m k = false → getKey! (insertManyIfNewUnit m l) k' = k := - Raw₀.Const.getKey!_insertManyIfNewUnit_list_of_mem_of_contains_eq_false ⟨m.1, _⟩ m.2 k_beq - distinct mem + getKey! (insertManyIfNewUnit m l) k' = k := + Raw₀.Const.getKey!_insertManyIfNewUnit_list_of_contains_eq_false_of_mem ⟨m.1, _⟩ m.2 k_beq + distinct mem contains_eq_false -theorem getKey!_insertManyIfNewUnit_list_mem_of_contains_of_contains [EquivBEq α] [LawfulHashable α] - [Inhabited α] {l : List α} {k : α} - (distinct : l.Pairwise (fun a b => (a == b) = false)) - (h : l.contains k) : - m.contains k → getKey! (insertManyIfNewUnit m l) k = getKey! m k := - Raw₀.Const.getKey!_insertManyIfNewUnit_list_mem_of_contains_of_contains ⟨m.1, _⟩ m.2 distinct h +theorem getKey!_insertManyIfNewUnit_list_of_mem [EquivBEq α] [LawfulHashable α] + [Inhabited α] {l : List α} {k : α} (mem : k ∈ m): + getKey! (insertManyIfNewUnit m l) k = getKey! m k := + Raw₀.Const.getKey!_insertManyIfNewUnit_list_of_contains ⟨m.1, _⟩ m.2 mem -theorem getKeyD_insertManyIfNewUnit_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] - {l : List α} {k fallback : α} - (contains_eq_false : l.contains k = false) : - getKeyD (insertManyIfNewUnit m l) k fallback = getKeyD m k fallback := - Raw₀.Const.getKeyD_insertManyIfNewUnit_list_of_contains_eq_false ⟨m.1, _⟩ m.2 contains_eq_false +theorem getKeyD_insertManyIfNewUnit_list_of_contains_eq_false_of_contains_eq_false + [EquivBEq α] [LawfulHashable α] {l : List α} {k fallback : α} + (contains_eq_false : m.contains k = false) (contains_eq_false' : l.contains k = false) : + getKeyD (insertManyIfNewUnit m l) k fallback = fallback := + Raw₀.Const.getKeyD_insertManyIfNewUnit_list_of_contains_eq_false_of_contains_eq_false + ⟨m.1, _⟩ m.2 contains_eq_false' contains_eq_false -theorem getKeyD_insertManyIfNewUnit_list_of_mem_of_contains_eq_false [EquivBEq α] [LawfulHashable α] +theorem getKeyD_insertManyIfNewUnit_list_of_contains_eq_false_of_mem [EquivBEq α] [LawfulHashable α] {l : List α} {k k' fallback : α} (k_beq : k == k') (distinct : l.Pairwise (fun a b => (a == b) = false)) - (mem : k ∈ l ) (contains_eq_false : m.contains k = false) : + (contains_eq_false : m.contains k = false) (mem : k ∈ l ) : getKeyD (insertManyIfNewUnit m l) k' fallback = k := - Raw₀.Const.getKeyD_insertManyIfNewUnit_list_of_mem_of_contains_eq_false ⟨m.1, _⟩ m.2 k_beq + Raw₀.Const.getKeyD_insertManyIfNewUnit_list_of_contains_eq_false_of_mem ⟨m.1, _⟩ m.2 k_beq distinct mem contains_eq_false -theorem getKeyD_insertManyIfNewUnit_list_of_contains_of_contains [EquivBEq α] [LawfulHashable α] - {l : List α} {k fallback : α} - (distinct : l.Pairwise (fun a b => (a == b) = false)) - (h : l.contains k) : - m.contains k → getKeyD (insertManyIfNewUnit m l) k fallback = getKeyD m k fallback := - Raw₀.Const.getKeyD_insertManyIfNewUnit_list_of_contains_of_contains ⟨m.1, _⟩ m.2 distinct h +theorem getKeyD_insertManyIfNewUnit_list_of_mem [EquivBEq α] [LawfulHashable α] + {l : List α} {k fallback : α} (mem : k ∈ m): + getKeyD (insertManyIfNewUnit m l) k fallback = getKeyD m k fallback := + Raw₀.Const.getKeyD_insertManyIfNewUnit_list_of_contains ⟨m.1, _⟩ m.2 mem theorem size_insertManyIfNewUnit_list [EquivBEq α] [LawfulHashable α] {l : List α} diff --git a/src/Std/Data/DHashMap/RawLemmas.lean b/src/Std/Data/DHashMap/RawLemmas.lean index c409eaf006c1..6fb514e196dd 100644 --- a/src/Std/Data/DHashMap/RawLemmas.lean +++ b/src/Std/Data/DHashMap/RawLemmas.lean @@ -1423,84 +1423,76 @@ theorem mem_of_mem_insertManyIfNewUnit_list [EquivBEq α] [LawfulHashable α] (h simp only [mem_iff_contains] simp_to_raw using Raw₀.Const.contains_of_contains_insertManyIfNewUnit_list -theorem getKey?_insertManyIfNewUnit_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] - (h : m.WF) {l : List α} {k : α} (contains_eq_false : l.contains k = false) : - getKey? (insertManyIfNewUnit m l) k = getKey? m k := by - simp_to_raw using Raw₀.Const.getKey?_insertManyIfNewUnit_list_of_contains_eq_false +theorem getKey?_insertManyIfNewUnit_list_of_contains_eq_false_of_contains_eq_false + [EquivBEq α] [LawfulHashable α] (h : m.WF) {l : List α} {k : α} + (contains_eq_false : l.contains k = false) : + m.contains k = false → getKey? (insertManyIfNewUnit m l) k = none := by + simp_to_raw using Raw₀.Const.getKey?_insertManyIfNewUnit_list_of_contains_eq_false_of_contains_eq_false -theorem getKey?_insertManyIfNewUnit_list_of_mem_of_contains_eq_false [EquivBEq α] [LawfulHashable α] +theorem getKey?_insertManyIfNewUnit_list_of_contains_eq_false_of_mem [EquivBEq α] [LawfulHashable α] (h : m.WF) {l : List α} {k k' : α} (k_beq : k == k') (distinct : l.Pairwise (fun a b => (a == b) = false)) (mem : k ∈ l) : m.contains k = false → getKey? (insertManyIfNewUnit m l) k' = some k := by - simp_to_raw using Raw₀.Const.getKey?_insertManyIfNewUnit_list_of_mem_of_contains_eq_false + simp_to_raw using Raw₀.Const.getKey?_insertManyIfNewUnit_list_of_contains_eq_false_of_mem -theorem getKey?_insertManyIfNewUnit_list_of_contains_of_contains [EquivBEq α] [LawfulHashable α] - (h : m.WF) {l : List α} {k : α} - (distinct : l.Pairwise (fun a b => (a == b) = false)) - (h' : l.contains k) : - m.contains k → getKey? (insertManyIfNewUnit m l) k = getKey? m k := by - simp_to_raw using Raw₀.Const.getKey?_insertManyIfNewUnit_list_of_contains_of_contains +theorem getKey?_insertManyIfNewUnit_list_of_mem [EquivBEq α] [LawfulHashable α] + (h : m.WF) {l : List α} {k : α} : + k ∈ m → getKey? (insertManyIfNewUnit m l) k = getKey? m k := by + simp only [mem_iff_contains] + simp_to_raw using Raw₀.Const.getKey?_insertManyIfNewUnit_list_of_contains -theorem getKey_insertManyIfNewUnit_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] - (h : m.WF) {l : List α} {k : α} (contains_eq_false : l.contains k = false) {h'} : - getKey (insertManyIfNewUnit m l) k h' = - getKey m k (mem_of_mem_insertManyIfNewUnit_list h contains_eq_false h') := by - simp_to_raw using Raw₀.Const.getKey_insertManyIfNewUnit_list_of_contains_eq_false +theorem getKey_insertManyIfNewUnit_list_of_mem + [EquivBEq α] [LawfulHashable α] (h : m.WF) {l : List α} {k : α} {h'} (mem : k ∈ m): + getKey (insertManyIfNewUnit m l) k h' = getKey m k mem := by + simp_to_raw using Raw₀.Const.getKey_insertManyIfNewUnit_list_of_contains -theorem getKey_insertManyIfNewUnit_list_of_mem_of_contains_eq_false [EquivBEq α] [LawfulHashable α] +theorem getKey_insertManyIfNewUnit_list_of_contains_eq_false_of_mem [EquivBEq α] [LawfulHashable α] (h : m.WF) {l : List α} {k k' : α} (k_beq : k == k') (distinct : l.Pairwise (fun a b => (a == b) = false)) (mem : k ∈ l) {h'} : m.contains k = false → getKey (insertManyIfNewUnit m l) k' h' = k := by - simp_to_raw using Raw₀.Const.getKey_insertManyIfNewUnit_list_of_mem_of_contains_eq_false - -theorem getKey_insertManyIfNewUnit_list_mem_of_contains_of_contains [EquivBEq α] [LawfulHashable α] - (h : m.WF) {l : List α} {k : α} - (distinct : l.Pairwise (fun a b => (a == b) = false)) - (h₁ : l.contains k) (h₂ : m.contains k) {h₃} : - getKey (insertManyIfNewUnit m l) k h₃ = getKey m k h₂ := by - simp_to_raw using Raw₀.Const.getKey_insertManyIfNewUnit_list_mem_of_contains_of_contains + simp_to_raw using Raw₀.Const.getKey_insertManyIfNewUnit_list_of_contains_eq_false_of_mem -theorem getKey!_insertManyIfNewUnit_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] - [Inhabited α] (h : m.WF) {l : List α} {k : α} +theorem getKey!_insertManyIfNewUnit_list_of_contains_eq_false_of_contains_eq_false + [EquivBEq α] [LawfulHashable α] [Inhabited α] (h : m.WF) {l : List α} {k : α} (contains_eq_false : l.contains k = false) : - getKey! (insertManyIfNewUnit m l) k = getKey! m k := by - simp_to_raw using Raw₀.Const.getKey!_insertManyIfNewUnit_list_of_contains_eq_false + m.contains k = false → getKey! (insertManyIfNewUnit m l) k = default := by + simp_to_raw using + Raw₀.Const.getKey!_insertManyIfNewUnit_list_of_contains_eq_false_of_contains_eq_false -theorem getKey!_insertManyIfNewUnit_list_of_mem_of_contains_eq_false [EquivBEq α] [LawfulHashable α] +theorem getKey!_insertManyIfNewUnit_list_of_contains_eq_false_of_mem [EquivBEq α] [LawfulHashable α] [Inhabited α] (h : m.WF) {l : List α} {k k' : α} (k_beq : k == k') (distinct : l.Pairwise (fun a b => (a == b) = false)) (mem : k ∈ l) : contains m k = false → getKey! (insertManyIfNewUnit m l) k' = k := by - simp_to_raw using Raw₀.Const.getKey!_insertManyIfNewUnit_list_of_mem_of_contains_eq_false + simp_to_raw using Raw₀.Const.getKey!_insertManyIfNewUnit_list_of_contains_eq_false_of_mem -theorem getKey!_insertManyIfNewUnit_list_mem_of_contains_of_contains [EquivBEq α] [LawfulHashable α] - [Inhabited α] (h : m.WF) {l : List α} {k : α} - (distinct : l.Pairwise (fun a b => (a == b) = false)) - (h' : l.contains k) : - m.contains k → getKey! (insertManyIfNewUnit m l) k = getKey! m k := by - simp_to_raw using Raw₀.Const.getKey!_insertManyIfNewUnit_list_mem_of_contains_of_contains +theorem getKey!_insertManyIfNewUnit_list_of_mem [EquivBEq α] [LawfulHashable α] + [Inhabited α] (h : m.WF) {l : List α} {k : α} : + k ∈ m → getKey! (insertManyIfNewUnit m l) k = getKey! m k := by + simp only [mem_iff_contains] + simp_to_raw using Raw₀.Const.getKey!_insertManyIfNewUnit_list_of_contains -theorem getKeyD_insertManyIfNewUnit_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] - (h : m.WF) {l : List α} {k fallback : α} +theorem getKeyD_insertManyIfNewUnit_list_of_contains_eq_false_of_contains_eq_false + [EquivBEq α] [LawfulHashable α] (h : m.WF) {l : List α} {k fallback : α} (contains_eq_false : l.contains k = false) : - getKeyD (insertManyIfNewUnit m l) k fallback = getKeyD m k fallback := by - simp_to_raw using Raw₀.Const.getKeyD_insertManyIfNewUnit_list_of_contains_eq_false + m.contains k = false → getKeyD (insertManyIfNewUnit m l) k fallback = fallback := by + simp_to_raw using + Raw₀.Const.getKeyD_insertManyIfNewUnit_list_of_contains_eq_false_of_contains_eq_false -theorem getKeyD_insertManyIfNewUnit_list_of_mem_of_contains_eq_false [EquivBEq α] [LawfulHashable α] +theorem getKeyD_insertManyIfNewUnit_list_of_contains_eq_false_of_mem [EquivBEq α] [LawfulHashable α] (h : m.WF) {l : List α} {k k' fallback : α} (k_beq : k == k') (distinct : l.Pairwise (fun a b => (a == b) = false)) (mem : k ∈ l) : m.contains k = false → getKeyD (insertManyIfNewUnit m l) k' fallback = k := by - simp_to_raw using Raw₀.Const.getKeyD_insertManyIfNewUnit_list_of_mem_of_contains_eq_false + simp_to_raw using Raw₀.Const.getKeyD_insertManyIfNewUnit_list_of_contains_eq_false_of_mem -theorem getKeyD_insertManyIfNewUnit_list_of_contains_of_contains [EquivBEq α] [LawfulHashable α] - (h : m.WF) {l : List α} {k fallback : α} - (distinct : l.Pairwise (fun a b => (a == b) = false)) - (h' : l.contains k) : - m.contains k → getKeyD (insertManyIfNewUnit m l) k fallback = getKeyD m k fallback := by - simp_to_raw using Raw₀.Const.getKeyD_insertManyIfNewUnit_list_of_contains_of_contains +theorem getKeyD_insertManyIfNewUnit_list_of_mem [EquivBEq α] [LawfulHashable α] + (h : m.WF) {l : List α} {k fallback : α} : + k ∈ m → getKeyD (insertManyIfNewUnit m l) k fallback = getKeyD m k fallback := by + simp only [mem_iff_contains] + simp_to_raw using Raw₀.Const.getKeyD_insertManyIfNewUnit_list_of_contains theorem size_insertManyIfNewUnit_list [EquivBEq α] [LawfulHashable α] (h : m.WF) {l : List α} diff --git a/src/Std/Data/HashMap/Lemmas.lean b/src/Std/Data/HashMap/Lemmas.lean index 83a6faa720fa..f172ff1a1d3c 100644 --- a/src/Std/Data/HashMap/Lemmas.lean +++ b/src/Std/Data/HashMap/Lemmas.lean @@ -910,87 +910,76 @@ theorem getD_insertManyIfNewUnit_list getD (insertManyIfNewUnit m l) k fallback = () := by simp -theorem getKey?_insertManyIfNewUnit_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] - {l : List α} {k : α} (contains_eq_false : l.contains k = false) : - getKey? (insertManyIfNewUnit m l) k = getKey? m k := - DHashMap.Const.getKey?_insertManyIfNewUnit_list_of_contains_eq_false contains_eq_false - -theorem getKey?_insertManyIfNewUnit_list_of_mem_of_contains_eq_false [EquivBEq α] [LawfulHashable α] - {l : List α} {k k' : α} (k_beq : k == k') - (distinct : l.Pairwise (fun a b => (a == b) = false)) (mem : k ∈ l) - (contains_eq_false : m.contains k = false) : +theorem getKey?_insertManyIfNewUnit_list_of_contains_eq_false_of_contains_eq_false + [EquivBEq α] [LawfulHashable α] {l : List α} {k : α} + (contains_eq_false : m.contains k = false) (contains_eq_false' : l.contains k = false) : + getKey? (insertManyIfNewUnit m l) k = none := + DHashMap.Const.getKey?_insertManyIfNewUnit_list_of_contains_eq_false_of_contains_eq_false + contains_eq_false contains_eq_false' + +theorem getKey?_insertManyIfNewUnit_list_of_contains_eq_false_of_mem [EquivBEq α] [LawfulHashable α] + {l : List α} {k k' : α} (k_beq : k == k') (distinct : l.Pairwise (fun a b => (a == b) = false)) + (contains_eq_false : m.contains k = false) (mem : k ∈ l) : getKey? (insertManyIfNewUnit m l) k' = some k := - DHashMap.Const.getKey?_insertManyIfNewUnit_list_of_mem_of_contains_eq_false k_beq distinct mem contains_eq_false + DHashMap.Const.getKey?_insertManyIfNewUnit_list_of_contains_eq_false_of_mem + k_beq distinct contains_eq_false mem -theorem getKey?_insertManyIfNewUnit_list_of_contains_of_contains [EquivBEq α] [LawfulHashable α] - {l : List α} {k : α} - (distinct : l.Pairwise (fun a b => (a == b) = false)) - (h : l.contains k) - (h' : m.contains k) : +theorem getKey?_insertManyIfNewUnit_list_of_mem [EquivBEq α] [LawfulHashable α] + {l : List α} {k : α} (mem : k ∈ m) : getKey? (insertManyIfNewUnit m l) k = getKey? m k := - DHashMap.Const.getKey?_insertManyIfNewUnit_list_of_contains_of_contains distinct h h' - -theorem getKey_insertManyIfNewUnit_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] - {l : List α} {k : α} (contains_eq_false : l.contains k = false) {h} : - getKey (insertManyIfNewUnit m l) k h = - getKey m k (mem_of_mem_insertManyIfNewUnit_list contains_eq_false h) := - DHashMap.Const.getKey_insertManyIfNewUnit_list_of_contains_eq_false contains_eq_false + DHashMap.Const.getKey?_insertManyIfNewUnit_list_of_mem mem -theorem getKey_insertManyIfNewUnit_list_of_mem_of_contains_eq_false [EquivBEq α] [LawfulHashable α] - {l : List α} - {k k' : α} (k_beq : k == k') - (distinct : l.Pairwise (fun a b => (a == b) = false)) - (mem : k ∈ l) {h} - (contains_eq_false : m.contains k = false) : +theorem getKey_insertManyIfNewUnit_list_of_contains_eq_false_of_mem [EquivBEq α] [LawfulHashable α] + {l : List α} {k k' : α} (k_beq : k == k') (distinct : l.Pairwise (fun a b => (a == b) = false)) + (contains_eq_false : m.contains k = false) (mem : k ∈ l) {h} : getKey (insertManyIfNewUnit m l) k' h = k := - DHashMap.Const.getKey_insertManyIfNewUnit_list_of_mem_of_contains_eq_false k_beq distinct mem contains_eq_false - -theorem getKey_insertManyIfNewUnit_list_mem_of_contains_of_contains [EquivBEq α] [LawfulHashable α] - {l : List α} {k : α} - (distinct : l.Pairwise (fun a b => (a == b) = false)) - (h₁ : l.contains k) {h₂ : m.contains k} {h} : - getKey (insertManyIfNewUnit m l) k h = getKey m k h₂ := - DHashMap.Const.getKey_insertManyIfNewUnit_list_mem_of_contains_of_contains distinct h₁ - -theorem getKey!_insertManyIfNewUnit_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] - [Inhabited α] {l : List α} {k : α} - (contains_eq_false : l.contains k = false) : - getKey! (insertManyIfNewUnit m l) k = getKey! m k := - DHashMap.Const.getKey!_insertManyIfNewUnit_list_of_contains_eq_false contains_eq_false - -theorem getKey!_insertManyIfNewUnit_list_of_mem_of_contains_eq_false [EquivBEq α] [LawfulHashable α] + DHashMap.Const.getKey_insertManyIfNewUnit_list_of_contains_eq_false_of_mem + k_beq distinct contains_eq_false mem + +theorem getKey_insertManyIfNewUnit_list_mem_of_mem [EquivBEq α] [LawfulHashable α] + {l : List α} {k : α} (mem : k ∈ m) {h} : + getKey (insertManyIfNewUnit m l) k h = getKey m k mem := + DHashMap.Const.getKey_insertManyIfNewUnit_list_of_mem mem + +theorem getKey!_insertManyIfNewUnit_list_of_contains_eq_false_of_contains_eq_false + [EquivBEq α] [LawfulHashable α] [Inhabited α] {l : List α} {k : α} + (contains_eq_false : m.contains k = false) (contains_eq_false' : l.contains k = false) : + getKey! (insertManyIfNewUnit m l) k = default := + DHashMap.Const.getKey!_insertManyIfNewUnit_list_of_contains_eq_false_of_contains_eq_false + contains_eq_false contains_eq_false' + +theorem getKey!_insertManyIfNewUnit_list_of_contains_eq_false_of_mem [EquivBEq α] [LawfulHashable α] [Inhabited α] {l : List α} {k k' : α} (k_beq : k == k') (distinct : l.Pairwise (fun a b => (a == b) = false)) - (mem : k ∈ l) : - contains m k = false → getKey! (insertManyIfNewUnit m l) k' = k := - DHashMap.Const.getKey!_insertManyIfNewUnit_list_of_mem_of_contains_eq_false k_beq distinct mem + (contains_eq_false : m.contains k = false) (mem : k ∈ l) : + getKey! (insertManyIfNewUnit m l) k' = k := + DHashMap.Const.getKey!_insertManyIfNewUnit_list_of_contains_eq_false_of_mem + k_beq distinct contains_eq_false mem -theorem getKey!_insertManyIfNewUnit_list_mem_of_contains_of_contains [EquivBEq α] [LawfulHashable α] - [Inhabited α] {l : List α} {k : α} - (distinct : l.Pairwise (fun a b => (a == b) = false)) - (h : l.contains k) : - m.contains k → getKey! (insertManyIfNewUnit m l) k = getKey! m k := - DHashMap.Const.getKey!_insertManyIfNewUnit_list_mem_of_contains_of_contains distinct h +theorem getKey!_insertManyIfNewUnit_list_mem_of_mem [EquivBEq α] [LawfulHashable α] + [Inhabited α] {l : List α} {k : α} (mem : k ∈ m): + getKey! (insertManyIfNewUnit m l) k = getKey! m k := + DHashMap.Const.getKey!_insertManyIfNewUnit_list_of_mem mem -theorem getKeyD_insertManyIfNewUnit_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] - {l : List α} {k fallback : α} - (contains_eq_false : l.contains k = false) : - getKeyD (insertManyIfNewUnit m l) k fallback = getKeyD m k fallback := - DHashMap.Const.getKeyD_insertManyIfNewUnit_list_of_contains_eq_false contains_eq_false +theorem getKeyD_insertManyIfNewUnit_list_of_contains_eq_false_of_contains_eq_false + [EquivBEq α] [LawfulHashable α] {l : List α} {k fallback : α} + (contains_eq_false : m.contains k = false) (contains_eq_false' : l.contains k = false) : + getKeyD (insertManyIfNewUnit m l) k fallback = fallback := + DHashMap.Const.getKeyD_insertManyIfNewUnit_list_of_contains_eq_false_of_contains_eq_false + contains_eq_false contains_eq_false' -theorem getKeyD_insertManyIfNewUnit_list_of_mem_of_contains_eq_false [EquivBEq α] [LawfulHashable α] +theorem getKeyD_insertManyIfNewUnit_list_of_contains_eq_false_of_mem [EquivBEq α] [LawfulHashable α] {l : List α} {k k' fallback : α} (k_beq : k == k') (distinct : l.Pairwise (fun a b => (a == b) = false)) - (mem : k ∈ l ) (contains_eq_false : m.contains k = false) : + (contains_eq_false : m.contains k = false) (mem : k ∈ l ) : getKeyD (insertManyIfNewUnit m l) k' fallback = k := - DHashMap.Const.getKeyD_insertManyIfNewUnit_list_of_mem_of_contains_eq_false k_beq distinct mem contains_eq_false + DHashMap.Const.getKeyD_insertManyIfNewUnit_list_of_contains_eq_false_of_mem + k_beq distinct contains_eq_false mem -theorem getKeyD_insertManyIfNewUnit_list_of_contains_of_contains [EquivBEq α] [LawfulHashable α] - {l : List α} {k fallback : α} - (distinct : l.Pairwise (fun a b => (a == b) = false)) - (h : l.contains k) : - m.contains k → getKeyD (insertManyIfNewUnit m l) k fallback = getKeyD m k fallback := - DHashMap.Const.getKeyD_insertManyIfNewUnit_list_of_contains_of_contains distinct h +theorem getKeyD_insertManyIfNewUnit_list_of_mem [EquivBEq α] [LawfulHashable α] + {l : List α} {k fallback : α} (mem : k ∈ m): + getKeyD (insertManyIfNewUnit m l) k fallback = getKeyD m k fallback := + DHashMap.Const.getKeyD_insertManyIfNewUnit_list_of_mem mem theorem size_insertManyIfNewUnit_list [EquivBEq α] [LawfulHashable α] {l : List α} diff --git a/src/Std/Data/HashMap/RawLemmas.lean b/src/Std/Data/HashMap/RawLemmas.lean index d2f5c5b3672e..358af7204a5f 100644 --- a/src/Std/Data/HashMap/RawLemmas.lean +++ b/src/Std/Data/HashMap/RawLemmas.lean @@ -898,89 +898,79 @@ theorem mem_of_mem_insertManyIfNewUnit_list [EquivBEq α] [LawfulHashable α] (h k ∈ (insertManyIfNewUnit m l) → k ∈ m := DHashMap.Raw.Const.mem_of_mem_insertManyIfNewUnit_list h.out contains_eq_false -theorem getKey?_insertManyIfNewUnit_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] - (h : m.WF) {l : List α} {k : α} (contains_eq_false : l.contains k = false) : - getKey? (insertManyIfNewUnit m l) k = getKey? m k := - DHashMap.Raw.Const.getKey?_insertManyIfNewUnit_list_of_contains_eq_false h.out contains_eq_false - -theorem getKey?_insertManyIfNewUnit_list_of_mem_of_contains_eq_false [EquivBEq α] [LawfulHashable α] +theorem getKey?_insertManyIfNewUnit_list_of_contains_eq_false_of_contains_eq_false + [EquivBEq α] [LawfulHashable α] (h : m.WF) {l : List α} {k : α} + (contains_eq_false : m.contains k = false) (contains_eq_false': l.contains k = false) : + getKey? (insertManyIfNewUnit m l) k = none := + DHashMap.Raw.Const.getKey?_insertManyIfNewUnit_list_of_contains_eq_false_of_contains_eq_false + h.out contains_eq_false' contains_eq_false + +theorem getKey?_insertManyIfNewUnit_list_of_contains_eq_false_of_mem [EquivBEq α] [LawfulHashable α] (h : m.WF) {l : List α} {k k' : α} (k_beq : k == k') - (distinct : l.Pairwise (fun a b => (a == b) = false)) (mem : k ∈ l) : - m.contains k = false → getKey? (insertManyIfNewUnit m l) k' = some k := - DHashMap.Raw.Const.getKey?_insertManyIfNewUnit_list_of_mem_of_contains_eq_false - h.out k_beq distinct mem - -theorem getKey?_insertManyIfNewUnit_list_of_contains_of_contains [EquivBEq α] [LawfulHashable α] - (h : m.WF) {l : List α} {k : α} (distinct : l.Pairwise (fun a b => (a == b) = false)) - (h' : l.contains k) : - m.contains k → getKey? (insertManyIfNewUnit m l) k = getKey? m k := - DHashMap.Raw.Const.getKey?_insertManyIfNewUnit_list_of_contains_of_contains h.out distinct h' + (contains_eq_false : m.contains k = false) (mem : k ∈ l) : + getKey? (insertManyIfNewUnit m l) k' = some k := + DHashMap.Raw.Const.getKey?_insertManyIfNewUnit_list_of_contains_eq_false_of_mem + h.out k_beq distinct mem contains_eq_false -theorem getKey_insertManyIfNewUnit_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] - (h : m.WF) {l : List α} {k : α} (contains_eq_false : l.contains k = false) {h'} : - getKey (insertManyIfNewUnit m l) k h' = - getKey m k (mem_of_mem_insertManyIfNewUnit_list h contains_eq_false h') := - DHashMap.Raw.Const.getKey_insertManyIfNewUnit_list_of_contains_eq_false h.out contains_eq_false +theorem getKey?_insertManyIfNewUnit_list_of_mem [EquivBEq α] [LawfulHashable α] + (h : m.WF) {l : List α} {k : α} (mem : k ∈ m) : + getKey? (insertManyIfNewUnit m l) k = getKey? m k := + DHashMap.Raw.Const.getKey?_insertManyIfNewUnit_list_of_mem h.out mem -theorem getKey_insertManyIfNewUnit_list_of_mem_of_contains_eq_false [EquivBEq α] [LawfulHashable α] +theorem getKey_insertManyIfNewUnit_list_of_contains_eq_false_of_mem [EquivBEq α] [LawfulHashable α] (h : m.WF) {l : List α} {k k' : α} (k_beq : k == k') (distinct : l.Pairwise (fun a b => (a == b) = false)) - (mem : k ∈ l) {h'} : - m.contains k = false → getKey (insertManyIfNewUnit m l) k' h' = k := - DHashMap.Raw.Const.getKey_insertManyIfNewUnit_list_of_mem_of_contains_eq_false - h.out k_beq distinct mem - -theorem getKey_insertManyIfNewUnit_list_mem_of_contains_of_contains [EquivBEq α] [LawfulHashable α] - (h : m.WF) {l : List α} {k : α} - (distinct : l.Pairwise (fun a b => (a == b) = false)) - (h₁ : l.contains k) (h₂ : m.contains k) {h₃} : - getKey (insertManyIfNewUnit m l) k h₃ = getKey m k h₂ := - DHashMap.Raw.Const.getKey_insertManyIfNewUnit_list_mem_of_contains_of_contains - h.out distinct h₁ h₂ - -theorem getKey!_insertManyIfNewUnit_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] - [Inhabited α] (h : m.WF) {l : List α} {k : α} - (contains_eq_false : l.contains k = false) : - getKey! (insertManyIfNewUnit m l) k = getKey! m k := - DHashMap.Raw.Const.getKey!_insertManyIfNewUnit_list_of_contains_eq_false h.out contains_eq_false - -theorem getKey!_insertManyIfNewUnit_list_of_mem_of_contains_eq_false [EquivBEq α] [LawfulHashable α] + (contains_eq_false : m.contains k = false) (mem : k ∈ l) {h'} : + getKey (insertManyIfNewUnit m l) k' h' = k := + DHashMap.Raw.Const.getKey_insertManyIfNewUnit_list_of_contains_eq_false_of_mem + h.out k_beq distinct mem contains_eq_false + +theorem getKey_insertManyIfNewUnit_list_of_mem [EquivBEq α] [LawfulHashable α] + (h : m.WF) {l : List α} {k : α} (mem: k ∈ m) {h₃} : + getKey (insertManyIfNewUnit m l) k h₃ = getKey m k mem := + DHashMap.Raw.Const.getKey_insertManyIfNewUnit_list_of_mem h.out mem + +theorem getKey!_insertManyIfNewUnit_list_of_contains_eq_false_of_contains_eq_false + [EquivBEq α] [LawfulHashable α] [Inhabited α] (h : m.WF) {l : List α} {k : α} + (contains_eq_false : m.contains k = false) (contains_eq_false' : l.contains k = false) : + getKey! (insertManyIfNewUnit m l) k = default := + DHashMap.Raw.Const.getKey!_insertManyIfNewUnit_list_of_contains_eq_false_of_contains_eq_false + h.out contains_eq_false' contains_eq_false + +theorem getKey!_insertManyIfNewUnit_list_of_contains_eq_false_of_mem [EquivBEq α] [LawfulHashable α] [Inhabited α] (h : m.WF) {l : List α} {k k' : α} (k_beq : k == k') (distinct : l.Pairwise (fun a b => (a == b) = false)) - (mem : k ∈ l) : - contains m k = false → getKey! (insertManyIfNewUnit m l) k' = k := - DHashMap.Raw.Const.getKey!_insertManyIfNewUnit_list_of_mem_of_contains_eq_false - h.out k_beq distinct mem + (contains_eq_false : m.contains k = false) (mem : k ∈ l) : + getKey! (insertManyIfNewUnit m l) k' = k := + DHashMap.Raw.Const.getKey!_insertManyIfNewUnit_list_of_contains_eq_false_of_mem + h.out k_beq distinct mem contains_eq_false -theorem getKey!_insertManyIfNewUnit_list_mem_of_contains_of_contains [EquivBEq α] [LawfulHashable α] - [Inhabited α] (h : m.WF) {l : List α} {k : α} - (distinct : l.Pairwise (fun a b => (a == b) = false)) - (h' : l.contains k) : - m.contains k → getKey! (insertManyIfNewUnit m l) k = getKey! m k := - DHashMap.Raw.Const.getKey!_insertManyIfNewUnit_list_mem_of_contains_of_contains h.out distinct h' +theorem getKey!_insertManyIfNewUnit_list_of_mem [EquivBEq α] [LawfulHashable α] + [Inhabited α] (h : m.WF) {l : List α} {k : α} (mem : k ∈ m) : + getKey! (insertManyIfNewUnit m l) k = getKey! m k := + DHashMap.Raw.Const.getKey!_insertManyIfNewUnit_list_of_mem h.out mem -theorem getKeyD_insertManyIfNewUnit_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] - (h : m.WF) {l : List α} {k fallback : α} - (contains_eq_false : l.contains k = false) : - getKeyD (insertManyIfNewUnit m l) k fallback = getKeyD m k fallback := - DHashMap.Raw.Const.getKeyD_insertManyIfNewUnit_list_of_contains_eq_false h.out contains_eq_false +theorem getKeyD_insertManyIfNewUnit_list_of_contains_eq_false_of_contains_eq_false + [EquivBEq α] [LawfulHashable α] (h : m.WF) {l : List α} {k fallback : α} + (contains_eq_false : m.contains k = false) (contains_eq_false' : l.contains k = false) : + getKeyD (insertManyIfNewUnit m l) k fallback = fallback := + DHashMap.Raw.Const.getKeyD_insertManyIfNewUnit_list_of_contains_eq_false_of_contains_eq_false + h.out contains_eq_false' contains_eq_false -theorem getKeyD_insertManyIfNewUnit_list_of_mem_of_contains_eq_false [EquivBEq α] [LawfulHashable α] +theorem getKeyD_insertManyIfNewUnit_list_of_contains_eq_false_of_mem [EquivBEq α] [LawfulHashable α] (h : m.WF) {l : List α} {k k' fallback : α} (k_beq : k == k') (distinct : l.Pairwise (fun a b => (a == b) = false)) - (mem : k ∈ l) : - m.contains k = false → getKeyD (insertManyIfNewUnit m l) k' fallback = k := - DHashMap.Raw.Const.getKeyD_insertManyIfNewUnit_list_of_mem_of_contains_eq_false - h.out k_beq distinct mem + (contains_eq_false : m.contains k = false) (mem : k ∈ l) : + getKeyD (insertManyIfNewUnit m l) k' fallback = k := + DHashMap.Raw.Const.getKeyD_insertManyIfNewUnit_list_of_contains_eq_false_of_mem + h.out k_beq distinct mem contains_eq_false -theorem getKeyD_insertManyIfNewUnit_list_of_contains_of_contains [EquivBEq α] [LawfulHashable α] - (h : m.WF) {l : List α} {k fallback : α} - (distinct : l.Pairwise (fun a b => (a == b) = false)) - (h' : l.contains k) : - m.contains k → getKeyD (insertManyIfNewUnit m l) k fallback = getKeyD m k fallback := - DHashMap.Raw.Const.getKeyD_insertManyIfNewUnit_list_of_contains_of_contains h.out distinct h' +theorem getKeyD_insertManyIfNewUnit_list_of_mem [EquivBEq α] [LawfulHashable α] + (h : m.WF) {l : List α} {k fallback : α} (mem : k ∈ m): + getKeyD (insertManyIfNewUnit m l) k fallback = getKeyD m k fallback := + DHashMap.Raw.Const.getKeyD_insertManyIfNewUnit_list_of_mem h.out mem theorem size_insertManyIfNewUnit_list [EquivBEq α] [LawfulHashable α] (h : m.WF) {l : List α} diff --git a/src/Std/Data/HashSet/Lemmas.lean b/src/Std/Data/HashSet/Lemmas.lean index 340689731f10..f01c79f14279 100644 --- a/src/Std/Data/HashSet/Lemmas.lean +++ b/src/Std/Data/HashSet/Lemmas.lean @@ -405,87 +405,79 @@ theorem mem_of_mem_insertMany_list [EquivBEq α] [LawfulHashable α] k ∈ (insertMany m l) → k ∈ m := HashMap.mem_of_mem_insertManyIfNewUnit_list contains_eq_false -theorem get?_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] - {l : List α} {k : α} (contains_eq_false : l.contains k = false) : - get? (insertMany m l) k = get? m k := - HashMap.getKey?_insertManyIfNewUnit_list_of_contains_eq_false contains_eq_false - -theorem get?_insertMany_list_of_mem_of_contains_eq_false [EquivBEq α] [LawfulHashable α] +theorem get?_insertMany_list_of_contains_eq_false_of_contains_eq_false + [EquivBEq α] [LawfulHashable α] {l : List α} {k : α} + (contains_eq_false : m.contains k = false) (contains_eq_false' : l.contains k = false) : + get? (insertMany m l) k = none := + HashMap.getKey?_insertManyIfNewUnit_list_of_contains_eq_false_of_contains_eq_false + contains_eq_false contains_eq_false' + +theorem get?_insertMany_list_of_contains_eq_false_of_mem [EquivBEq α] [LawfulHashable α] {l : List α} {k k' : α} (k_beq : k == k') - (distinct : l.Pairwise (fun a b => (a == b) = false)) (mem : k ∈ l) - (contains_eq_false : m.contains k = false) : + (distinct : l.Pairwise (fun a b => (a == b) = false)) + (contains_eq_false : m.contains k = false) (mem : k ∈ l) : get? (insertMany m l) k' = some k := - HashMap.getKey?_insertManyIfNewUnit_list_of_mem_of_contains_eq_false k_beq distinct mem contains_eq_false + HashMap.getKey?_insertManyIfNewUnit_list_of_contains_eq_false_of_mem + k_beq distinct contains_eq_false mem -theorem get?_insertMany_list_of_contains_of_contains [EquivBEq α] [LawfulHashable α] - {l : List α} {k : α} - (distinct : l.Pairwise (fun a b => (a == b) = false)) - (h : l.contains k) - (h' : m.contains k) : +theorem get?_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] + {l : List α} {k : α} (mem : k ∈ m): get? (insertMany m l) k = get? m k := - HashMap.getKey?_insertManyIfNewUnit_list_of_contains_of_contains distinct h h' - -theorem get_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] - {l : List α} {k : α} (contains_eq_false : l.contains k = false) {h} : - get (insertMany m l) k h = - get m k (mem_of_mem_insertMany_list contains_eq_false h) := - HashMap.getKey_insertManyIfNewUnit_list_of_contains_eq_false contains_eq_false + HashMap.getKey?_insertManyIfNewUnit_list_of_mem mem -theorem get_insertMany_list_of_mem_of_contains_eq_false [EquivBEq α] [LawfulHashable α] +theorem get_insertMany_list_of_contains_eq_false_of_mem [EquivBEq α] [LawfulHashable α] {l : List α} {k k' : α} (k_beq : k == k') (distinct : l.Pairwise (fun a b => (a == b) = false)) - (mem : k ∈ l) {h} - (contains_eq_false : m.contains k = false) : + (contains_eq_false : m.contains k = false) (mem : k ∈ l) {h} : get (insertMany m l) k' h = k := - HashMap.getKey_insertManyIfNewUnit_list_of_mem_of_contains_eq_false k_beq distinct mem contains_eq_false - -theorem get_insertMany_list_mem_of_contains_of_contains [EquivBEq α] [LawfulHashable α] - {l : List α} {k : α} - (distinct : l.Pairwise (fun a b => (a == b) = false)) - (h₁ : l.contains k) {h₂ : m.contains k} {h} : - get (insertMany m l) k h = get m k h₂ := - HashMap.getKey_insertManyIfNewUnit_list_mem_of_contains_of_contains distinct h₁ - -theorem get!_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] - [Inhabited α] {l : List α} {k : α} - (contains_eq_false : l.contains k = false) : - get! (insertMany m l) k = get! m k := - HashMap.getKey!_insertManyIfNewUnit_list_of_contains_eq_false contains_eq_false - -theorem get!_insertMany_list_of_mem_of_contains_eq_false [EquivBEq α] [LawfulHashable α] + HashMap.getKey_insertManyIfNewUnit_list_of_contains_eq_false_of_mem + k_beq distinct contains_eq_false mem + +theorem get_insertMany_list_mem_of_mem [EquivBEq α] [LawfulHashable α] + {l : List α} {k : α} (mem : k ∈ m) {h}: + get (insertMany m l) k h = get m k mem := + HashMap.getKey_insertManyIfNewUnit_list_mem_of_mem mem + +theorem get!_insertMany_list_of_contains_eq_false_of_contains_eq_false + [EquivBEq α] [LawfulHashable α] [Inhabited α] {l : List α} {k : α} + (contains_eq_false : m.contains k = false) (contains_eq_false' : l.contains k = false) : + get! (insertMany m l) k = default := + HashMap.getKey!_insertManyIfNewUnit_list_of_contains_eq_false_of_contains_eq_false + contains_eq_false contains_eq_false' + +theorem get!_insertMany_list_of_contains_eq_false_of_mem [EquivBEq α] [LawfulHashable α] [Inhabited α] {l : List α} {k k' : α} (k_beq : k == k') (distinct : l.Pairwise (fun a b => (a == b) = false)) - (mem : k ∈ l) : - contains m k = false → get! (insertMany m l) k' = k := - HashMap.getKey!_insertManyIfNewUnit_list_of_mem_of_contains_eq_false k_beq distinct mem + (contains_eq_false : m.contains k = false) (mem : k ∈ l) : + get! (insertMany m l) k' = k := + HashMap.getKey!_insertManyIfNewUnit_list_of_contains_eq_false_of_mem + k_beq distinct contains_eq_false mem -theorem get!_insertMany_list_mem_of_contains_of_contains [EquivBEq α] [LawfulHashable α] - [Inhabited α] {l : List α} {k : α} - (distinct : l.Pairwise (fun a b => (a == b) = false)) - (h : l.contains k) : - m.contains k → get! (insertMany m l) k = get! m k := - HashMap.getKey!_insertManyIfNewUnit_list_mem_of_contains_of_contains distinct h +theorem get!_insertMany_list_mem_of_mem [EquivBEq α] [LawfulHashable α] + [Inhabited α] {l : List α} {k : α} (mem : k ∈ m) : + get! (insertMany m l) k = get! m k := + HashMap.getKey!_insertManyIfNewUnit_list_mem_of_mem mem -theorem getD_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] - {l : List α} {k fallback : α} - (contains_eq_false : l.contains k = false) : - getD (insertMany m l) k fallback = getD m k fallback := - HashMap.getKeyD_insertManyIfNewUnit_list_of_contains_eq_false contains_eq_false +theorem getD_insertMany_list_of_contains_eq_false_of_contains_eq_false + [EquivBEq α] [LawfulHashable α] {l : List α} {k fallback : α} + (contains_eq_false : m.contains k = false) (contains_eq_false' : l.contains k = false) : + getD (insertMany m l) k fallback = fallback := + HashMap.getKeyD_insertManyIfNewUnit_list_of_contains_eq_false_of_contains_eq_false + contains_eq_false contains_eq_false' -theorem getD_insertMany_list_of_mem_of_contains_eq_false [EquivBEq α] [LawfulHashable α] +theorem getD_insertMany_list_of_contains_eq_false_of_mem [EquivBEq α] [LawfulHashable α] {l : List α} {k k' fallback : α} (k_beq : k == k') (distinct : l.Pairwise (fun a b => (a == b) = false)) (mem : k ∈ l ) (contains_eq_false : m.contains k = false) : getD (insertMany m l) k' fallback = k := - HashMap.getKeyD_insertManyIfNewUnit_list_of_mem_of_contains_eq_false k_beq distinct mem contains_eq_false + HashMap.getKeyD_insertManyIfNewUnit_list_of_contains_eq_false_of_mem + k_beq distinct contains_eq_false mem -theorem getD_insertMany_list_of_contains_of_contains [EquivBEq α] [LawfulHashable α] - {l : List α} {k fallback : α} - (distinct : l.Pairwise (fun a b => (a == b) = false)) - (h : l.contains k) : - m.contains k → getD (insertMany m l) k fallback = getD m k fallback := - HashMap.getKeyD_insertManyIfNewUnit_list_of_contains_of_contains distinct h +theorem getD_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] + {l : List α} {k fallback : α} (mem : k ∈ m): + getD (insertMany m l) k fallback = getD m k fallback := + HashMap.getKeyD_insertManyIfNewUnit_list_of_mem mem theorem size_insertMany_list [EquivBEq α] [LawfulHashable α] {l : List α} diff --git a/src/Std/Data/HashSet/RawLemmas.lean b/src/Std/Data/HashSet/RawLemmas.lean index f86a52f78a3d..93178271dec8 100644 --- a/src/Std/Data/HashSet/RawLemmas.lean +++ b/src/Std/Data/HashSet/RawLemmas.lean @@ -421,84 +421,79 @@ theorem mem_of_mem_insertMany_list [EquivBEq α] [LawfulHashable α] (h : m.WF) k ∈ (insertMany m l) → k ∈ m := HashMap.Raw.mem_of_mem_insertManyIfNewUnit_list h.1 contains_eq_false -theorem get?_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] - (h : m.WF) {l : List α} {k : α} (contains_eq_false : l.contains k = false) : - get? (insertMany m l) k = get? m k := - HashMap.Raw.getKey?_insertManyIfNewUnit_list_of_contains_eq_false h.1 contains_eq_false - -theorem get?_insertMany_list_of_mem_of_contains_eq_false [EquivBEq α] [LawfulHashable α] +theorem get?_insertMany_list_of_contains_eq_false_of_contains_eq_false + [EquivBEq α] [LawfulHashable α] (h : m.WF) {l : List α} {k : α} + (contains_eq_false : m.contains k = false) (contains_eq_false' : l.contains k = false) : + get? (insertMany m l) k = none := + HashMap.Raw.getKey?_insertManyIfNewUnit_list_of_contains_eq_false_of_contains_eq_false + h.1 contains_eq_false contains_eq_false' + +theorem get?_insertMany_list_of_contains_eq_false_of_mem [EquivBEq α] [LawfulHashable α] (h : m.WF) {l : List α} {k k' : α} (k_beq : k == k') - (distinct : l.Pairwise (fun a b => (a == b) = false)) (mem : k ∈ l) : - m.contains k = false → get? (insertMany m l) k' = some k := - HashMap.Raw.getKey?_insertManyIfNewUnit_list_of_mem_of_contains_eq_false h.1 k_beq distinct mem - -theorem get?_insertMany_list_of_contains_of_contains [EquivBEq α] [LawfulHashable α] - (h : m.WF) {l : List α} {k : α} (distinct : l.Pairwise (fun a b => (a == b) = false)) - (h' : l.contains k) : - m.contains k → get? (insertMany m l) k = get? m k := - HashMap.Raw.getKey?_insertManyIfNewUnit_list_of_contains_of_contains h.1 distinct h' + (contains_eq_false : m.contains k = false) (mem : k ∈ l) : + get? (insertMany m l) k' = some k := + HashMap.Raw.getKey?_insertManyIfNewUnit_list_of_contains_eq_false_of_mem + h.1 k_beq distinct contains_eq_false mem -theorem get_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] - (h : m.WF) {l : List α} {k : α} (contains_eq_false : l.contains k = false) {h'} : - get (insertMany m l) k h' = - get m k (mem_of_mem_insertMany_list h contains_eq_false h') := - HashMap.Raw.getKey_insertManyIfNewUnit_list_of_contains_eq_false h.1 contains_eq_false +theorem get?_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] + (h : m.WF) {l : List α} {k : α} (mem : k ∈ m) : + get? (insertMany m l) k = get? m k := + HashMap.Raw.getKey?_insertManyIfNewUnit_list_of_mem h.1 mem -theorem get_insertMany_list_of_mem_of_contains_eq_false [EquivBEq α] [LawfulHashable α] +theorem get_insertMany_list_of_contains_eq_false_of_mem [EquivBEq α] [LawfulHashable α] (h : m.WF) {l : List α} {k k' : α} (k_beq : k == k') (distinct : l.Pairwise (fun a b => (a == b) = false)) - (mem : k ∈ l) {h'} : - m.contains k = false → get (insertMany m l) k' h' = k := - HashMap.Raw.getKey_insertManyIfNewUnit_list_of_mem_of_contains_eq_false h.1 k_beq distinct mem - -theorem get_insertMany_list_mem_of_contains_of_contains [EquivBEq α] [LawfulHashable α] - (h : m.WF) {l : List α} {k : α} - (distinct : l.Pairwise (fun a b => (a == b) = false)) - (h₁ : l.contains k) (h₂ : m.contains k) {h₃} : - get (insertMany m l) k h₃ = get m k h₂ := - HashMap.Raw.getKey_insertManyIfNewUnit_list_mem_of_contains_of_contains h.1 distinct h₁ h₂ - -theorem get!_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] - [Inhabited α] (h : m.WF) {l : List α} {k : α} - (contains_eq_false : l.contains k = false) : - get! (insertMany m l) k = get! m k := - HashMap.Raw.getKey!_insertManyIfNewUnit_list_of_contains_eq_false h.1 contains_eq_false - -theorem get!_insertMany_list_of_mem_of_contains_eq_false [EquivBEq α] [LawfulHashable α] + (contains_eq_false : m.contains k = false) (mem : k ∈ l) {h'} : + get (insertMany m l) k' h' = k := + HashMap.Raw.getKey_insertManyIfNewUnit_list_of_contains_eq_false_of_mem + h.1 k_beq distinct contains_eq_false mem + +theorem get_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] + (h : m.WF) {l : List α} {k : α} (mem : k ∈ m) {h₃} : + get (insertMany m l) k h₃ = get m k mem := + HashMap.Raw.getKey_insertManyIfNewUnit_list_of_mem h.1 mem + +theorem get!_insertMany_list_of_contains_eq_false_of_contains_eq_false + [EquivBEq α] [LawfulHashable α] [Inhabited α] (h : m.WF) {l : List α} {k : α} + (contains_eq_false : m.contains k = false) (contains_eq_false' : l.contains k = false) : + get! (insertMany m l) k = default := + HashMap.Raw.getKey!_insertManyIfNewUnit_list_of_contains_eq_false_of_contains_eq_false + h.1 contains_eq_false contains_eq_false' + +theorem get!_insertMany_list_of_contains_eq_false_of_mem [EquivBEq α] [LawfulHashable α] [Inhabited α] (h : m.WF) {l : List α} {k k' : α} (k_beq : k == k') (distinct : l.Pairwise (fun a b => (a == b) = false)) - (mem : k ∈ l) : - contains m k = false → get! (insertMany m l) k' = k := - HashMap.Raw.getKey!_insertManyIfNewUnit_list_of_mem_of_contains_eq_false h.1 k_beq distinct mem + (contains_eq_false : contains m k = false) (mem : k ∈ l) : + get! (insertMany m l) k' = k := + HashMap.Raw.getKey!_insertManyIfNewUnit_list_of_contains_eq_false_of_mem + h.1 k_beq distinct contains_eq_false mem -theorem get!_insertMany_list_mem_of_contains_of_contains [EquivBEq α] [LawfulHashable α] - [Inhabited α] (h : m.WF) {l : List α} {k : α} - (distinct : l.Pairwise (fun a b => (a == b) = false)) - (h' : l.contains k) : - m.contains k → get! (insertMany m l) k = get! m k := - HashMap.Raw.getKey!_insertManyIfNewUnit_list_mem_of_contains_of_contains h.1 distinct h' +theorem get!_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] + [Inhabited α] (h : m.WF) {l : List α} {k : α} (mem : k ∈ m) : + get! (insertMany m l) k = get! m k := + HashMap.Raw.getKey!_insertManyIfNewUnit_list_of_mem h.1 mem -theorem getD_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] - (h : m.WF) {l : List α} {k fallback : α} - (contains_eq_false : l.contains k = false) : - getD (insertMany m l) k fallback = getD m k fallback := - HashMap.Raw.getKeyD_insertManyIfNewUnit_list_of_contains_eq_false h.1 contains_eq_false +theorem getD_insertMany_list_of_contains_eq_false_of_contains_eq_false + [EquivBEq α] [LawfulHashable α] (h : m.WF) {l : List α} {k fallback : α} + (contains_eq_false : m.contains k = false) (contains_eq_false' : l.contains k = false) : + getD (insertMany m l) k fallback = fallback := + HashMap.Raw.getKeyD_insertManyIfNewUnit_list_of_contains_eq_false_of_contains_eq_false + h.1 contains_eq_false contains_eq_false' -theorem getD_insertMany_list_of_mem_of_contains_eq_false [EquivBEq α] [LawfulHashable α] +theorem getD_insertMany_list_of_contains_eq_false_of_mem [EquivBEq α] [LawfulHashable α] (h : m.WF) {l : List α} {k k' fallback : α} (k_beq : k == k') (distinct : l.Pairwise (fun a b => (a == b) = false)) - (mem : k ∈ l) : - m.contains k = false → getD (insertMany m l) k' fallback = k := - HashMap.Raw.getKeyD_insertManyIfNewUnit_list_of_mem_of_contains_eq_false h.1 k_beq distinct mem + (contains_eq_false : m.contains k = false) (mem : k ∈ l) : + getD (insertMany m l) k' fallback = k := + HashMap.Raw.getKeyD_insertManyIfNewUnit_list_of_contains_eq_false_of_mem + h.1 k_beq distinct contains_eq_false mem -theorem getD_insertMany_list_of_contains_of_contains [EquivBEq α] [LawfulHashable α] - (h : m.WF) {l : List α} {k fallback : α} - (distinct : l.Pairwise (fun a b => (a == b) = false)) - (h' : l.contains k) : - m.contains k → getD (insertMany m l) k fallback = getD m k fallback := - HashMap.Raw.getKeyD_insertManyIfNewUnit_list_of_contains_of_contains h.1 distinct h' +theorem getD_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] + (h : m.WF) {l : List α} {k fallback : α} (mem : k ∈ m) : + getD (insertMany m l) k fallback = getD m k fallback := + HashMap.Raw.getKeyD_insertManyIfNewUnit_list_of_mem h.1 mem theorem size_insertMany_list [EquivBEq α] [LawfulHashable α] (h : m.WF) {l : List α} From 281c080b7b23c84f0016a2f5426742fe4e9b9e04 Mon Sep 17 00:00:00 2001 From: monsterkrampe Date: Mon, 13 Jan 2025 08:01:08 +0100 Subject: [PATCH 77/83] Fix typo in HashMap/RawLemmas --- src/Std/Data/HashMap/RawLemmas.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/src/Std/Data/HashMap/RawLemmas.lean b/src/Std/Data/HashMap/RawLemmas.lean index 358af7204a5f..6c1ea5a45d47 100644 --- a/src/Std/Data/HashMap/RawLemmas.lean +++ b/src/Std/Data/HashMap/RawLemmas.lean @@ -998,7 +998,7 @@ theorem getElem?_insertManyIfNewUnit_list [EquivBEq α] [LawfulHashable α] (h : DHashMap.Raw.Const.get?_insertManyIfNewUnit_list h.out @[simp] -theorem ggetElem_insertManyIfNewUnit_list +theorem getElem_insertManyIfNewUnit_list {l : List α} {k : α} {h} : (insertManyIfNewUnit m l)[k] = () := DHashMap.Raw.Const.get_insertManyIfNewUnit_list (h:=h) From 4a24852f166330a48ed7d444e1dcdef279a3cd2d Mon Sep 17 00:00:00 2001 From: jt0202 Date: Mon, 13 Jan 2025 09:18:28 +0100 Subject: [PATCH 78/83] Move ofList out of unverified --- src/Std/Data/DHashMap/Basic.lean | 24 +++++++++++----------- src/Std/Data/DHashMap/Raw.lean | 34 +++++++++++++++++--------------- src/Std/Data/HashMap/Basic.lean | 16 +++++++-------- src/Std/Data/HashMap/Raw.lean | 16 +++++++-------- src/Std/Data/HashSet/Basic.lean | 16 +++++++-------- src/Std/Data/HashSet/Raw.lean | 16 +++++++-------- 6 files changed, 62 insertions(+), 60 deletions(-) diff --git a/src/Std/Data/DHashMap/Basic.lean b/src/Std/Data/DHashMap/Basic.lean index 8ddbfd80216a..1f8ee095f152 100644 --- a/src/Std/Data/DHashMap/Basic.lean +++ b/src/Std/Data/DHashMap/Basic.lean @@ -297,24 +297,12 @@ but will later become a primitive operation. ⟨(Raw₀.Const.insertManyIfNewUnit ⟨m.1, m.2.size_buckets_pos⟩ l).1, (Raw₀.Const.insertManyIfNewUnit ⟨m.1, m.2.size_buckets_pos⟩ l).2 _ Raw.WF.insertIfNew₀ m.2⟩ -@[inline, inherit_doc Raw.ofList] def ofList [BEq α] [Hashable α] (l : List ((a : α) × β a)) : - DHashMap α β := - insertMany ∅ l - /-- Computes the union of the given hash maps, by traversing `m₂` and inserting its elements into `m₁`. -/ @[inline] def union [BEq α] [Hashable α] (m₁ m₂ : DHashMap α β) : DHashMap α β := m₂.fold (init := m₁) fun acc x => acc.insert x instance [BEq α] [Hashable α] : Union (DHashMap α β) := ⟨union⟩ -@[inline, inherit_doc Raw.Const.ofList] def Const.ofList {β : Type v} [BEq α] [Hashable α] - (l : List (α × β)) : DHashMap α (fun _ => β) := - Const.insertMany ∅ l - -@[inline, inherit_doc Raw.Const.unitOfList] def Const.unitOfList [BEq α] [Hashable α] (l : List α) : - DHashMap α (fun _ => Unit) := - Const.insertManyIfNewUnit ∅ l - @[inline, inherit_doc Raw.Const.unitOfArray] def Const.unitOfArray [BEq α] [Hashable α] (l : Array α) : DHashMap α (fun _ => Unit) := Const.insertManyIfNewUnit ∅ l @@ -328,4 +316,16 @@ instance [BEq α] [Hashable α] [Repr α] [(a : α) → Repr (β a)] : Repr (DHa end Unverified +@[inline, inherit_doc Raw.ofList] def ofList [BEq α] [Hashable α] (l : List ((a : α) × β a)) : + DHashMap α β := + insertMany ∅ l + +@[inline, inherit_doc Raw.Const.ofList] def Const.ofList {β : Type v} [BEq α] [Hashable α] + (l : List (α × β)) : DHashMap α (fun _ => β) := + Const.insertMany ∅ l + +@[inline, inherit_doc Raw.Const.unitOfList] def Const.unitOfList [BEq α] [Hashable α] (l : List α) : + DHashMap α (fun _ => Unit) := + Const.insertManyIfNewUnit ∅ l + end Std.DHashMap diff --git a/src/Std/Data/DHashMap/Raw.lean b/src/Std/Data/DHashMap/Raw.lean index 8a6c7788c7b6..5319672f6692 100644 --- a/src/Std/Data/DHashMap/Raw.lean +++ b/src/Std/Data/DHashMap/Raw.lean @@ -422,10 +422,6 @@ This is mainly useful to implement `HashSet.insertMany`, so if you are consideri (Raw₀.Const.insertManyIfNewUnit ⟨m, h⟩ l).1 else m -- will never happen for well-formed inputs -/-- Creates a hash map from a list of mappings. If the same key appears multiple times, the last -occurrence takes precedence. -/ -@[inline] def ofList [BEq α] [Hashable α] (l : List ((a : α) × β a)) : Raw α β := - insertMany ∅ l /-- Computes the union of the given hash maps, by traversing `m₂` and inserting its elements into `m₁`. -/ @[inline] def union [BEq α] [Hashable α] (m₁ m₂ : Raw α β) : Raw α β := @@ -433,18 +429,6 @@ occurrence takes precedence. -/ instance [BEq α] [Hashable α] : Union (Raw α β) := ⟨union⟩ -@[inline, inherit_doc Raw.ofList] def Const.ofList {β : Type v} [BEq α] [Hashable α] - (l : List (α × β)) : Raw α (fun _ => β) := - Const.insertMany ∅ l - -/-- 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. -/ -@[inline] def Const.unitOfList [BEq α] [Hashable α] (l : List α) : - Raw α (fun _ => Unit) := - Const.insertManyIfNewUnit ∅ l - /-- 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, @@ -470,6 +454,24 @@ end Unverified @[inline] def keys (m : Raw α β) : List α := m.foldRev (fun acc k _ => k :: acc) [] +/-- Creates a hash map from a list of mappings. If the same key appears multiple times, the last +occurrence takes precedence. -/ +@[inline] def ofList [BEq α] [Hashable α] (l : List ((a : α) × β a)) : Raw α β := + insertMany ∅ l + + +@[inline, inherit_doc Raw.ofList] def Const.ofList {β : Type v} [BEq α] [Hashable α] + (l : List (α × β)) : Raw α (fun _ => β) := + Const.insertMany ∅ l + +/-- 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. -/ +@[inline] def Const.unitOfList [BEq α] [Hashable α] (l : List α) : + Raw α (fun _ => Unit) := + Const.insertManyIfNewUnit ∅ l + section WF /-- diff --git a/src/Std/Data/HashMap/Basic.lean b/src/Std/Data/HashMap/Basic.lean index 0063ae3e4839..a9060733e4ad 100644 --- a/src/Std/Data/HashMap/Basic.lean +++ b/src/Std/Data/HashMap/Basic.lean @@ -191,6 +191,14 @@ instance [BEq α] [Hashable α] : GetElem? (HashMap α β) α β (fun m a => a @[inline, inherit_doc DHashMap.keys] def keys (m : HashMap α β) : List α := m.inner.keys +@[inline, inherit_doc DHashMap.Const.ofList] def ofList [BEq α] [Hashable α] (l : List (α × β)) : + HashMap α β := + ⟨DHashMap.Const.ofList l⟩ + +@[inline, inherit_doc DHashMap.Const.unitOfList] def unitOfList [BEq α] [Hashable α] (l : List α) : + HashMap α Unit := + ⟨DHashMap.Const.unitOfList l⟩ + section Unverified /-! We currently do not provide lemmas for the functions below. -/ @@ -269,20 +277,12 @@ instance [BEq α] [Hashable α] {m : Type w → Type w} : ForIn m (HashMap α β {ρ : Type w} [ForIn Id ρ α] (m : HashMap α Unit) (l : ρ) : HashMap α Unit := ⟨DHashMap.Const.insertManyIfNewUnit m.inner l⟩ -@[inline, inherit_doc DHashMap.Const.ofList] def ofList [BEq α] [Hashable α] (l : List (α × β)) : - HashMap α β := - ⟨DHashMap.Const.ofList l⟩ - /-- Computes the union of the given hash maps, by traversing `m₂` and inserting its elements into `m₁`. -/ @[inline] def union [BEq α] [Hashable α] (m₁ m₂ : HashMap α β) : HashMap α β := m₂.fold (init := m₁) fun acc x => acc.insert x instance [BEq α] [Hashable α] : Union (HashMap α β) := ⟨union⟩ -@[inline, inherit_doc DHashMap.Const.unitOfList] def unitOfList [BEq α] [Hashable α] (l : List α) : - HashMap α Unit := - ⟨DHashMap.Const.unitOfList l⟩ - @[inline, inherit_doc DHashMap.Const.unitOfArray] def unitOfArray [BEq α] [Hashable α] (l : Array α) : HashMap α Unit := ⟨DHashMap.Const.unitOfArray l⟩ diff --git a/src/Std/Data/HashMap/Raw.lean b/src/Std/Data/HashMap/Raw.lean index 91718cb8f368..5db2a768e678 100644 --- a/src/Std/Data/HashMap/Raw.lean +++ b/src/Std/Data/HashMap/Raw.lean @@ -173,6 +173,14 @@ instance [BEq α] [Hashable α] : GetElem? (Raw α β) α β (fun m a => a ∈ m @[inline, inherit_doc DHashMap.Raw.keys] def keys (m : Raw α β) : List α := m.inner.keys +@[inline, inherit_doc DHashMap.Raw.Const.ofList] def ofList [BEq α] [Hashable α] + (l : List (α × β)) : Raw α β := + ⟨DHashMap.Raw.Const.ofList l⟩ + +@[inline, inherit_doc DHashMap.Raw.Const.unitOfList] def unitOfList [BEq α] [Hashable α] + (l : List α) : Raw α Unit := + ⟨DHashMap.Raw.Const.unitOfList l⟩ + section Unverified /-! We currently do not provide lemmas for the functions below. -/ @@ -233,20 +241,12 @@ m.inner.values [Hashable α] {ρ : Type w} [ForIn Id ρ α] (m : Raw α Unit) (l : ρ) : Raw α Unit := ⟨DHashMap.Raw.Const.insertManyIfNewUnit m.inner l⟩ -@[inline, inherit_doc DHashMap.Raw.Const.ofList] def ofList [BEq α] [Hashable α] - (l : List (α × β)) : Raw α β := - ⟨DHashMap.Raw.Const.ofList l⟩ - /-- Computes the union of the given hash maps, by traversing `m₂` and inserting its elements into `m₁`. -/ @[inline] def union [BEq α] [Hashable α] (m₁ m₂ : Raw α β) : Raw α β := m₂.fold (init := m₁) fun acc x => acc.insert x instance [BEq α] [Hashable α] : Union (Raw α β) := ⟨union⟩ -@[inline, inherit_doc DHashMap.Raw.Const.unitOfList] def unitOfList [BEq α] [Hashable α] - (l : List α) : Raw α Unit := - ⟨DHashMap.Raw.Const.unitOfList l⟩ - @[inline, inherit_doc DHashMap.Raw.Const.unitOfArray] def unitOfArray [BEq α] [Hashable α] (l : Array α) : Raw α Unit := ⟨DHashMap.Raw.Const.unitOfArray l⟩ diff --git a/src/Std/Data/HashSet/Basic.lean b/src/Std/Data/HashSet/Basic.lean index bbbb3517b4dd..2a75fd63fc39 100644 --- a/src/Std/Data/HashSet/Basic.lean +++ b/src/Std/Data/HashSet/Basic.lean @@ -162,6 +162,14 @@ for all `a`. @[inline] def toList (m : HashSet α) : List α := m.inner.keys +/-- +Creates a hash set from a list of elements. Note that unlike repeatedly calling `insert`, if the +collection contains multiple elements that are equal (with regard to `==`), then the last element +in the collection will be present in the returned hash set. +-/ +@[inline] def ofList [BEq α] [Hashable α] (l : List α) : HashSet α := + ⟨HashMap.unitOfList l⟩ + section Unverified /-! We currently do not provide lemmas for the functions below. -/ @@ -233,14 +241,6 @@ appearance. HashSet α := ⟨m.inner.insertManyIfNewUnit l⟩ -/-- -Creates a hash set from a list of elements. Note that unlike repeatedly calling `insert`, if the -collection contains multiple elements that are equal (with regard to `==`), then the last element -in the collection will be present in the returned hash set. --/ -@[inline] def ofList [BEq α] [Hashable α] (l : List α) : HashSet α := - ⟨HashMap.unitOfList l⟩ - /-- Creates a hash set from an array of elements. Note that unlike repeatedly calling `insert`, if the collection contains multiple elements that are equal (with regard to `==`), then the last element diff --git a/src/Std/Data/HashSet/Raw.lean b/src/Std/Data/HashSet/Raw.lean index d0a7dd010fcf..699dd5c8e66d 100644 --- a/src/Std/Data/HashSet/Raw.lean +++ b/src/Std/Data/HashSet/Raw.lean @@ -165,6 +165,14 @@ for all `a`. @[inline] def toList (m : Raw α) : List α := m.inner.keys +/-- +Creates a hash set from a list of elements. Note that unlike repeatedly calling `insert`, if the +collection contains multiple elements that are equal (with regard to `==`), then the last element +in the collection will be present in the returned hash set. +-/ +@[inline] def ofList [BEq α] [Hashable α] (l : List α) : Raw α := + ⟨HashMap.Raw.unitOfList l⟩ + section Unverified /-! We currently do not provide lemmas for the functions below. -/ @@ -231,14 +239,6 @@ appearance. Raw α := ⟨m.inner.insertManyIfNewUnit l⟩ -/-- -Creates a hash set from a list of elements. Note that unlike repeatedly calling `insert`, if the -collection contains multiple elements that are equal (with regard to `==`), then the last element -in the collection will be present in the returned hash set. --/ -@[inline] def ofList [BEq α] [Hashable α] (l : List α) : Raw α := - ⟨HashMap.Raw.unitOfList l⟩ - /-- Creates a hash set from an array of elements. Note that unlike repeatedly calling `insert`, if the collection contains multiple elements that are equal (with regard to `==`), then the last element From 42ee903620b20d2f96af1647d8d313926bb73dba Mon Sep 17 00:00:00 2001 From: jt0202 Date: Tue, 14 Jan 2025 10:57:27 +0100 Subject: [PATCH 79/83] Style fixes --- .../DHashMap/Internal/List/Associative.lean | 31 ++++---- src/Std/Data/DHashMap/Internal/RawLemmas.lean | 71 +++++++++---------- src/Std/Data/DHashMap/Lemmas.lean | 44 ++++++------ src/Std/Data/DHashMap/Raw.lean | 2 - src/Std/Data/DHashMap/RawLemmas.lean | 66 ++++++++--------- src/Std/Data/HashMap/Lemmas.lean | 42 +++++------ src/Std/Data/HashMap/RawLemmas.lean | 48 ++++++------- src/Std/Data/HashSet/Lemmas.lean | 45 ++++++------ src/Std/Data/HashSet/RawLemmas.lean | 36 +++++----- 9 files changed, 190 insertions(+), 195 deletions(-) diff --git a/src/Std/Data/DHashMap/Internal/List/Associative.lean b/src/Std/Data/DHashMap/Internal/List/Associative.lean index 26b849ad39a4..3fb4723cf2ec 100644 --- a/src/Std/Data/DHashMap/Internal/List/Associative.lean +++ b/src/Std/Data/DHashMap/Internal/List/Associative.lean @@ -2493,7 +2493,7 @@ theorem containsKey_of_containsKey_insertListIfNewUnit [BEq α] [PartialEquivBEq theorem getKey?_insertListIfNewUnit_of_contains_eq_false_of_contains_eq_false [BEq α] [EquivBEq α] {l : List ((_ : α) × Unit)} {toInsert : List α} {k : α} - (h : toInsert.contains k = false) (h': containsKey k l = false) : + (h': containsKey k l = false) (h : toInsert.contains k = false) : getKey? k (insertListIfNewUnit l toInsert) = none := by rw [getKey?_eq_getEntry?, getEntry?_insertListIfNewUnit_of_contains_eq_false h, Option.map_eq_none', getEntry?_eq_none] @@ -2502,8 +2502,8 @@ theorem getKey?_insertListIfNewUnit_of_contains_eq_false_of_contains_eq_false [B theorem getKey?_insertListIfNewUnit_of_contains_eq_false_of_mem [BEq α] [EquivBEq α] {l : List ((_ : α) × Unit)} {toInsert : List α} {k k' : α} (k_beq : k == k') - (distinct : toInsert.Pairwise (fun a b => (a == b) = false)) - (mem : k ∈ toInsert) (mem' : containsKey k l = false) : + (mem' : containsKey k l = false) + (distinct : toInsert.Pairwise (fun a b => (a == b) = false)) (mem : k ∈ toInsert): getKey? k' (insertListIfNewUnit l toInsert) = some k := by simp only [getKey?_eq_getEntry?, getEntry?_insertListIfNewUnit, Option.map_eq_some', Option.or_eq_some, getEntry?_eq_none] @@ -2528,12 +2528,12 @@ theorem getKey?_insertListIfNewUnit_of_contains [BEq α] [EquivBEq α] theorem getKey_insertListIfNewUnit_of_contains_eq_false_of_mem [BEq α] [EquivBEq α] {l : List ((_ : α) × Unit)} {toInsert : List α} {k k' : α} (k_beq : k == k') + {h} (contains_eq_false : containsKey k l = false) (distinct : toInsert.Pairwise (fun a b => (a == b) = false)) - (mem : k ∈ toInsert) - {h} (contains_eq_false : containsKey k l = false) : + (mem : k ∈ toInsert) : getKey k' (insertListIfNewUnit l toInsert) h = k := by rw [← Option.some_inj, ← getKey?_eq_some_getKey, - getKey?_insertListIfNewUnit_of_contains_eq_false_of_mem k_beq distinct mem contains_eq_false] + getKey?_insertListIfNewUnit_of_contains_eq_false_of_mem k_beq contains_eq_false distinct mem] theorem getKey_insertListIfNewUnit_of_contains [BEq α] [EquivBEq α] {l : List ((_ : α) × Unit)} {toInsert : List α} @@ -2545,7 +2545,8 @@ theorem getKey_insertListIfNewUnit_of_contains [BEq α] [EquivBEq α] theorem getKey!_insertListIfNewUnit_of_contains_eq_false_of_contains_eq_false [BEq α] [EquivBEq α] [Inhabited α] {l : List ((_ : α) × Unit)} {toInsert : List α} {k : α} - (contains_eq_false : toInsert.contains k = false) (contains_eq_false' : containsKey k l = false): + (contains_eq_false : containsKey k l = false) + (contains_eq_false' : toInsert.contains k = false) : getKey! k (insertListIfNewUnit l toInsert) = default := by rw [getKey!_eq_getKey?, getKey?_insertListIfNewUnit_of_contains_eq_false_of_contains_eq_false contains_eq_false contains_eq_false'] @@ -2554,11 +2555,11 @@ theorem getKey!_insertListIfNewUnit_of_contains_eq_false_of_contains_eq_false [B theorem getKey!_insertListIfNewUnit_of_contains_eq_false_of_mem [BEq α] [EquivBEq α] [Inhabited α] {l : List ((_ : α) × Unit)} {toInsert : List α} {k k' : α} (k_beq : k == k') - (distinct : toInsert.Pairwise (fun a b => (a == b) = false)) - (mem : k ∈ toInsert) (h : containsKey k l = false) : + (h : containsKey k l = false) + (distinct : toInsert.Pairwise (fun a b => (a == b) = false)) (mem : k ∈ toInsert) : getKey! k' (insertListIfNewUnit l toInsert) = k := by rw [getKey!_eq_getKey?, - getKey?_insertListIfNewUnit_of_contains_eq_false_of_mem k_beq distinct mem h, Option.get!_some] + getKey?_insertListIfNewUnit_of_contains_eq_false_of_mem k_beq h distinct mem, Option.get!_some] theorem getKey!_insertListIfNewUnit_of_contains [BEq α] [EquivBEq α] [Inhabited α] {l : List ((_ : α) × Unit)} {toInsert : List α} @@ -2569,7 +2570,7 @@ theorem getKey!_insertListIfNewUnit_of_contains [BEq α] [EquivBEq α] [Inhabite theorem getKeyD_insertListIfNewUnit_of_contains_eq_false_of_contains_eq_false [BEq α] [EquivBEq α] {l : List ((_ : α) × Unit)} {toInsert : List α} {k fallback : α} - (contains_eq_false : toInsert.contains k = false) (contains_eq_false' : containsKey k l = false): + (contains_eq_false : containsKey k l = false) (contains_eq_false' : toInsert.contains k = false) : getKeyD k (insertListIfNewUnit l toInsert) fallback = fallback := by rw [getKeyD_eq_getKey?, getKey?_insertListIfNewUnit_of_contains_eq_false_of_contains_eq_false contains_eq_false contains_eq_false'] @@ -2578,11 +2579,11 @@ theorem getKeyD_insertListIfNewUnit_of_contains_eq_false_of_contains_eq_false [B theorem getKeyD_insertListIfNewUnit_of_contains_eq_false_of_mem [BEq α] [EquivBEq α] {l : List ((_ : α) × Unit)} {toInsert : List α} {k k' fallback : α} (k_beq : k == k') - (distinct : toInsert.Pairwise (fun a b => (a == b) = false)) - (mem : k ∈ toInsert) (h : containsKey k l = false) : + (h : containsKey k l = false) + (distinct : toInsert.Pairwise (fun a b => (a == b) = false)) (mem : k ∈ toInsert) : getKeyD k' (insertListIfNewUnit l toInsert) fallback = k := by - rw [getKeyD_eq_getKey?, getKey?_insertListIfNewUnit_of_contains_eq_false_of_mem k_beq distinct - mem h, Option.getD_some] + rw [getKeyD_eq_getKey?, getKey?_insertListIfNewUnit_of_contains_eq_false_of_mem k_beq h + distinct mem, Option.getD_some] theorem getKeyD_insertListIfNewUnit_of_contains [BEq α] [EquivBEq α] {l : List ((_ : α) × Unit)} {toInsert : List α} diff --git a/src/Std/Data/DHashMap/Internal/RawLemmas.lean b/src/Std/Data/DHashMap/Internal/RawLemmas.lean index ce9ee38cd51b..b99baf9bf643 100644 --- a/src/Std/Data/DHashMap/Internal/RawLemmas.lean +++ b/src/Std/Data/DHashMap/Internal/RawLemmas.lean @@ -914,7 +914,7 @@ theorem get!_insertMany_list_of_mem [LawfulBEq α] (h : m.1.WF) theorem getD_insertMany_list_of_contains_eq_false [LawfulBEq α] (h : m.1.WF) {l : List ((a : α) × β a)} {k : α} {fallback : β k} - (not_mem : (l.map Sigma.fst).contains k = false) : + (contains_eq_false : (l.map Sigma.fst).contains k = false) : (m.insertMany l).1.getD k fallback = m.getD k fallback := by simp_to_model using getValueCastD_insertList_of_contains_eq_false @@ -1191,14 +1191,14 @@ theorem contains_of_contains_insertManyIfNewUnit_list [EquivBEq α] [LawfulHasha theorem getKey?_insertManyIfNewUnit_list_of_contains_eq_false_of_contains_eq_false [EquivBEq α] [LawfulHashable α] - (h : m.1.WF) {l : List α} {k : α} (h' : l.contains k = false) : - m.contains k = false → getKey? (insertManyIfNewUnit m l).1 k = none := by + (h : m.1.WF) {l : List α} {k : α} : + m.contains k = false → l.contains k = false → getKey? (insertManyIfNewUnit m l).1 k = none := by simp_to_model using getKey?_insertListIfNewUnit_of_contains_eq_false_of_contains_eq_false theorem getKey?_insertManyIfNewUnit_list_of_contains_eq_false_of_mem [EquivBEq α] [LawfulHashable α] - (h : m.1.WF) {l : List α} {k k' : α} (k_beq : k == k') - (distinct : l.Pairwise (fun a b => (a == b) = false)) (mem : k ∈ l) : - m.contains k = false → getKey? (insertManyIfNewUnit m l).1 k' = some k := by + (h : m.1.WF) {l : List α} {k k' : α} (k_beq : k == k') : + m.contains k = false → l.Pairwise (fun a b => (a == b) = false) → k ∈ l → + getKey? (insertManyIfNewUnit m l).1 k' = some k := by simp_to_model using getKey?_insertListIfNewUnit_of_contains_eq_false_of_mem theorem getKey?_insertManyIfNewUnit_list_of_contains [EquivBEq α] [LawfulHashable α] @@ -1213,10 +1213,9 @@ theorem getKey_insertManyIfNewUnit_list_of_contains [EquivBEq α] [LawfulHashabl theorem getKey_insertManyIfNewUnit_list_of_contains_eq_false_of_mem [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l : List α} - {k k' : α} (k_beq : k == k') - (distinct : l.Pairwise (fun a b => (a == b) = false)) - (mem : k ∈ l) {h'} : - m.contains k = false → getKey (insertManyIfNewUnit m l).1 k' h' = k := by + {k k' : α} (k_beq : k == k') {h'} : + m.contains k = false → l.Pairwise (fun a b => (a == b) = false) → k ∈ l → + getKey (insertManyIfNewUnit m l).1 k' h' = k := by simp_to_model using getKey_insertListIfNewUnit_of_contains_eq_false_of_mem theorem getKey_insertManyIfNewUnit_list_mem_of_contains [EquivBEq α] [LawfulHashable α] @@ -1225,16 +1224,15 @@ theorem getKey_insertManyIfNewUnit_list_mem_of_contains [EquivBEq α] [LawfulHas simp_to_model using getKey_insertListIfNewUnit_of_contains theorem getKey!_insertManyIfNewUnit_list_of_contains_eq_false_of_contains_eq_false - [EquivBEq α] [LawfulHashable α] [Inhabited α] (h : m.1.WF) {l : List α} {k : α} - (contains_eq_false : l.contains k = false) : - m.contains k = false → getKey! (insertManyIfNewUnit m l).1 k = default := by + [EquivBEq α] [LawfulHashable α] [Inhabited α] (h : m.1.WF) {l : List α} {k : α} : + m.contains k = false → l.contains k = false → + getKey! (insertManyIfNewUnit m l).1 k = default := by simp_to_model using getKey!_insertListIfNewUnit_of_contains_eq_false_of_contains_eq_false theorem getKey!_insertManyIfNewUnit_list_of_contains_eq_false_of_mem [EquivBEq α] [LawfulHashable α] - [Inhabited α] (h : m.1.WF) {l : List α} {k k' : α} (k_beq : k == k') - (distinct : l.Pairwise (fun a b => (a == b) = false)) - (mem : k ∈ l) : - contains m k = false → getKey! (insertManyIfNewUnit m l).1 k' = k := by + [Inhabited α] (h : m.1.WF) {l : List α} {k k' : α} (k_beq : k == k') : + contains m k = false → l.Pairwise (fun a b => (a == b) = false) → k ∈ l → + getKey! (insertManyIfNewUnit m l).1 k' = k := by simp_to_model using getKey!_insertListIfNewUnit_of_contains_eq_false_of_mem theorem getKey!_insertManyIfNewUnit_list_of_contains [EquivBEq α] [LawfulHashable α] @@ -1243,16 +1241,14 @@ theorem getKey!_insertManyIfNewUnit_list_of_contains [EquivBEq α] [LawfulHashab simp_to_model using getKey!_insertListIfNewUnit_of_contains theorem getKeyD_insertManyIfNewUnit_list_of_contains_eq_false_of_contains_eq_false - [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l : List α} {k fallback : α} - (h' : l.contains k = false) : - m.contains k = false → getKeyD (insertManyIfNewUnit m l).1 k fallback = fallback := by + [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l : List α} {k fallback : α} : + m.contains k = false → l.contains k = false → getKeyD (insertManyIfNewUnit m l).1 k fallback = fallback := by simp_to_model using getKeyD_insertListIfNewUnit_of_contains_eq_false_of_contains_eq_false theorem getKeyD_insertManyIfNewUnit_list_of_contains_eq_false_of_mem [EquivBEq α] [LawfulHashable α] - (h : m.1.WF) {l : List α} {k k' fallback : α} (k_beq : k == k') - (distinct : l.Pairwise (fun a b => (a == b) = false)) - (mem : k ∈ l) : - m.contains k = false → getKeyD (insertManyIfNewUnit m l).1 k' fallback = k := by + (h : m.1.WF) {l : List α} {k k' fallback : α} (k_beq : k == k') : + m.contains k = false → l.Pairwise (fun a b => (a == b) = false) → k ∈ l → + getKeyD (insertManyIfNewUnit m l).1 k' fallback = k := by simp_to_model using getKeyD_insertListIfNewUnit_of_contains_eq_false_of_mem theorem getKeyD_insertManyIfNewUnit_list_of_contains [EquivBEq α] [LawfulHashable α] @@ -1618,15 +1614,15 @@ theorem contains_insertManyIfNewUnit_empty_list [EquivBEq α] [LawfulHashable α theorem getKey?_insertManyIfNewUnit_empty_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] {l : List α} {k : α} (h' : l.contains k = false) : getKey? (insertManyIfNewUnit (empty : Raw₀ α (fun _ => Unit)) l).1 k = none := by - simp [getKey?_insertManyIfNewUnit_list_of_contains_eq_false_of_contains_eq_false - _ Raw.WF.empty₀ h'] + exact getKey?_insertManyIfNewUnit_list_of_contains_eq_false_of_contains_eq_false _ Raw.WF.empty₀ + contains_empty h' theorem getKey?_insertManyIfNewUnit_empty_list_of_mem [EquivBEq α] [LawfulHashable α] {l : List α} {k k' : α} (k_beq : k == k') (distinct : l.Pairwise (fun a b => (a == b) = false)) (mem : k ∈ l) : getKey? (insertManyIfNewUnit (empty : Raw₀ α (fun _ => Unit)) l).1 k' = some k := by - simp [getKey?_insertManyIfNewUnit_list_of_contains_eq_false_of_mem _ Raw.WF.empty₀ k_beq - distinct mem contains_empty] + exact getKey?_insertManyIfNewUnit_list_of_contains_eq_false_of_mem _ Raw.WF.empty₀ k_beq + contains_empty distinct mem theorem getKey_insertManyIfNewUnit_empty_list_of_mem [EquivBEq α] [LawfulHashable α] {l : List α} @@ -1634,37 +1630,38 @@ theorem getKey_insertManyIfNewUnit_empty_list_of_mem [EquivBEq α] [LawfulHashab (distinct : l.Pairwise (fun a b => (a == b) = false)) (mem : k ∈ l) {h'} : getKey (insertManyIfNewUnit (empty : Raw₀ α (fun _ => Unit)) l).1 k' h' = k := by - simp [getKey_insertManyIfNewUnit_list_of_contains_eq_false_of_mem _ Raw.WF.empty₀ k_beq - distinct mem contains_empty] + exact getKey_insertManyIfNewUnit_list_of_contains_eq_false_of_mem _ Raw.WF.empty₀ k_beq + contains_empty distinct mem theorem getKey!_insertManyIfNewUnit_empty_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] [Inhabited α] {l : List α} {k : α} (h' : l.contains k = false) : getKey! (insertManyIfNewUnit (empty : Raw₀ α (fun _ => Unit)) l).1 k = default := by - simp [getKey!_insertManyIfNewUnit_list_of_contains_eq_false_of_contains_eq_false _ Raw.WF.empty₀ h'] + exact getKey!_insertManyIfNewUnit_list_of_contains_eq_false_of_contains_eq_false _ Raw.WF.empty₀ + contains_empty h' theorem getKey!_insertManyIfNewUnit_empty_list_of_mem [EquivBEq α] [LawfulHashable α] [Inhabited α] {l : List α} {k k' : α} (k_beq : k == k') (distinct : l.Pairwise (fun a b => (a == b) = false)) (mem : k ∈ l) : getKey! (insertManyIfNewUnit (empty : Raw₀ α (fun _ => Unit)) l).1 k' = k := by - simp [getKey!_insertManyIfNewUnit_list_of_contains_eq_false_of_mem _ Raw.WF.empty₀ k_beq - distinct mem contains_empty] + exact getKey!_insertManyIfNewUnit_list_of_contains_eq_false_of_mem _ Raw.WF.empty₀ k_beq + contains_empty distinct mem theorem getKeyD_insertManyIfNewUnit_empty_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] {l : List α} {k fallback : α} (h' : l.contains k = false) : getKeyD (insertManyIfNewUnit (empty : Raw₀ α (fun _ => Unit)) l).1 k fallback = fallback := by - simp [getKeyD_insertManyIfNewUnit_list_of_contains_eq_false_of_contains_eq_false - _ Raw.WF.empty₀ h'] + exact getKeyD_insertManyIfNewUnit_list_of_contains_eq_false_of_contains_eq_false + _ Raw.WF.empty₀ contains_empty h' theorem getKeyD_insertManyIfNewUnit_empty_list_of_mem [EquivBEq α] [LawfulHashable α] {l : List α} {k k' fallback : α} (k_beq : k == k') (distinct : l.Pairwise (fun a b => (a == b) = false)) (mem : k ∈ l) : getKeyD (insertManyIfNewUnit (empty : Raw₀ α (fun _ => Unit)) l).1 k' fallback = k := by - simp [getKeyD_insertManyIfNewUnit_list_of_contains_eq_false_of_mem _ Raw.WF.empty₀ k_beq - distinct mem contains_empty] + exact getKeyD_insertManyIfNewUnit_list_of_contains_eq_false_of_mem _ Raw.WF.empty₀ k_beq + contains_empty distinct mem theorem size_insertManyIfNewUnit_empty_list [EquivBEq α] [LawfulHashable α] {l : List α} diff --git a/src/Std/Data/DHashMap/Lemmas.lean b/src/Std/Data/DHashMap/Lemmas.lean index 747815e1e3f1..2294cfc9d250 100644 --- a/src/Std/Data/DHashMap/Lemmas.lean +++ b/src/Std/Data/DHashMap/Lemmas.lean @@ -1111,7 +1111,7 @@ theorem getKeyD_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] theorem size_insertMany_list [EquivBEq α] [LawfulHashable α] {l : List ((a : α) × β a)} (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) : - (∀ (a : α), m.contains a → (l.map Sigma.fst).contains a = false) → + (∀ (a : α), a ∈ m → (l.map Sigma.fst).contains a = false) → (m.insertMany l).size = m.size + l.length := Raw₀.size_insertMany_list ⟨m.1, _⟩ m.2 distinct @@ -1158,7 +1158,7 @@ theorem mem_insertMany_list [EquivBEq α] [LawfulHashable α] simp [mem_iff_contains] theorem mem_of_mem_insertMany_list [EquivBEq α] [LawfulHashable α] - {l : List ( α × β )} {k : α} (mem : k ∈ (insertMany m l)) + {l : List (α × β)} {k : α} (mem : k ∈ insertMany m l) (contains_eq_false : (l.map Prod.fst).contains k = false) : k ∈ m := Raw₀.Const.contains_of_contains_insertMany_list ⟨m.1, _⟩ m.2 mem contains_eq_false @@ -1225,7 +1225,7 @@ theorem getKeyD_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] theorem size_insertMany_list [EquivBEq α] [LawfulHashable α] {l : List (α × β)} (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) : - (∀ (a : α), m.contains a → (l.map Prod.fst).contains a = false) → + (∀ (a : α), a ∈ m → (l.map Prod.fst).contains a = false) → (insertMany m l).size = m.size + l.length := Raw₀.Const.size_insertMany_list ⟨m.1, _⟩ m.2 distinct @@ -1322,23 +1322,23 @@ theorem mem_insertManyIfNewUnit_list [EquivBEq α] [LawfulHashable α] theorem mem_of_mem_insertManyIfNewUnit_list [EquivBEq α] [LawfulHashable α] {l : List α} {k : α} (contains_eq_false : l.contains k = false) : - k ∈ (insertManyIfNewUnit m l) → k ∈ m := + k ∈ insertManyIfNewUnit m l → k ∈ m := Raw₀.Const.contains_of_contains_insertManyIfNewUnit_list ⟨m.1, _⟩ m.2 contains_eq_false theorem getKey?_insertManyIfNewUnit_list_of_contains_eq_false_of_contains_eq_false [EquivBEq α] [LawfulHashable α] {l : List α} {k : α} - (contains_eq_false : m.contains k = false) (contains_eq_false' : l.contains k = false): + (contains_eq_false : m.contains k = false) (contains_eq_false' : l.contains k = false) : getKey? (insertManyIfNewUnit m l) k = none := Raw₀.Const.getKey?_insertManyIfNewUnit_list_of_contains_eq_false_of_contains_eq_false - ⟨m.1, _⟩ m.2 contains_eq_false' contains_eq_false + ⟨m.1, _⟩ m.2 contains_eq_false contains_eq_false' theorem getKey?_insertManyIfNewUnit_list_of_contains_eq_false_of_mem [EquivBEq α] [LawfulHashable α] {l : List α} {k k' : α} (k_beq : k == k') - (distinct : l.Pairwise (fun a b => (a == b) = false)) - (contains_eq_false : m.contains k = false) (mem : k ∈ l) : + (contains_eq_false : m.contains k = false) + (distinct : l.Pairwise (fun a b => (a == b) = false)) (mem : k ∈ l) : getKey? (insertManyIfNewUnit m l) k' = some k := Raw₀.Const.getKey?_insertManyIfNewUnit_list_of_contains_eq_false_of_mem ⟨m.1, _⟩ - m.2 k_beq distinct mem contains_eq_false + m.2 k_beq contains_eq_false distinct mem theorem getKey?_insertManyIfNewUnit_list_of_mem [EquivBEq α] [LawfulHashable α] {l : List α} {k : α} (h' : k ∈ m) : @@ -1348,11 +1348,11 @@ theorem getKey?_insertManyIfNewUnit_list_of_mem [EquivBEq α] [LawfulHashable α theorem getKey_insertManyIfNewUnit_list_of_contains_eq_false_of_mem [EquivBEq α] [LawfulHashable α] {l : List α} {k k' : α} (k_beq : k == k') - (distinct : l.Pairwise (fun a b => (a == b) = false)) (contains_eq_false : m.contains k = false) + (contains_eq_false : m.contains k = false) (distinct : l.Pairwise (fun a b => (a == b) = false)) (mem : k ∈ l) {h} : getKey (insertManyIfNewUnit m l) k' h = k := Raw₀.Const.getKey_insertManyIfNewUnit_list_of_contains_eq_false_of_mem ⟨m.1, _⟩ m.2 k_beq - distinct mem contains_eq_false + contains_eq_false distinct mem theorem getKey_insertManyIfNewUnit_list_of_mem [EquivBEq α] [LawfulHashable α] {l : List α} {k : α} (mem : k ∈ m) {h} : @@ -1364,18 +1364,18 @@ theorem getKey!_insertManyIfNewUnit_list_of_contains_eq_false_of_contains_eq_fal (contains_eq_false : m.contains k = false) (contains_eq_false' : l.contains k = false) : getKey! (insertManyIfNewUnit m l) k = default := Raw₀.Const.getKey!_insertManyIfNewUnit_list_of_contains_eq_false_of_contains_eq_false - ⟨m.1, _⟩ m.2 contains_eq_false' contains_eq_false + ⟨m.1, _⟩ m.2 contains_eq_false contains_eq_false' theorem getKey!_insertManyIfNewUnit_list_of_contains_eq_false_of_mem [EquivBEq α] [LawfulHashable α] [Inhabited α] {l : List α} {k k' : α} (k_beq : k == k') - (distinct : l.Pairwise (fun a b => (a == b) = false)) (contains_eq_false: m.contains k = false) + (contains_eq_false: m.contains k = false) (distinct : l.Pairwise (fun a b => (a == b) = false)) (mem : k ∈ l) : getKey! (insertManyIfNewUnit m l) k' = k := Raw₀.Const.getKey!_insertManyIfNewUnit_list_of_contains_eq_false_of_mem ⟨m.1, _⟩ m.2 k_beq - distinct mem contains_eq_false + contains_eq_false distinct mem theorem getKey!_insertManyIfNewUnit_list_of_mem [EquivBEq α] [LawfulHashable α] - [Inhabited α] {l : List α} {k : α} (mem : k ∈ m): + [Inhabited α] {l : List α} {k : α} (mem : k ∈ m) : getKey! (insertManyIfNewUnit m l) k = getKey! m k := Raw₀.Const.getKey!_insertManyIfNewUnit_list_of_contains ⟨m.1, _⟩ m.2 mem @@ -1384,25 +1384,25 @@ theorem getKeyD_insertManyIfNewUnit_list_of_contains_eq_false_of_contains_eq_fal (contains_eq_false : m.contains k = false) (contains_eq_false' : l.contains k = false) : getKeyD (insertManyIfNewUnit m l) k fallback = fallback := Raw₀.Const.getKeyD_insertManyIfNewUnit_list_of_contains_eq_false_of_contains_eq_false - ⟨m.1, _⟩ m.2 contains_eq_false' contains_eq_false + ⟨m.1, _⟩ m.2 contains_eq_false contains_eq_false' theorem getKeyD_insertManyIfNewUnit_list_of_contains_eq_false_of_mem [EquivBEq α] [LawfulHashable α] {l : List α} {k k' fallback : α} (k_beq : k == k') - (distinct : l.Pairwise (fun a b => (a == b) = false)) - (contains_eq_false : m.contains k = false) (mem : k ∈ l ) : + (contains_eq_false : m.contains k = false) + (distinct : l.Pairwise (fun a b => (a == b) = false)) (mem : k ∈ l) : getKeyD (insertManyIfNewUnit m l) k' fallback = k := Raw₀.Const.getKeyD_insertManyIfNewUnit_list_of_contains_eq_false_of_mem ⟨m.1, _⟩ m.2 k_beq - distinct mem contains_eq_false + contains_eq_false distinct mem theorem getKeyD_insertManyIfNewUnit_list_of_mem [EquivBEq α] [LawfulHashable α] - {l : List α} {k fallback : α} (mem : k ∈ m): + {l : List α} {k fallback : α} (mem : k ∈ m) : getKeyD (insertManyIfNewUnit m l) k fallback = getKeyD m k fallback := Raw₀.Const.getKeyD_insertManyIfNewUnit_list_of_contains ⟨m.1, _⟩ m.2 mem theorem size_insertManyIfNewUnit_list [EquivBEq α] [LawfulHashable α] {l : List α} (distinct : l.Pairwise (fun a b => (a == b) = false)) : - (∀ (a : α), m.contains a → l.contains a = false) → + (∀ (a : α), a ∈ m → l.contains a = false) → (insertManyIfNewUnit m l).size = m.size + l.length := Raw₀.Const.size_insertManyIfNewUnit_list ⟨m.1, _⟩ m.2 distinct @@ -1420,7 +1420,7 @@ theorem isEmpty_insertManyIfNewUnit_list [EquivBEq α] [LawfulHashable α] theorem get?_insertManyIfNewUnit_list [EquivBEq α] [LawfulHashable α] {l : List α} {k : α} : get? (insertManyIfNewUnit m l) k = - if m.contains k ∨ l.contains k then some () else none := + if k ∈ m ∨ l.contains k then some () else none := Raw₀.Const.get?_insertManyIfNewUnit_list ⟨m.1, _⟩ m.2 theorem get_insertManyIfNewUnit_list diff --git a/src/Std/Data/DHashMap/Raw.lean b/src/Std/Data/DHashMap/Raw.lean index 5319672f6692..3f841ef563ab 100644 --- a/src/Std/Data/DHashMap/Raw.lean +++ b/src/Std/Data/DHashMap/Raw.lean @@ -422,7 +422,6 @@ This is mainly useful to implement `HashSet.insertMany`, so if you are consideri (Raw₀.Const.insertManyIfNewUnit ⟨m, h⟩ l).1 else m -- will never happen for well-formed inputs - /-- Computes the union of the given hash maps, by traversing `m₂` and inserting its elements into `m₁`. -/ @[inline] def union [BEq α] [Hashable α] (m₁ m₂ : Raw α β) : Raw α β := m₂.fold (init := m₁) fun acc x => acc.insert x @@ -459,7 +458,6 @@ occurrence takes precedence. -/ @[inline] def ofList [BEq α] [Hashable α] (l : List ((a : α) × β a)) : Raw α β := insertMany ∅ l - @[inline, inherit_doc Raw.ofList] def Const.ofList {β : Type v} [BEq α] [Hashable α] (l : List (α × β)) : Raw α (fun _ => β) := Const.insertMany ∅ l diff --git a/src/Std/Data/DHashMap/RawLemmas.lean b/src/Std/Data/DHashMap/RawLemmas.lean index 6fb514e196dd..c533e3e9be3b 100644 --- a/src/Std/Data/DHashMap/RawLemmas.lean +++ b/src/Std/Data/DHashMap/RawLemmas.lean @@ -1203,8 +1203,9 @@ theorem getKeyD_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] (h : m. theorem size_insertMany_list [EquivBEq α] [LawfulHashable α] (h : m.WF) {l : List ((a : α) × β a)} (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) : - (∀ (a : α), m.contains a → (l.map Sigma.fst).contains a = false) → + (∀ (a : α), a ∈ m → (l.map Sigma.fst).contains a = false) → (m.insertMany l).size = m.size + l.length := by + simp only [mem_iff_contains] simp_to_raw using Raw₀.size_insertMany_list theorem size_insertMany_list_le [EquivBEq α] [LawfulHashable α] (h : m.WF) @@ -1250,12 +1251,12 @@ theorem contains_insertMany_list [EquivBEq α] [LawfulHashable α] (h : m.WF) @[simp] theorem mem_insertMany_list [EquivBEq α] [LawfulHashable α] (h : m.WF) {l : List (α × β)} {k : α} : - k ∈ (insertMany m l) ↔ k ∈ m ∨ (l.map Prod.fst).contains k := by + k ∈ insertMany m l ↔ k ∈ m ∨ (l.map Prod.fst).contains k := by simp [mem_iff_contains, contains_insertMany_list h] theorem mem_of_mem_insertMany_list [EquivBEq α] [LawfulHashable α] (h : m.WF) - {l : List ( α × β )} {k : α} : - k ∈ (insertMany m l) → (l.map Prod.fst).contains k = false → k ∈ m := by + {l : List (α × β)} {k : α} : + k ∈ insertMany m l → (l.map Prod.fst).contains k = false → k ∈ m := by simp only [mem_iff_contains] simp_to_raw using Raw₀.Const.contains_of_contains_insertMany_list @@ -1321,8 +1322,9 @@ theorem getKeyD_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] (h : m. theorem size_insertMany_list [EquivBEq α] [LawfulHashable α] (h : m.WF) {l : List (α × β)} (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) : - (∀ (a : α), m.contains a → (l.map Prod.fst).contains a = false) → + (∀ (a : α), a ∈ m → (l.map Prod.fst).contains a = false) → (insertMany m l).size = m.size + l.length := by + simp [mem_iff_contains] simp_to_raw using Raw₀.Const.size_insertMany_list theorem size_insertMany_list_le [EquivBEq α] [LawfulHashable α] (h : m.WF) @@ -1419,20 +1421,20 @@ theorem mem_insertManyIfNewUnit_list [EquivBEq α] [LawfulHashable α] (h : m.WF theorem mem_of_mem_insertManyIfNewUnit_list [EquivBEq α] [LawfulHashable α] (h : m.WF) {l : List α} {k : α} (contains_eq_false : l.contains k = false) : - k ∈ (insertManyIfNewUnit m l) → k ∈ m := by + k ∈ insertManyIfNewUnit m l → k ∈ m := by simp only [mem_iff_contains] simp_to_raw using Raw₀.Const.contains_of_contains_insertManyIfNewUnit_list theorem getKey?_insertManyIfNewUnit_list_of_contains_eq_false_of_contains_eq_false - [EquivBEq α] [LawfulHashable α] (h : m.WF) {l : List α} {k : α} - (contains_eq_false : l.contains k = false) : - m.contains k = false → getKey? (insertManyIfNewUnit m l) k = none := by + [EquivBEq α] [LawfulHashable α] (h : m.WF) {l : List α} {k : α} : + m.contains k = false → l.contains k = false → + getKey? (insertManyIfNewUnit m l) k = none := by simp_to_raw using Raw₀.Const.getKey?_insertManyIfNewUnit_list_of_contains_eq_false_of_contains_eq_false theorem getKey?_insertManyIfNewUnit_list_of_contains_eq_false_of_mem [EquivBEq α] [LawfulHashable α] - (h : m.WF) {l : List α} {k k' : α} (k_beq : k == k') - (distinct : l.Pairwise (fun a b => (a == b) = false)) (mem : k ∈ l) : - m.contains k = false → getKey? (insertManyIfNewUnit m l) k' = some k := by + (h : m.WF) {l : List α} {k k' : α} (k_beq : k == k') : + m.contains k = false → l.Pairwise (fun a b => (a == b) = false) → k ∈ l → + getKey? (insertManyIfNewUnit m l) k' = some k := by simp_to_raw using Raw₀.Const.getKey?_insertManyIfNewUnit_list_of_contains_eq_false_of_mem theorem getKey?_insertManyIfNewUnit_list_of_mem [EquivBEq α] [LawfulHashable α] @@ -1442,30 +1444,27 @@ theorem getKey?_insertManyIfNewUnit_list_of_mem [EquivBEq α] [LawfulHashable α simp_to_raw using Raw₀.Const.getKey?_insertManyIfNewUnit_list_of_contains theorem getKey_insertManyIfNewUnit_list_of_mem - [EquivBEq α] [LawfulHashable α] (h : m.WF) {l : List α} {k : α} {h'} (mem : k ∈ m): + [EquivBEq α] [LawfulHashable α] (h : m.WF) {l : List α} {k : α} {h'} (mem : k ∈ m) : getKey (insertManyIfNewUnit m l) k h' = getKey m k mem := by simp_to_raw using Raw₀.Const.getKey_insertManyIfNewUnit_list_of_contains theorem getKey_insertManyIfNewUnit_list_of_contains_eq_false_of_mem [EquivBEq α] [LawfulHashable α] (h : m.WF) {l : List α} - {k k' : α} (k_beq : k == k') - (distinct : l.Pairwise (fun a b => (a == b) = false)) - (mem : k ∈ l) {h'} : - m.contains k = false → getKey (insertManyIfNewUnit m l) k' h' = k := by + {k k' : α} (k_beq : k == k') {h'} : + m.contains k = false → l.Pairwise (fun a b => (a == b) = false) → k ∈ l → + getKey (insertManyIfNewUnit m l) k' h' = k := by simp_to_raw using Raw₀.Const.getKey_insertManyIfNewUnit_list_of_contains_eq_false_of_mem theorem getKey!_insertManyIfNewUnit_list_of_contains_eq_false_of_contains_eq_false - [EquivBEq α] [LawfulHashable α] [Inhabited α] (h : m.WF) {l : List α} {k : α} - (contains_eq_false : l.contains k = false) : - m.contains k = false → getKey! (insertManyIfNewUnit m l) k = default := by + [EquivBEq α] [LawfulHashable α] [Inhabited α] (h : m.WF) {l : List α} {k : α} : + m.contains k = false → l.contains k = false → getKey! (insertManyIfNewUnit m l) k = default := by simp_to_raw using Raw₀.Const.getKey!_insertManyIfNewUnit_list_of_contains_eq_false_of_contains_eq_false theorem getKey!_insertManyIfNewUnit_list_of_contains_eq_false_of_mem [EquivBEq α] [LawfulHashable α] - [Inhabited α] (h : m.WF) {l : List α} {k k' : α} (k_beq : k == k') - (distinct : l.Pairwise (fun a b => (a == b) = false)) - (mem : k ∈ l) : - contains m k = false → getKey! (insertManyIfNewUnit m l) k' = k := by + [Inhabited α] (h : m.WF) {l : List α} {k k' : α} (k_beq : k == k') : + contains m k = false → l.Pairwise (fun a b => (a == b) = false) → k ∈ l → + getKey! (insertManyIfNewUnit m l) k' = k := by simp_to_raw using Raw₀.Const.getKey!_insertManyIfNewUnit_list_of_contains_eq_false_of_mem theorem getKey!_insertManyIfNewUnit_list_of_mem [EquivBEq α] [LawfulHashable α] @@ -1475,17 +1474,16 @@ theorem getKey!_insertManyIfNewUnit_list_of_mem [EquivBEq α] [LawfulHashable α simp_to_raw using Raw₀.Const.getKey!_insertManyIfNewUnit_list_of_contains theorem getKeyD_insertManyIfNewUnit_list_of_contains_eq_false_of_contains_eq_false - [EquivBEq α] [LawfulHashable α] (h : m.WF) {l : List α} {k fallback : α} - (contains_eq_false : l.contains k = false) : - m.contains k = false → getKeyD (insertManyIfNewUnit m l) k fallback = fallback := by + [EquivBEq α] [LawfulHashable α] (h : m.WF) {l : List α} {k fallback : α} : + m.contains k = false → l.contains k = false → + getKeyD (insertManyIfNewUnit m l) k fallback = fallback := by simp_to_raw using Raw₀.Const.getKeyD_insertManyIfNewUnit_list_of_contains_eq_false_of_contains_eq_false theorem getKeyD_insertManyIfNewUnit_list_of_contains_eq_false_of_mem [EquivBEq α] [LawfulHashable α] - (h : m.WF) {l : List α} {k k' fallback : α} (k_beq : k == k') - (distinct : l.Pairwise (fun a b => (a == b) = false)) - (mem : k ∈ l) : - m.contains k = false → getKeyD (insertManyIfNewUnit m l) k' fallback = k := by + (h : m.WF) {l : List α} {k k' fallback : α} (k_beq : k == k') : + m.contains k = false → l.Pairwise (fun a b => (a == b) = false) → k ∈ l → + getKeyD (insertManyIfNewUnit m l) k' fallback = k := by simp_to_raw using Raw₀.Const.getKeyD_insertManyIfNewUnit_list_of_contains_eq_false_of_mem theorem getKeyD_insertManyIfNewUnit_list_of_mem [EquivBEq α] [LawfulHashable α] @@ -1497,8 +1495,9 @@ theorem getKeyD_insertManyIfNewUnit_list_of_mem [EquivBEq α] [LawfulHashable α theorem size_insertManyIfNewUnit_list [EquivBEq α] [LawfulHashable α] (h : m.WF) {l : List α} (distinct : l.Pairwise (fun a b => (a == b) = false)) : - (∀ (a : α), m.contains a → l.contains a = false) → + (∀ (a : α), a ∈ m → l.contains a = false) → (insertManyIfNewUnit m l).size = m.size + l.length := by + simp only [mem_iff_contains] simp_to_raw using Raw₀.Const.size_insertManyIfNewUnit_list theorem size_insertManyIfNewUnit_list_le [EquivBEq α] [LawfulHashable α] (h : m.WF) @@ -1516,7 +1515,8 @@ theorem isEmpty_insertManyIfNewUnit_list [EquivBEq α] [LawfulHashable α] (h : theorem get?_insertManyIfNewUnit_list [EquivBEq α] [LawfulHashable α] (h : m.WF) {l : List α} {k : α} : get? (insertManyIfNewUnit m l) k = - if m.contains k ∨ l.contains k then some () else none := by + if k ∈ m ∨ l.contains k then some () else none := by + simp only [mem_iff_contains] simp_to_raw using Raw₀.Const.get?_insertManyIfNewUnit_list @[simp] diff --git a/src/Std/Data/HashMap/Lemmas.lean b/src/Std/Data/HashMap/Lemmas.lean index f172ff1a1d3c..be3e42211c4f 100644 --- a/src/Std/Data/HashMap/Lemmas.lean +++ b/src/Std/Data/HashMap/Lemmas.lean @@ -725,7 +725,7 @@ theorem mem_insertMany_list [EquivBEq α] [LawfulHashable α] DHashMap.Const.mem_insertMany_list theorem mem_of_mem_insertMany_list [EquivBEq α] [LawfulHashable α] - {l : List (α × β)} {k : α} (mem : k ∈ (insertMany m l)) + {l : List (α × β)} {k : α} (mem : k ∈ insertMany m l) (contains_eq_false : (l.map Prod.fst).contains k = false) : k ∈ m := DHashMap.Const.mem_of_mem_insertMany_list mem contains_eq_false @@ -841,7 +841,7 @@ theorem getKeyD_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] theorem size_insertMany_list [EquivBEq α] [LawfulHashable α] {l : List (α × β)} (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) : - (∀ (a : α), m.contains a → (l.map Prod.fst).contains a = false) → + (∀ (a : α), a ∈ m → (l.map Prod.fst).contains a = false) → (insertMany m l).size = m.size + l.length := DHashMap.Const.size_insertMany_list distinct @@ -886,13 +886,13 @@ theorem mem_insertManyIfNewUnit_list [EquivBEq α] [LawfulHashable α] theorem mem_of_mem_insertManyIfNewUnit_list [EquivBEq α] [LawfulHashable α] {l : List α} {k : α} (contains_eq_false : l.contains k = false) : - k ∈ (insertManyIfNewUnit m l) → k ∈ m := + k ∈ insertManyIfNewUnit m l → k ∈ m := DHashMap.Const.mem_of_mem_insertManyIfNewUnit_list contains_eq_false theorem getElem?_insertManyIfNewUnit_list [EquivBEq α] [LawfulHashable α] {l : List α} {k : α} : (insertManyIfNewUnit m l)[k]? = - if m.contains k ∨ l.contains k then some () else none := + if k ∈ m ∨ l.contains k then some () else none := DHashMap.Const.get?_insertManyIfNewUnit_list theorem getElem_insertManyIfNewUnit_list @@ -918,11 +918,11 @@ theorem getKey?_insertManyIfNewUnit_list_of_contains_eq_false_of_contains_eq_fal contains_eq_false contains_eq_false' theorem getKey?_insertManyIfNewUnit_list_of_contains_eq_false_of_mem [EquivBEq α] [LawfulHashable α] - {l : List α} {k k' : α} (k_beq : k == k') (distinct : l.Pairwise (fun a b => (a == b) = false)) - (contains_eq_false : m.contains k = false) (mem : k ∈ l) : + {l : List α} {k k' : α} (k_beq : k == k') (contains_eq_false : m.contains k = false) + (distinct : l.Pairwise (fun a b => (a == b) = false)) (mem : k ∈ l) : getKey? (insertManyIfNewUnit m l) k' = some k := DHashMap.Const.getKey?_insertManyIfNewUnit_list_of_contains_eq_false_of_mem - k_beq distinct contains_eq_false mem + k_beq contains_eq_false distinct mem theorem getKey?_insertManyIfNewUnit_list_of_mem [EquivBEq α] [LawfulHashable α] {l : List α} {k : α} (mem : k ∈ m) : @@ -930,13 +930,13 @@ theorem getKey?_insertManyIfNewUnit_list_of_mem [EquivBEq α] [LawfulHashable α DHashMap.Const.getKey?_insertManyIfNewUnit_list_of_mem mem theorem getKey_insertManyIfNewUnit_list_of_contains_eq_false_of_mem [EquivBEq α] [LawfulHashable α] - {l : List α} {k k' : α} (k_beq : k == k') (distinct : l.Pairwise (fun a b => (a == b) = false)) - (contains_eq_false : m.contains k = false) (mem : k ∈ l) {h} : + {l : List α} {k k' : α} (k_beq : k == k') (contains_eq_false : m.contains k = false) + (distinct : l.Pairwise (fun a b => (a == b) = false)) (mem : k ∈ l) {h} : getKey (insertManyIfNewUnit m l) k' h = k := DHashMap.Const.getKey_insertManyIfNewUnit_list_of_contains_eq_false_of_mem - k_beq distinct contains_eq_false mem + k_beq contains_eq_false distinct mem -theorem getKey_insertManyIfNewUnit_list_mem_of_mem [EquivBEq α] [LawfulHashable α] +theorem getKey_insertManyIfNewUnit_list_of_mem [EquivBEq α] [LawfulHashable α] {l : List α} {k : α} (mem : k ∈ m) {h} : getKey (insertManyIfNewUnit m l) k h = getKey m k mem := DHashMap.Const.getKey_insertManyIfNewUnit_list_of_mem mem @@ -950,14 +950,14 @@ theorem getKey!_insertManyIfNewUnit_list_of_contains_eq_false_of_contains_eq_fal theorem getKey!_insertManyIfNewUnit_list_of_contains_eq_false_of_mem [EquivBEq α] [LawfulHashable α] [Inhabited α] {l : List α} {k k' : α} (k_beq : k == k') - (distinct : l.Pairwise (fun a b => (a == b) = false)) - (contains_eq_false : m.contains k = false) (mem : k ∈ l) : + (contains_eq_false : m.contains k = false) + (distinct : l.Pairwise (fun a b => (a == b) = false)) (mem : k ∈ l) : getKey! (insertManyIfNewUnit m l) k' = k := DHashMap.Const.getKey!_insertManyIfNewUnit_list_of_contains_eq_false_of_mem - k_beq distinct contains_eq_false mem + k_beq contains_eq_false distinct mem -theorem getKey!_insertManyIfNewUnit_list_mem_of_mem [EquivBEq α] [LawfulHashable α] - [Inhabited α] {l : List α} {k : α} (mem : k ∈ m): +theorem getKey!_insertManyIfNewUnit_list_of_mem [EquivBEq α] [LawfulHashable α] + [Inhabited α] {l : List α} {k : α} (mem : k ∈ m) : getKey! (insertManyIfNewUnit m l) k = getKey! m k := DHashMap.Const.getKey!_insertManyIfNewUnit_list_of_mem mem @@ -970,21 +970,21 @@ theorem getKeyD_insertManyIfNewUnit_list_of_contains_eq_false_of_contains_eq_fal theorem getKeyD_insertManyIfNewUnit_list_of_contains_eq_false_of_mem [EquivBEq α] [LawfulHashable α] {l : List α} {k k' fallback : α} (k_beq : k == k') - (distinct : l.Pairwise (fun a b => (a == b) = false)) - (contains_eq_false : m.contains k = false) (mem : k ∈ l ) : + (contains_eq_false : m.contains k = false) + (distinct : l.Pairwise (fun a b => (a == b) = false)) (mem : k ∈ l ) : getKeyD (insertManyIfNewUnit m l) k' fallback = k := DHashMap.Const.getKeyD_insertManyIfNewUnit_list_of_contains_eq_false_of_mem - k_beq distinct contains_eq_false mem + k_beq contains_eq_false distinct mem theorem getKeyD_insertManyIfNewUnit_list_of_mem [EquivBEq α] [LawfulHashable α] - {l : List α} {k fallback : α} (mem : k ∈ m): + {l : List α} {k fallback : α} (mem : k ∈ m) : getKeyD (insertManyIfNewUnit m l) k fallback = getKeyD m k fallback := DHashMap.Const.getKeyD_insertManyIfNewUnit_list_of_mem mem theorem size_insertManyIfNewUnit_list [EquivBEq α] [LawfulHashable α] {l : List α} (distinct : l.Pairwise (fun a b => (a == b) = false)) : - (∀ (a : α), m.contains a → l.contains a = false) → + (∀ (a : α), a ∈ m → l.contains a = false) → (insertManyIfNewUnit m l).size = m.size + l.length := DHashMap.Const.size_insertManyIfNewUnit_list distinct diff --git a/src/Std/Data/HashMap/RawLemmas.lean b/src/Std/Data/HashMap/RawLemmas.lean index 6c1ea5a45d47..c5eb10a57133 100644 --- a/src/Std/Data/HashMap/RawLemmas.lean +++ b/src/Std/Data/HashMap/RawLemmas.lean @@ -687,7 +687,7 @@ theorem length_keys [EquivBEq α] [LawfulHashable α] (h : m.WF) : DHashMap.Raw.length_keys h.out @[simp] -theorem isEmpty_keys [EquivBEq α] [LawfulHashable α] (h : m.WF): +theorem isEmpty_keys [EquivBEq α] [LawfulHashable α] (h : m.WF) : m.keys.isEmpty = m.isEmpty := DHashMap.Raw.isEmpty_keys h.out @@ -730,12 +730,12 @@ theorem contains_insertMany_list [EquivBEq α] [LawfulHashable α] (h : m.WF) @[simp] theorem mem_insertMany_list [EquivBEq α] [LawfulHashable α] (h : m.WF) {l : List (α × β)} {k : α} : - k ∈ (insertMany m l) ↔ k ∈ m ∨ (l.map Prod.fst).contains k := + k ∈ insertMany m l ↔ k ∈ m ∨ (l.map Prod.fst).contains k := DHashMap.Raw.Const.mem_insertMany_list h.out theorem mem_of_mem_insertMany_list [EquivBEq α] [LawfulHashable α] (h : m.WF) - {l : List ( α × β )} {k : α} : - k ∈ (insertMany m l) → (l.map Prod.fst).contains k = false → k ∈ m := + {l : List (α × β)} {k : α} : + k ∈ insertMany m l → (l.map Prod.fst).contains k = false → k ∈ m := DHashMap.Raw.Const.mem_of_mem_insertMany_list h.out theorem getKey?_insertMany_list_of_contains_eq_false [EquivBEq α] [LawfulHashable α] @@ -800,7 +800,7 @@ theorem getKeyD_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] theorem size_insertMany_list [EquivBEq α] [LawfulHashable α] (h : m.WF) {l : List (α × β)} (distinct : l.Pairwise (fun a b => (a.1 == b.1) = false)) : - (∀ (a : α), m.contains a → (l.map Prod.fst).contains a = false) → + (∀ (a : α), a ∈ m → (l.map Prod.fst).contains a = false) → (insertMany m l).size = m.size + l.length := DHashMap.Raw.Const.size_insertMany_list h.out distinct @@ -895,7 +895,7 @@ theorem mem_insertManyIfNewUnit_list [EquivBEq α] [LawfulHashable α] (h : m.WF theorem mem_of_mem_insertManyIfNewUnit_list [EquivBEq α] [LawfulHashable α] (h : m.WF) {l : List α} {k : α} (contains_eq_false : l.contains k = false) : - k ∈ (insertManyIfNewUnit m l) → k ∈ m := + k ∈ insertManyIfNewUnit m l → k ∈ m := DHashMap.Raw.Const.mem_of_mem_insertManyIfNewUnit_list h.out contains_eq_false theorem getKey?_insertManyIfNewUnit_list_of_contains_eq_false_of_contains_eq_false @@ -903,15 +903,15 @@ theorem getKey?_insertManyIfNewUnit_list_of_contains_eq_false_of_contains_eq_fal (contains_eq_false : m.contains k = false) (contains_eq_false': l.contains k = false) : getKey? (insertManyIfNewUnit m l) k = none := DHashMap.Raw.Const.getKey?_insertManyIfNewUnit_list_of_contains_eq_false_of_contains_eq_false - h.out contains_eq_false' contains_eq_false + h.out contains_eq_false contains_eq_false' theorem getKey?_insertManyIfNewUnit_list_of_contains_eq_false_of_mem [EquivBEq α] [LawfulHashable α] (h : m.WF) {l : List α} {k k' : α} (k_beq : k == k') - (distinct : l.Pairwise (fun a b => (a == b) = false)) - (contains_eq_false : m.contains k = false) (mem : k ∈ l) : + (contains_eq_false : m.contains k = false) + (distinct : l.Pairwise (fun a b => (a == b) = false)) (mem : k ∈ l) : getKey? (insertManyIfNewUnit m l) k' = some k := DHashMap.Raw.Const.getKey?_insertManyIfNewUnit_list_of_contains_eq_false_of_mem - h.out k_beq distinct mem contains_eq_false + h.out k_beq contains_eq_false distinct mem theorem getKey?_insertManyIfNewUnit_list_of_mem [EquivBEq α] [LawfulHashable α] (h : m.WF) {l : List α} {k : α} (mem : k ∈ m) : @@ -921,11 +921,11 @@ theorem getKey?_insertManyIfNewUnit_list_of_mem [EquivBEq α] [LawfulHashable α theorem getKey_insertManyIfNewUnit_list_of_contains_eq_false_of_mem [EquivBEq α] [LawfulHashable α] (h : m.WF) {l : List α} {k k' : α} (k_beq : k == k') - (distinct : l.Pairwise (fun a b => (a == b) = false)) - (contains_eq_false : m.contains k = false) (mem : k ∈ l) {h'} : + (contains_eq_false : m.contains k = false) + (distinct : l.Pairwise (fun a b => (a == b) = false)) (mem : k ∈ l) {h'} : getKey (insertManyIfNewUnit m l) k' h' = k := DHashMap.Raw.Const.getKey_insertManyIfNewUnit_list_of_contains_eq_false_of_mem - h.out k_beq distinct mem contains_eq_false + h.out k_beq contains_eq_false distinct mem theorem getKey_insertManyIfNewUnit_list_of_mem [EquivBEq α] [LawfulHashable α] (h : m.WF) {l : List α} {k : α} (mem: k ∈ m) {h₃} : @@ -937,15 +937,15 @@ theorem getKey!_insertManyIfNewUnit_list_of_contains_eq_false_of_contains_eq_fal (contains_eq_false : m.contains k = false) (contains_eq_false' : l.contains k = false) : getKey! (insertManyIfNewUnit m l) k = default := DHashMap.Raw.Const.getKey!_insertManyIfNewUnit_list_of_contains_eq_false_of_contains_eq_false - h.out contains_eq_false' contains_eq_false + h.out contains_eq_false contains_eq_false' theorem getKey!_insertManyIfNewUnit_list_of_contains_eq_false_of_mem [EquivBEq α] [LawfulHashable α] [Inhabited α] (h : m.WF) {l : List α} {k k' : α} (k_beq : k == k') - (distinct : l.Pairwise (fun a b => (a == b) = false)) - (contains_eq_false : m.contains k = false) (mem : k ∈ l) : + (contains_eq_false : m.contains k = false) + (distinct : l.Pairwise (fun a b => (a == b) = false)) (mem : k ∈ l) : getKey! (insertManyIfNewUnit m l) k' = k := DHashMap.Raw.Const.getKey!_insertManyIfNewUnit_list_of_contains_eq_false_of_mem - h.out k_beq distinct mem contains_eq_false + h.out k_beq contains_eq_false distinct mem theorem getKey!_insertManyIfNewUnit_list_of_mem [EquivBEq α] [LawfulHashable α] [Inhabited α] (h : m.WF) {l : List α} {k : α} (mem : k ∈ m) : @@ -957,25 +957,25 @@ theorem getKeyD_insertManyIfNewUnit_list_of_contains_eq_false_of_contains_eq_fal (contains_eq_false : m.contains k = false) (contains_eq_false' : l.contains k = false) : getKeyD (insertManyIfNewUnit m l) k fallback = fallback := DHashMap.Raw.Const.getKeyD_insertManyIfNewUnit_list_of_contains_eq_false_of_contains_eq_false - h.out contains_eq_false' contains_eq_false + h.out contains_eq_false contains_eq_false' theorem getKeyD_insertManyIfNewUnit_list_of_contains_eq_false_of_mem [EquivBEq α] [LawfulHashable α] (h : m.WF) {l : List α} {k k' fallback : α} (k_beq : k == k') - (distinct : l.Pairwise (fun a b => (a == b) = false)) - (contains_eq_false : m.contains k = false) (mem : k ∈ l) : + (contains_eq_false : m.contains k = false) + (distinct : l.Pairwise (fun a b => (a == b) = false)) (mem : k ∈ l) : getKeyD (insertManyIfNewUnit m l) k' fallback = k := DHashMap.Raw.Const.getKeyD_insertManyIfNewUnit_list_of_contains_eq_false_of_mem - h.out k_beq distinct mem contains_eq_false + h.out k_beq contains_eq_false distinct mem theorem getKeyD_insertManyIfNewUnit_list_of_mem [EquivBEq α] [LawfulHashable α] - (h : m.WF) {l : List α} {k fallback : α} (mem : k ∈ m): + (h : m.WF) {l : List α} {k fallback : α} (mem : k ∈ m) : getKeyD (insertManyIfNewUnit m l) k fallback = getKeyD m k fallback := DHashMap.Raw.Const.getKeyD_insertManyIfNewUnit_list_of_mem h.out mem theorem size_insertManyIfNewUnit_list [EquivBEq α] [LawfulHashable α] (h : m.WF) {l : List α} (distinct : l.Pairwise (fun a b => (a == b) = false)) : - (∀ (a : α), m.contains a → l.contains a = false) → + (∀ (a : α), a ∈ m → l.contains a = false) → (insertManyIfNewUnit m l).size = m.size + l.length := DHashMap.Raw.Const.size_insertManyIfNewUnit_list h.out distinct @@ -994,7 +994,7 @@ theorem isEmpty_insertManyIfNewUnit_list [EquivBEq α] [LawfulHashable α] (h : theorem getElem?_insertManyIfNewUnit_list [EquivBEq α] [LawfulHashable α] (h : m.WF) {l : List α} {k : α} : (insertManyIfNewUnit m l)[k]? = - if m.contains k ∨ l.contains k then some () else none := + if k ∈ m ∨ l.contains k then some () else none := DHashMap.Raw.Const.get?_insertManyIfNewUnit_list h.out @[simp] diff --git a/src/Std/Data/HashSet/Lemmas.lean b/src/Std/Data/HashSet/Lemmas.lean index f01c79f14279..ef26858f446d 100644 --- a/src/Std/Data/HashSet/Lemmas.lean +++ b/src/Std/Data/HashSet/Lemmas.lean @@ -361,7 +361,7 @@ theorem isEmpty_toList [EquivBEq α] [LawfulHashable α] : HashMap.isEmpty_keys @[simp] -theorem contains_toList [EquivBEq α] [LawfulHashable α] {k : α}: +theorem contains_toList [EquivBEq α] [LawfulHashable α] {k : α} : m.toList.contains k = m.contains k := HashMap.contains_keys @@ -402,7 +402,7 @@ theorem mem_insertMany_list [EquivBEq α] [LawfulHashable α] theorem mem_of_mem_insertMany_list [EquivBEq α] [LawfulHashable α] {l : List α} {k : α} (contains_eq_false : l.contains k = false) : - k ∈ (insertMany m l) → k ∈ m := + k ∈ insertMany m l → k ∈ m := HashMap.mem_of_mem_insertManyIfNewUnit_list contains_eq_false theorem get?_insertMany_list_of_contains_eq_false_of_contains_eq_false @@ -414,30 +414,30 @@ theorem get?_insertMany_list_of_contains_eq_false_of_contains_eq_false theorem get?_insertMany_list_of_contains_eq_false_of_mem [EquivBEq α] [LawfulHashable α] {l : List α} {k k' : α} (k_beq : k == k') - (distinct : l.Pairwise (fun a b => (a == b) = false)) - (contains_eq_false : m.contains k = false) (mem : k ∈ l) : + (contains_eq_false : m.contains k = false) + (distinct : l.Pairwise (fun a b => (a == b) = false)) (mem : k ∈ l) : get? (insertMany m l) k' = some k := HashMap.getKey?_insertManyIfNewUnit_list_of_contains_eq_false_of_mem - k_beq distinct contains_eq_false mem + k_beq contains_eq_false distinct mem theorem get?_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] - {l : List α} {k : α} (mem : k ∈ m): + {l : List α} {k : α} (mem : k ∈ m) : get? (insertMany m l) k = get? m k := HashMap.getKey?_insertManyIfNewUnit_list_of_mem mem theorem get_insertMany_list_of_contains_eq_false_of_mem [EquivBEq α] [LawfulHashable α] {l : List α} {k k' : α} (k_beq : k == k') - (distinct : l.Pairwise (fun a b => (a == b) = false)) - (contains_eq_false : m.contains k = false) (mem : k ∈ l) {h} : + (contains_eq_false : m.contains k = false) + (distinct : l.Pairwise (fun a b => (a == b) = false)) (mem : k ∈ l) {h} : get (insertMany m l) k' h = k := HashMap.getKey_insertManyIfNewUnit_list_of_contains_eq_false_of_mem - k_beq distinct contains_eq_false mem + k_beq contains_eq_false distinct mem -theorem get_insertMany_list_mem_of_mem [EquivBEq α] [LawfulHashable α] - {l : List α} {k : α} (mem : k ∈ m) {h}: +theorem get_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] + {l : List α} {k : α} (mem : k ∈ m) {h} : get (insertMany m l) k h = get m k mem := - HashMap.getKey_insertManyIfNewUnit_list_mem_of_mem mem + HashMap.getKey_insertManyIfNewUnit_list_of_mem mem theorem get!_insertMany_list_of_contains_eq_false_of_contains_eq_false [EquivBEq α] [LawfulHashable α] [Inhabited α] {l : List α} {k : α} @@ -448,16 +448,16 @@ theorem get!_insertMany_list_of_contains_eq_false_of_contains_eq_false theorem get!_insertMany_list_of_contains_eq_false_of_mem [EquivBEq α] [LawfulHashable α] [Inhabited α] {l : List α} {k k' : α} (k_beq : k == k') - (distinct : l.Pairwise (fun a b => (a == b) = false)) - (contains_eq_false : m.contains k = false) (mem : k ∈ l) : + (contains_eq_false : m.contains k = false) + (distinct : l.Pairwise (fun a b => (a == b) = false)) (mem : k ∈ l) : get! (insertMany m l) k' = k := HashMap.getKey!_insertManyIfNewUnit_list_of_contains_eq_false_of_mem - k_beq distinct contains_eq_false mem + k_beq contains_eq_false distinct mem -theorem get!_insertMany_list_mem_of_mem [EquivBEq α] [LawfulHashable α] +theorem get!_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] [Inhabited α] {l : List α} {k : α} (mem : k ∈ m) : get! (insertMany m l) k = get! m k := - HashMap.getKey!_insertManyIfNewUnit_list_mem_of_mem mem + HashMap.getKey!_insertManyIfNewUnit_list_of_mem mem theorem getD_insertMany_list_of_contains_eq_false_of_contains_eq_false [EquivBEq α] [LawfulHashable α] {l : List α} {k fallback : α} @@ -468,21 +468,21 @@ theorem getD_insertMany_list_of_contains_eq_false_of_contains_eq_false theorem getD_insertMany_list_of_contains_eq_false_of_mem [EquivBEq α] [LawfulHashable α] {l : List α} {k k' fallback : α} (k_beq : k == k') - (distinct : l.Pairwise (fun a b => (a == b) = false)) - (mem : k ∈ l ) (contains_eq_false : m.contains k = false) : + (contains_eq_false : m.contains k = false) + (distinct : l.Pairwise (fun a b => (a == b) = false)) (mem : k ∈ l) : getD (insertMany m l) k' fallback = k := HashMap.getKeyD_insertManyIfNewUnit_list_of_contains_eq_false_of_mem - k_beq distinct contains_eq_false mem + k_beq contains_eq_false distinct mem theorem getD_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] - {l : List α} {k fallback : α} (mem : k ∈ m): + {l : List α} {k fallback : α} (mem : k ∈ m) : getD (insertMany m l) k fallback = getD m k fallback := HashMap.getKeyD_insertManyIfNewUnit_list_of_mem mem theorem size_insertMany_list [EquivBEq α] [LawfulHashable α] {l : List α} (distinct : l.Pairwise (fun a b => (a == b) = false)) : - (∀ (a : α), m.contains a → l.contains a = false) → + (∀ (a : α), a ∈ m → l.contains a = false) → (insertMany m l).size = m.size + l.length := HashMap.size_insertManyIfNewUnit_list distinct @@ -584,7 +584,6 @@ theorem isEmpty_ofList [EquivBEq α] [LawfulHashable α] (ofList l).isEmpty = l.isEmpty := HashMap.isEmpty_unitOfList - end end Std.HashSet diff --git a/src/Std/Data/HashSet/RawLemmas.lean b/src/Std/Data/HashSet/RawLemmas.lean index 93178271dec8..37fd3d89e996 100644 --- a/src/Std/Data/HashSet/RawLemmas.lean +++ b/src/Std/Data/HashSet/RawLemmas.lean @@ -367,22 +367,22 @@ theorem containsThenInsert_snd (h : m.WF) {k : α} : (m.containsThenInsert k).2 ext (HashMap.Raw.containsThenInsertIfNew_snd h.out) @[simp] -theorem length_toList [EquivBEq α] [LawfulHashable α] (h : m.WF): +theorem length_toList [EquivBEq α] [LawfulHashable α] (h : m.WF) : m.toList.length = m.size := HashMap.Raw.length_keys h.1 @[simp] -theorem isEmpty_toList [EquivBEq α] [LawfulHashable α] (h : m.WF): +theorem isEmpty_toList [EquivBEq α] [LawfulHashable α] (h : m.WF) : m.toList.isEmpty = m.isEmpty := HashMap.Raw.isEmpty_keys h.1 @[simp] -theorem contains_toList [EquivBEq α] [LawfulHashable α] {k : α} (h : m.WF): +theorem contains_toList [EquivBEq α] [LawfulHashable α] {k : α} (h : m.WF) : m.toList.contains k = m.contains k := HashMap.Raw.contains_keys h.1 @[simp] -theorem mem_toList [LawfulBEq α] [LawfulHashable α] (h : m.WF) {k : α}: +theorem mem_toList [LawfulBEq α] [LawfulHashable α] (h : m.WF) {k : α} : k ∈ m.toList ↔ k ∈ m := HashMap.Raw.mem_keys h.1 @@ -418,7 +418,7 @@ theorem mem_insertMany_list [EquivBEq α] [LawfulHashable α] (h : m.WF) theorem mem_of_mem_insertMany_list [EquivBEq α] [LawfulHashable α] (h : m.WF) {l : List α} {k : α} (contains_eq_false : l.contains k = false) : - k ∈ (insertMany m l) → k ∈ m := + k ∈ insertMany m l → k ∈ m := HashMap.Raw.mem_of_mem_insertManyIfNewUnit_list h.1 contains_eq_false theorem get?_insertMany_list_of_contains_eq_false_of_contains_eq_false @@ -430,11 +430,11 @@ theorem get?_insertMany_list_of_contains_eq_false_of_contains_eq_false theorem get?_insertMany_list_of_contains_eq_false_of_mem [EquivBEq α] [LawfulHashable α] (h : m.WF) {l : List α} {k k' : α} (k_beq : k == k') - (distinct : l.Pairwise (fun a b => (a == b) = false)) - (contains_eq_false : m.contains k = false) (mem : k ∈ l) : + (contains_eq_false : m.contains k = false) + (distinct : l.Pairwise (fun a b => (a == b) = false)) (mem : k ∈ l) : get? (insertMany m l) k' = some k := HashMap.Raw.getKey?_insertManyIfNewUnit_list_of_contains_eq_false_of_mem - h.1 k_beq distinct contains_eq_false mem + h.1 k_beq contains_eq_false distinct mem theorem get?_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] (h : m.WF) {l : List α} {k : α} (mem : k ∈ m) : @@ -444,11 +444,11 @@ theorem get?_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] theorem get_insertMany_list_of_contains_eq_false_of_mem [EquivBEq α] [LawfulHashable α] (h : m.WF) {l : List α} {k k' : α} (k_beq : k == k') - (distinct : l.Pairwise (fun a b => (a == b) = false)) - (contains_eq_false : m.contains k = false) (mem : k ∈ l) {h'} : + (contains_eq_false : m.contains k = false) + (distinct : l.Pairwise (fun a b => (a == b) = false)) (mem : k ∈ l) {h'} : get (insertMany m l) k' h' = k := HashMap.Raw.getKey_insertManyIfNewUnit_list_of_contains_eq_false_of_mem - h.1 k_beq distinct contains_eq_false mem + h.1 k_beq contains_eq_false distinct mem theorem get_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] (h : m.WF) {l : List α} {k : α} (mem : k ∈ m) {h₃} : @@ -464,11 +464,11 @@ theorem get!_insertMany_list_of_contains_eq_false_of_contains_eq_false theorem get!_insertMany_list_of_contains_eq_false_of_mem [EquivBEq α] [LawfulHashable α] [Inhabited α] (h : m.WF) {l : List α} {k k' : α} (k_beq : k == k') - (distinct : l.Pairwise (fun a b => (a == b) = false)) - (contains_eq_false : contains m k = false) (mem : k ∈ l) : + (contains_eq_false : contains m k = false) + (distinct : l.Pairwise (fun a b => (a == b) = false)) (mem : k ∈ l) : get! (insertMany m l) k' = k := HashMap.Raw.getKey!_insertManyIfNewUnit_list_of_contains_eq_false_of_mem - h.1 k_beq distinct contains_eq_false mem + h.1 k_beq contains_eq_false distinct mem theorem get!_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] [Inhabited α] (h : m.WF) {l : List α} {k : α} (mem : k ∈ m) : @@ -484,11 +484,11 @@ theorem getD_insertMany_list_of_contains_eq_false_of_contains_eq_false theorem getD_insertMany_list_of_contains_eq_false_of_mem [EquivBEq α] [LawfulHashable α] (h : m.WF) {l : List α} {k k' fallback : α} (k_beq : k == k') - (distinct : l.Pairwise (fun a b => (a == b) = false)) - (contains_eq_false : m.contains k = false) (mem : k ∈ l) : + (contains_eq_false : m.contains k = false) + (distinct : l.Pairwise (fun a b => (a == b) = false)) (mem : k ∈ l) : getD (insertMany m l) k' fallback = k := HashMap.Raw.getKeyD_insertManyIfNewUnit_list_of_contains_eq_false_of_mem - h.1 k_beq distinct contains_eq_false mem + h.1 k_beq contains_eq_false distinct mem theorem getD_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] (h : m.WF) {l : List α} {k fallback : α} (mem : k ∈ m) : @@ -498,7 +498,7 @@ theorem getD_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] theorem size_insertMany_list [EquivBEq α] [LawfulHashable α] (h : m.WF) {l : List α} (distinct : l.Pairwise (fun a b => (a == b) = false)) : - (∀ (a : α), m.contains a → l.contains a = false) → + (∀ (a : α), a ∈ m → l.contains a = false) → (insertMany m l).size = m.size + l.length := HashMap.Raw.size_insertManyIfNewUnit_list h.1 distinct From b521a65f5f1583c241b7bedc39ba5e4b3768a04c Mon Sep 17 00:00:00 2001 From: jt0202 Date: Tue, 14 Jan 2025 11:10:34 +0100 Subject: [PATCH 80/83] Added docstring that went missing during merge --- src/Std/Data/DHashMap/Internal/List/Associative.lean | 1 + 1 file changed, 1 insertion(+) diff --git a/src/Std/Data/DHashMap/Internal/List/Associative.lean b/src/Std/Data/DHashMap/Internal/List/Associative.lean index 6e6b5d22aba1..5d78184dcd09 100644 --- a/src/Std/Data/DHashMap/Internal/List/Associative.lean +++ b/src/Std/Data/DHashMap/Internal/List/Associative.lean @@ -2695,6 +2695,7 @@ theorem getValue?_insertListIfNewUnit [BEq α] [PartialEquivBEq α] end +/-- Internal implementation detail of the hash map -/ def alterKey [BEq α] [LawfulBEq α] (k : α) (f : Option (β k) → Option (β k)) (l : List ((a : α) × β a)) : List ((a : α) × β a) := match f (getValueCast? k l) with From 97f9854c3cd052bbd7286f4875676ffaa6cd0a71 Mon Sep 17 00:00:00 2001 From: monsterkrampe Date: Tue, 14 Jan 2025 16:45:10 +0100 Subject: [PATCH 81/83] Add HashMap size_le_size_insertMany_list theorems --- .../DHashMap/Internal/List/Associative.lean | 21 +++++++++++++++++++ src/Std/Data/DHashMap/Internal/RawLemmas.lean | 15 +++++++++++++ src/Std/Data/DHashMap/Lemmas.lean | 15 +++++++++++++ src/Std/Data/DHashMap/RawLemmas.lean | 15 +++++++++++++ src/Std/Data/HashMap/Lemmas.lean | 10 +++++++++ src/Std/Data/HashMap/RawLemmas.lean | 10 +++++++++ src/Std/Data/HashSet/Lemmas.lean | 5 +++++ src/Std/Data/HashSet/RawLemmas.lean | 5 +++++ 8 files changed, 96 insertions(+) diff --git a/src/Std/Data/DHashMap/Internal/List/Associative.lean b/src/Std/Data/DHashMap/Internal/List/Associative.lean index 5d78184dcd09..e56d6864e7a9 100644 --- a/src/Std/Data/DHashMap/Internal/List/Associative.lean +++ b/src/Std/Data/DHashMap/Internal/List/Associative.lean @@ -2172,6 +2172,13 @@ theorem length_insertList [BEq α] [EquivBEq α] (insertList l toInsert).length = l.length + toInsert.length := by simpa using (perm_insertList distinct_l distinct_toInsert distinct_both).length_eq +theorem length_le_length_insertList [BEq α] + {l toInsert : List ((a : α) × β a)} : + l.length ≤ (insertList l toInsert).length := by + induction toInsert generalizing l with + | nil => apply Nat.le_refl + | cons hd tl ih => exact Nat.le_trans length_le_length_insertEntry ih + theorem length_insertList_le [BEq α] {l toInsert : List ((a : α) × β a)} : (insertList l toInsert).length ≤ l.length + toInsert.length := by @@ -2312,6 +2319,13 @@ theorem length_insertListConst [BEq α] [EquivBEq α] · simpa [List.pairwise_map] · simpa using distinct_both +theorem length_le_length_insertListConst [BEq α] + {l : List ((_ : α) × β)} {toInsert : List (α × β)} : + l.length ≤ (insertListConst l toInsert).length := by + induction toInsert generalizing l with + | nil => apply Nat.le_refl + | cons hd tl ih => exact Nat.le_trans length_le_length_insertEntry ih + theorem length_insertListConst_le [BEq α] {l : List ((_ : α) × β)} {toInsert : List (α × β)} : (insertListConst l toInsert).length ≤ l.length + toInsert.length := by @@ -2656,6 +2670,13 @@ theorem length_insertListIfNewUnit [BEq α] [EquivBEq α] rw [Bool.or_eq_false_iff] at distinct_both apply And.right distinct_both +theorem length_le_length_insertListIfNewUnit [BEq α] [EquivBEq α] + {l : List ((_ : α) × Unit)} {toInsert : List α}: + l.length ≤ (insertListIfNewUnit l toInsert).length := by + induction toInsert generalizing l with + | nil => apply Nat.le_refl + | cons hd tl ih => exact Nat.le_trans length_le_length_insertEntryIfNew ih + theorem length_insertListIfNewUnit_le [BEq α] [EquivBEq α] {l : List ((_ : α) × Unit)} {toInsert : List α}: (insertListIfNewUnit l toInsert).length ≤ l.length + toInsert.length := by diff --git a/src/Std/Data/DHashMap/Internal/RawLemmas.lean b/src/Std/Data/DHashMap/Internal/RawLemmas.lean index b99baf9bf643..e08aa970e851 100644 --- a/src/Std/Data/DHashMap/Internal/RawLemmas.lean +++ b/src/Std/Data/DHashMap/Internal/RawLemmas.lean @@ -990,6 +990,11 @@ theorem size_insertMany_list [EquivBEq α] [LawfulHashable α] (h : m.1.WF) (m.insertMany l).1.1.size = m.1.size + l.length := by simp_to_model using length_insertList +theorem size_le_size_insertMany_list [EquivBEq α] [LawfulHashable α] (h : m.1.WF) + {l : List ((a : α) × β a)} : + m.1.size ≤ (m.insertMany l).1.1.size := by + simp_to_model using length_le_length_insertList + theorem size_insertMany_list_le [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l : List ((a : α) × β a)} : (m.insertMany l).1.1.size ≤ m.1.size + l.length := by @@ -1099,6 +1104,11 @@ theorem size_insertMany_list [EquivBEq α] [LawfulHashable α] (h : m.1.WF) (insertMany m l).1.1.size = m.1.size + l.length := by simp_to_model using length_insertListConst +theorem size_le_size_insertMany_list [EquivBEq α] [LawfulHashable α] (h : m.1.WF) + {l : List (α × β)} : + m.1.size ≤ (insertMany m l).1.1.size := by + simp_to_model using length_le_length_insertListConst + theorem size_insertMany_list_le [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l : List (α × β)} : (insertMany m l).1.1.size ≤ m.1.size + l.length := by @@ -1263,6 +1273,11 @@ theorem size_insertManyIfNewUnit_list [EquivBEq α] [LawfulHashable α] (h : m.1 (insertManyIfNewUnit m l).1.1.size = m.1.size + l.length := by simp_to_model using length_insertListIfNewUnit +theorem size_le_size_insertManyIfNewUnit_list [EquivBEq α] [LawfulHashable α] (h : m.1.WF) + {l : List α} : + m.1.size ≤ (insertManyIfNewUnit m l).1.1.size := by + simp_to_model using length_le_length_insertListIfNewUnit + theorem size_insertManyIfNewUnit_list_le [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {l : List α} : (insertManyIfNewUnit m l).1.1.size ≤ m.1.size + l.length := by diff --git a/src/Std/Data/DHashMap/Lemmas.lean b/src/Std/Data/DHashMap/Lemmas.lean index 2294cfc9d250..3b8721e6c6c9 100644 --- a/src/Std/Data/DHashMap/Lemmas.lean +++ b/src/Std/Data/DHashMap/Lemmas.lean @@ -1115,6 +1115,11 @@ theorem size_insertMany_list [EquivBEq α] [LawfulHashable α] (m.insertMany l).size = m.size + l.length := Raw₀.size_insertMany_list ⟨m.1, _⟩ m.2 distinct +theorem size_le_size_insertMany_list [EquivBEq α] [LawfulHashable α] + {l : List ((a : α) × β a)} : + m.size ≤ (m.insertMany l).size := + Raw₀.size_le_size_insertMany_list ⟨m.1, _⟩ m.2 + theorem size_insertMany_list_le [EquivBEq α] [LawfulHashable α] {l : List ((a : α) × β a)} : (m.insertMany l).size ≤ m.size + l.length := @@ -1229,6 +1234,11 @@ theorem size_insertMany_list [EquivBEq α] [LawfulHashable α] (insertMany m l).size = m.size + l.length := Raw₀.Const.size_insertMany_list ⟨m.1, _⟩ m.2 distinct +theorem size_le_size_insertMany_list [EquivBEq α] [LawfulHashable α] + {l : List (α × β)} : + m.size ≤ (insertMany m l).size := + Raw₀.Const.size_le_size_insertMany_list ⟨m.1, _⟩ m.2 + theorem size_insertMany_list_le [EquivBEq α] [LawfulHashable α] {l : List (α × β)} : (insertMany m l).size ≤ m.size + l.length := @@ -1406,6 +1416,11 @@ theorem size_insertManyIfNewUnit_list [EquivBEq α] [LawfulHashable α] (insertManyIfNewUnit m l).size = m.size + l.length := Raw₀.Const.size_insertManyIfNewUnit_list ⟨m.1, _⟩ m.2 distinct +theorem size_le_size_insertManyIfNewUnit_list [EquivBEq α] [LawfulHashable α] + {l : List α} : + m.size ≤ (insertManyIfNewUnit m l).size := + Raw₀.Const.size_le_size_insertManyIfNewUnit_list ⟨m.1, _⟩ m.2 + theorem size_insertManyIfNewUnit_list_le [EquivBEq α] [LawfulHashable α] {l : List α} : (insertManyIfNewUnit m l).size ≤ m.size + l.length := diff --git a/src/Std/Data/DHashMap/RawLemmas.lean b/src/Std/Data/DHashMap/RawLemmas.lean index c533e3e9be3b..dd6989fde67c 100644 --- a/src/Std/Data/DHashMap/RawLemmas.lean +++ b/src/Std/Data/DHashMap/RawLemmas.lean @@ -1208,6 +1208,11 @@ theorem size_insertMany_list [EquivBEq α] [LawfulHashable α] (h : m.WF) simp only [mem_iff_contains] simp_to_raw using Raw₀.size_insertMany_list +theorem size_le_size_insertMany_list [EquivBEq α] [LawfulHashable α] (h : m.WF) + {l : List ((a : α) × β a)} : + m.size ≤ (m.insertMany l).size := by + simp_to_raw using Raw₀.size_le_size_insertMany_list ⟨m, _⟩ + theorem size_insertMany_list_le [EquivBEq α] [LawfulHashable α] (h : m.WF) {l : List ((a : α) × β a)} : (m.insertMany l).size ≤ m.size + l.length := by @@ -1327,6 +1332,11 @@ theorem size_insertMany_list [EquivBEq α] [LawfulHashable α] (h : m.WF) simp [mem_iff_contains] simp_to_raw using Raw₀.Const.size_insertMany_list +theorem size_le_size_insertMany_list [EquivBEq α] [LawfulHashable α] (h : m.WF) + {l : List (α × β)} : + m.size ≤ (insertMany m l).size := by + simp_to_raw using Raw₀.Const.size_le_size_insertMany_list ⟨m, _⟩ + theorem size_insertMany_list_le [EquivBEq α] [LawfulHashable α] (h : m.WF) {l : List (α × β)} : (insertMany m l).size ≤ m.size + l.length := by @@ -1500,6 +1510,11 @@ theorem size_insertManyIfNewUnit_list [EquivBEq α] [LawfulHashable α] (h : m.W simp only [mem_iff_contains] simp_to_raw using Raw₀.Const.size_insertManyIfNewUnit_list +theorem size_le_size_insertManyIfNewUnit_list [EquivBEq α] [LawfulHashable α] (h : m.WF) + {l : List α} : + m.size ≤ (insertManyIfNewUnit m l).size := by + simp_to_raw using Raw₀.Const.size_le_size_insertManyIfNewUnit_list ⟨m, _⟩ + theorem size_insertManyIfNewUnit_list_le [EquivBEq α] [LawfulHashable α] (h : m.WF) {l : List α} : (insertManyIfNewUnit m l).size ≤ m.size + l.length := by diff --git a/src/Std/Data/HashMap/Lemmas.lean b/src/Std/Data/HashMap/Lemmas.lean index be3e42211c4f..e43a7635d851 100644 --- a/src/Std/Data/HashMap/Lemmas.lean +++ b/src/Std/Data/HashMap/Lemmas.lean @@ -845,6 +845,11 @@ theorem size_insertMany_list [EquivBEq α] [LawfulHashable α] (insertMany m l).size = m.size + l.length := DHashMap.Const.size_insertMany_list distinct +theorem size_le_size_insertMany_list [EquivBEq α] [LawfulHashable α] + {l : List (α × β)} : + m.size ≤ (insertMany m l).size := + DHashMap.Const.size_le_size_insertMany_list + theorem size_insertMany_list_le [EquivBEq α] [LawfulHashable α] {l : List (α × β)} : (insertMany m l).size ≤ m.size + l.length := @@ -988,6 +993,11 @@ theorem size_insertManyIfNewUnit_list [EquivBEq α] [LawfulHashable α] (insertManyIfNewUnit m l).size = m.size + l.length := DHashMap.Const.size_insertManyIfNewUnit_list distinct +theorem size_le_size_insertManyIfNewUnit_list [EquivBEq α] [LawfulHashable α] + {l : List α} : + m.size ≤ (insertManyIfNewUnit m l).size := + DHashMap.Const.size_le_size_insertManyIfNewUnit_list + theorem size_insertManyIfNewUnit_list_le [EquivBEq α] [LawfulHashable α] {l : List α} : (insertManyIfNewUnit m l).size ≤ m.size + l.length := diff --git a/src/Std/Data/HashMap/RawLemmas.lean b/src/Std/Data/HashMap/RawLemmas.lean index c5eb10a57133..fffeca95a0c5 100644 --- a/src/Std/Data/HashMap/RawLemmas.lean +++ b/src/Std/Data/HashMap/RawLemmas.lean @@ -804,6 +804,11 @@ theorem size_insertMany_list [EquivBEq α] [LawfulHashable α] (insertMany m l).size = m.size + l.length := DHashMap.Raw.Const.size_insertMany_list h.out distinct +theorem size_le_size_insertMany_list [EquivBEq α] [LawfulHashable α] + (h : m.WF) {l : List (α × β)} : + m.size ≤ (insertMany m l).size := + DHashMap.Raw.Const.size_le_size_insertMany_list h.out + theorem size_insertMany_list_le [EquivBEq α] [LawfulHashable α] (h : m.WF) {l : List (α × β)} : (insertMany m l).size ≤ m.size + l.length := @@ -979,6 +984,11 @@ theorem size_insertManyIfNewUnit_list [EquivBEq α] [LawfulHashable α] (h : m.W (insertManyIfNewUnit m l).size = m.size + l.length := DHashMap.Raw.Const.size_insertManyIfNewUnit_list h.out distinct +theorem size_le_size_insertManyIfNewUnit_list [EquivBEq α] [LawfulHashable α] (h : m.WF) + {l : List α} : + m.size ≤ (insertManyIfNewUnit m l).size := + DHashMap.Raw.Const.size_le_size_insertManyIfNewUnit_list h.out + theorem size_insertManyIfNewUnit_list_le [EquivBEq α] [LawfulHashable α] (h : m.WF) {l : List α} : (insertManyIfNewUnit m l).size ≤ m.size + l.length := diff --git a/src/Std/Data/HashSet/Lemmas.lean b/src/Std/Data/HashSet/Lemmas.lean index ef26858f446d..c5905415a923 100644 --- a/src/Std/Data/HashSet/Lemmas.lean +++ b/src/Std/Data/HashSet/Lemmas.lean @@ -486,6 +486,11 @@ theorem size_insertMany_list [EquivBEq α] [LawfulHashable α] (insertMany m l).size = m.size + l.length := HashMap.size_insertManyIfNewUnit_list distinct +theorem size_le_size_insertMany_list [EquivBEq α] [LawfulHashable α] + {l : List α} : + m.size ≤ (insertMany m l).size := + HashMap.size_le_size_insertManyIfNewUnit_list + theorem size_insertMany_list_le [EquivBEq α] [LawfulHashable α] {l : List α} : (insertMany m l).size ≤ m.size + l.length := diff --git a/src/Std/Data/HashSet/RawLemmas.lean b/src/Std/Data/HashSet/RawLemmas.lean index 37fd3d89e996..6ddd2fd99907 100644 --- a/src/Std/Data/HashSet/RawLemmas.lean +++ b/src/Std/Data/HashSet/RawLemmas.lean @@ -502,6 +502,11 @@ theorem size_insertMany_list [EquivBEq α] [LawfulHashable α] (h : m.WF) (insertMany m l).size = m.size + l.length := HashMap.Raw.size_insertManyIfNewUnit_list h.1 distinct +theorem size_le_size_insertMany_list [EquivBEq α] [LawfulHashable α] (h : m.WF) + {l : List α} : + m.size ≤ (insertMany m l).size := + HashMap.Raw.size_le_size_insertManyIfNewUnit_list h.1 + theorem size_insertMany_list_le [EquivBEq α] [LawfulHashable α] (h : m.WF) {l : List α} : (insertMany m l).size ≤ m.size + l.length := From aea42fb9b893479a9f4206a2a9de1dea292174eb Mon Sep 17 00:00:00 2001 From: jt0202 Date: Tue, 14 Jan 2025 19:24:06 +0100 Subject: [PATCH 82/83] =?UTF-8?q?Replace=20'm.contains=20k=20=3D=20false'?= =?UTF-8?q?=20by=20'=C2=AC=20k=20=E2=88=88=20m'=20in=20statements=20for=20?= =?UTF-8?q?insertIfNewUnit?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit --- src/Std/Data/DHashMap/Lemmas.lean | 79 +++++++++++++++------------- src/Std/Data/DHashMap/RawLemmas.lean | 35 +++++++----- src/Std/Data/HashMap/Lemmas.lean | 56 ++++++++++---------- src/Std/Data/HashMap/RawLemmas.lean | 56 ++++++++++---------- src/Std/Data/HashSet/Lemmas.lean | 56 ++++++++++---------- src/Std/Data/HashSet/RawLemmas.lean | 56 ++++++++++---------- 6 files changed, 176 insertions(+), 162 deletions(-) diff --git a/src/Std/Data/DHashMap/Lemmas.lean b/src/Std/Data/DHashMap/Lemmas.lean index 3b8721e6c6c9..c9e49d1c13be 100644 --- a/src/Std/Data/DHashMap/Lemmas.lean +++ b/src/Std/Data/DHashMap/Lemmas.lean @@ -990,7 +990,7 @@ theorem mem_insertMany_list [EquivBEq α] [LawfulHashable α] simp [mem_iff_contains] theorem mem_of_mem_insertMany_list [EquivBEq α] [LawfulHashable α] - {l : List ((a : α) × β a)} {k : α} (mem : k ∈ (m.insertMany l)) + {l : List ((a : α) × β a)} {k : α} (mem : k ∈ m.insertMany l) (contains_eq_false : (l.map Sigma.fst).contains k = false) : k ∈ m := Raw₀.contains_of_contains_insertMany_list ⟨m.1, _⟩ m.2 mem contains_eq_false @@ -1335,74 +1335,81 @@ theorem mem_of_mem_insertManyIfNewUnit_list [EquivBEq α] [LawfulHashable α] k ∈ insertManyIfNewUnit m l → k ∈ m := Raw₀.Const.contains_of_contains_insertManyIfNewUnit_list ⟨m.1, _⟩ m.2 contains_eq_false -theorem getKey?_insertManyIfNewUnit_list_of_contains_eq_false_of_contains_eq_false +theorem getKey?_insertManyIfNewUnit_list_of_not_mem_of_contains_eq_false [EquivBEq α] [LawfulHashable α] {l : List α} {k : α} - (contains_eq_false : m.contains k = false) (contains_eq_false' : l.contains k = false) : - getKey? (insertManyIfNewUnit m l) k = none := - Raw₀.Const.getKey?_insertManyIfNewUnit_list_of_contains_eq_false_of_contains_eq_false - ⟨m.1, _⟩ m.2 contains_eq_false contains_eq_false' + (not_mem : ¬ k ∈ m) (contains_eq_false' : l.contains k = false) : + getKey? (insertManyIfNewUnit m l) k = none := by + simp only [mem_iff_contains, Bool.not_eq_true] at not_mem + exact Raw₀.Const.getKey?_insertManyIfNewUnit_list_of_contains_eq_false_of_contains_eq_false + ⟨m.1, _⟩ m.2 not_mem contains_eq_false' -theorem getKey?_insertManyIfNewUnit_list_of_contains_eq_false_of_mem [EquivBEq α] [LawfulHashable α] +theorem getKey?_insertManyIfNewUnit_list_of_not_mem_of_mem [EquivBEq α] [LawfulHashable α] {l : List α} {k k' : α} (k_beq : k == k') - (contains_eq_false : m.contains k = false) + (not_mem : ¬ k ∈ m) (distinct : l.Pairwise (fun a b => (a == b) = false)) (mem : k ∈ l) : - getKey? (insertManyIfNewUnit m l) k' = some k := - Raw₀.Const.getKey?_insertManyIfNewUnit_list_of_contains_eq_false_of_mem ⟨m.1, _⟩ - m.2 k_beq contains_eq_false distinct mem + getKey? (insertManyIfNewUnit m l) k' = some k := by + simp only [mem_iff_contains, Bool.not_eq_true] at not_mem + exact Raw₀.Const.getKey?_insertManyIfNewUnit_list_of_contains_eq_false_of_mem ⟨m.1, _⟩ + m.2 k_beq not_mem distinct mem theorem getKey?_insertManyIfNewUnit_list_of_mem [EquivBEq α] [LawfulHashable α] {l : List α} {k : α} (h' : k ∈ m) : getKey? (insertManyIfNewUnit m l) k = getKey? m k := Raw₀.Const.getKey?_insertManyIfNewUnit_list_of_contains ⟨m.1, _⟩ m.2 h' -theorem getKey_insertManyIfNewUnit_list_of_contains_eq_false_of_mem [EquivBEq α] [LawfulHashable α] +theorem getKey_insertManyIfNewUnit_list_of_not_mem_of_mem [EquivBEq α] [LawfulHashable α] {l : List α} {k k' : α} (k_beq : k == k') - (contains_eq_false : m.contains k = false) (distinct : l.Pairwise (fun a b => (a == b) = false)) + (not_mem : ¬ k ∈ m) (distinct : l.Pairwise (fun a b => (a == b) = false)) (mem : k ∈ l) {h} : - getKey (insertManyIfNewUnit m l) k' h = k := - Raw₀.Const.getKey_insertManyIfNewUnit_list_of_contains_eq_false_of_mem ⟨m.1, _⟩ m.2 k_beq - contains_eq_false distinct mem + getKey (insertManyIfNewUnit m l) k' h = k := by + simp only [mem_iff_contains, Bool.not_eq_true] at not_mem + exact Raw₀.Const.getKey_insertManyIfNewUnit_list_of_contains_eq_false_of_mem ⟨m.1, _⟩ m.2 k_beq + not_mem distinct mem theorem getKey_insertManyIfNewUnit_list_of_mem [EquivBEq α] [LawfulHashable α] {l : List α} {k : α} (mem : k ∈ m) {h} : getKey (insertManyIfNewUnit m l) k h = getKey m k mem := Raw₀.Const.getKey_insertManyIfNewUnit_list_of_contains ⟨m.1, _⟩ m.2 _ -theorem getKey!_insertManyIfNewUnit_list_of_contains_eq_false_of_contains_eq_false +theorem getKey!_insertManyIfNewUnit_list_of_not_mem_of_contains_eq_false [EquivBEq α] [LawfulHashable α] [Inhabited α] {l : List α} {k : α} - (contains_eq_false : m.contains k = false) (contains_eq_false' : l.contains k = false) : - getKey! (insertManyIfNewUnit m l) k = default := - Raw₀.Const.getKey!_insertManyIfNewUnit_list_of_contains_eq_false_of_contains_eq_false - ⟨m.1, _⟩ m.2 contains_eq_false contains_eq_false' + (not_mem : ¬ k ∈ m) (contains_eq_false : l.contains k = false) : + getKey! (insertManyIfNewUnit m l) k = default := by + simp only [mem_iff_contains, Bool.not_eq_true] at not_mem + exact Raw₀.Const.getKey!_insertManyIfNewUnit_list_of_contains_eq_false_of_contains_eq_false + ⟨m.1, _⟩ m.2 not_mem contains_eq_false -theorem getKey!_insertManyIfNewUnit_list_of_contains_eq_false_of_mem [EquivBEq α] [LawfulHashable α] +theorem getKey!_insertManyIfNewUnit_list_of_not_mem_of_mem [EquivBEq α] [LawfulHashable α] [Inhabited α] {l : List α} {k k' : α} (k_beq : k == k') - (contains_eq_false: m.contains k = false) (distinct : l.Pairwise (fun a b => (a == b) = false)) + (not_mem : ¬ k ∈ m) (distinct : l.Pairwise (fun a b => (a == b) = false)) (mem : k ∈ l) : - getKey! (insertManyIfNewUnit m l) k' = k := - Raw₀.Const.getKey!_insertManyIfNewUnit_list_of_contains_eq_false_of_mem ⟨m.1, _⟩ m.2 k_beq - contains_eq_false distinct mem + getKey! (insertManyIfNewUnit m l) k' = k := by + simp only [mem_iff_contains, Bool.not_eq_true] at not_mem + exact Raw₀.Const.getKey!_insertManyIfNewUnit_list_of_contains_eq_false_of_mem ⟨m.1, _⟩ m.2 k_beq + not_mem distinct mem theorem getKey!_insertManyIfNewUnit_list_of_mem [EquivBEq α] [LawfulHashable α] [Inhabited α] {l : List α} {k : α} (mem : k ∈ m) : getKey! (insertManyIfNewUnit m l) k = getKey! m k := Raw₀.Const.getKey!_insertManyIfNewUnit_list_of_contains ⟨m.1, _⟩ m.2 mem -theorem getKeyD_insertManyIfNewUnit_list_of_contains_eq_false_of_contains_eq_false +theorem getKeyD_insertManyIfNewUnit_list_of_not_mem_of_contains_eq_false [EquivBEq α] [LawfulHashable α] {l : List α} {k fallback : α} - (contains_eq_false : m.contains k = false) (contains_eq_false' : l.contains k = false) : - getKeyD (insertManyIfNewUnit m l) k fallback = fallback := - Raw₀.Const.getKeyD_insertManyIfNewUnit_list_of_contains_eq_false_of_contains_eq_false - ⟨m.1, _⟩ m.2 contains_eq_false contains_eq_false' + (not_mem : ¬ k ∈ m) (contains_eq_false : l.contains k = false) : + getKeyD (insertManyIfNewUnit m l) k fallback = fallback := by + simp only [mem_iff_contains, Bool.not_eq_true] at not_mem + exact Raw₀.Const.getKeyD_insertManyIfNewUnit_list_of_contains_eq_false_of_contains_eq_false + ⟨m.1, _⟩ m.2 not_mem contains_eq_false -theorem getKeyD_insertManyIfNewUnit_list_of_contains_eq_false_of_mem [EquivBEq α] [LawfulHashable α] +theorem getKeyD_insertManyIfNewUnit_list_of_not_mem_of_mem [EquivBEq α] [LawfulHashable α] {l : List α} {k k' fallback : α} (k_beq : k == k') - (contains_eq_false : m.contains k = false) + (not_mem : ¬ k ∈ m) (distinct : l.Pairwise (fun a b => (a == b) = false)) (mem : k ∈ l) : - getKeyD (insertManyIfNewUnit m l) k' fallback = k := - Raw₀.Const.getKeyD_insertManyIfNewUnit_list_of_contains_eq_false_of_mem ⟨m.1, _⟩ m.2 k_beq - contains_eq_false distinct mem + getKeyD (insertManyIfNewUnit m l) k' fallback = k := by + simp only [mem_iff_contains, Bool.not_eq_true] at not_mem + exact Raw₀.Const.getKeyD_insertManyIfNewUnit_list_of_contains_eq_false_of_mem ⟨m.1, _⟩ m.2 k_beq + not_mem distinct mem theorem getKeyD_insertManyIfNewUnit_list_of_mem [EquivBEq α] [LawfulHashable α] {l : List α} {k fallback : α} (mem : k ∈ m) : diff --git a/src/Std/Data/DHashMap/RawLemmas.lean b/src/Std/Data/DHashMap/RawLemmas.lean index dd6989fde67c..4aa6d3e5774a 100644 --- a/src/Std/Data/DHashMap/RawLemmas.lean +++ b/src/Std/Data/DHashMap/RawLemmas.lean @@ -1435,16 +1435,18 @@ theorem mem_of_mem_insertManyIfNewUnit_list [EquivBEq α] [LawfulHashable α] (h simp only [mem_iff_contains] simp_to_raw using Raw₀.Const.contains_of_contains_insertManyIfNewUnit_list -theorem getKey?_insertManyIfNewUnit_list_of_contains_eq_false_of_contains_eq_false +theorem getKey?_insertManyIfNewUnit_list_of_not_mem_of_contains_eq_false [EquivBEq α] [LawfulHashable α] (h : m.WF) {l : List α} {k : α} : - m.contains k = false → l.contains k = false → + ¬ k ∈ m → l.contains k = false → getKey? (insertManyIfNewUnit m l) k = none := by + simp only [mem_iff_contains, Bool.not_eq_true] simp_to_raw using Raw₀.Const.getKey?_insertManyIfNewUnit_list_of_contains_eq_false_of_contains_eq_false -theorem getKey?_insertManyIfNewUnit_list_of_contains_eq_false_of_mem [EquivBEq α] [LawfulHashable α] +theorem getKey?_insertManyIfNewUnit_list_of_not_mem_of_mem [EquivBEq α] [LawfulHashable α] (h : m.WF) {l : List α} {k k' : α} (k_beq : k == k') : - m.contains k = false → l.Pairwise (fun a b => (a == b) = false) → k ∈ l → + ¬ k ∈ m → l.Pairwise (fun a b => (a == b) = false) → k ∈ l → getKey? (insertManyIfNewUnit m l) k' = some k := by + simp only [mem_iff_contains, Bool.not_eq_true] simp_to_raw using Raw₀.Const.getKey?_insertManyIfNewUnit_list_of_contains_eq_false_of_mem theorem getKey?_insertManyIfNewUnit_list_of_mem [EquivBEq α] [LawfulHashable α] @@ -1458,23 +1460,26 @@ theorem getKey_insertManyIfNewUnit_list_of_mem getKey (insertManyIfNewUnit m l) k h' = getKey m k mem := by simp_to_raw using Raw₀.Const.getKey_insertManyIfNewUnit_list_of_contains -theorem getKey_insertManyIfNewUnit_list_of_contains_eq_false_of_mem [EquivBEq α] [LawfulHashable α] +theorem getKey_insertManyIfNewUnit_list_of_not_mem_of_mem [EquivBEq α] [LawfulHashable α] (h : m.WF) {l : List α} {k k' : α} (k_beq : k == k') {h'} : - m.contains k = false → l.Pairwise (fun a b => (a == b) = false) → k ∈ l → + ¬ k ∈ m → l.Pairwise (fun a b => (a == b) = false) → k ∈ l → getKey (insertManyIfNewUnit m l) k' h' = k := by + simp only [mem_iff_contains, Bool.not_eq_true] simp_to_raw using Raw₀.Const.getKey_insertManyIfNewUnit_list_of_contains_eq_false_of_mem -theorem getKey!_insertManyIfNewUnit_list_of_contains_eq_false_of_contains_eq_false +theorem getKey!_insertManyIfNewUnit_list_of_not_mem_of_contains_eq_false [EquivBEq α] [LawfulHashable α] [Inhabited α] (h : m.WF) {l : List α} {k : α} : - m.contains k = false → l.contains k = false → getKey! (insertManyIfNewUnit m l) k = default := by + ¬ k ∈ m → l.contains k = false → getKey! (insertManyIfNewUnit m l) k = default := by + simp only [mem_iff_contains, Bool.not_eq_true] simp_to_raw using Raw₀.Const.getKey!_insertManyIfNewUnit_list_of_contains_eq_false_of_contains_eq_false -theorem getKey!_insertManyIfNewUnit_list_of_contains_eq_false_of_mem [EquivBEq α] [LawfulHashable α] +theorem getKey!_insertManyIfNewUnit_list_of_not_mem_of_mem [EquivBEq α] [LawfulHashable α] [Inhabited α] (h : m.WF) {l : List α} {k k' : α} (k_beq : k == k') : - contains m k = false → l.Pairwise (fun a b => (a == b) = false) → k ∈ l → + ¬ k ∈ m → l.Pairwise (fun a b => (a == b) = false) → k ∈ l → getKey! (insertManyIfNewUnit m l) k' = k := by + simp only [mem_iff_contains, Bool.not_eq_true] simp_to_raw using Raw₀.Const.getKey!_insertManyIfNewUnit_list_of_contains_eq_false_of_mem theorem getKey!_insertManyIfNewUnit_list_of_mem [EquivBEq α] [LawfulHashable α] @@ -1483,17 +1488,19 @@ theorem getKey!_insertManyIfNewUnit_list_of_mem [EquivBEq α] [LawfulHashable α simp only [mem_iff_contains] simp_to_raw using Raw₀.Const.getKey!_insertManyIfNewUnit_list_of_contains -theorem getKeyD_insertManyIfNewUnit_list_of_contains_eq_false_of_contains_eq_false +theorem getKeyD_insertManyIfNewUnit_list_of_not_mem_of_contains_eq_false [EquivBEq α] [LawfulHashable α] (h : m.WF) {l : List α} {k fallback : α} : - m.contains k = false → l.contains k = false → + ¬ k ∈ m → l.contains k = false → getKeyD (insertManyIfNewUnit m l) k fallback = fallback := by + simp only [mem_iff_contains, Bool.not_eq_true] simp_to_raw using Raw₀.Const.getKeyD_insertManyIfNewUnit_list_of_contains_eq_false_of_contains_eq_false -theorem getKeyD_insertManyIfNewUnit_list_of_contains_eq_false_of_mem [EquivBEq α] [LawfulHashable α] +theorem getKeyD_insertManyIfNewUnit_list_of_not_mem_of_mem [EquivBEq α] [LawfulHashable α] (h : m.WF) {l : List α} {k k' fallback : α} (k_beq : k == k') : - m.contains k = false → l.Pairwise (fun a b => (a == b) = false) → k ∈ l → + ¬ k ∈ m → l.Pairwise (fun a b => (a == b) = false) → k ∈ l → getKeyD (insertManyIfNewUnit m l) k' fallback = k := by + simp only [mem_iff_contains, Bool.not_eq_true] simp_to_raw using Raw₀.Const.getKeyD_insertManyIfNewUnit_list_of_contains_eq_false_of_mem theorem getKeyD_insertManyIfNewUnit_list_of_mem [EquivBEq α] [LawfulHashable α] diff --git a/src/Std/Data/HashMap/Lemmas.lean b/src/Std/Data/HashMap/Lemmas.lean index e43a7635d851..19e2d073bc57 100644 --- a/src/Std/Data/HashMap/Lemmas.lean +++ b/src/Std/Data/HashMap/Lemmas.lean @@ -915,71 +915,71 @@ theorem getD_insertManyIfNewUnit_list getD (insertManyIfNewUnit m l) k fallback = () := by simp -theorem getKey?_insertManyIfNewUnit_list_of_contains_eq_false_of_contains_eq_false +theorem getKey?_insertManyIfNewUnit_list_of_not_mem_of_contains_eq_false [EquivBEq α] [LawfulHashable α] {l : List α} {k : α} - (contains_eq_false : m.contains k = false) (contains_eq_false' : l.contains k = false) : + (not_mem : ¬ k ∈ m) (contains_eq_false : l.contains k = false) : getKey? (insertManyIfNewUnit m l) k = none := - DHashMap.Const.getKey?_insertManyIfNewUnit_list_of_contains_eq_false_of_contains_eq_false - contains_eq_false contains_eq_false' + DHashMap.Const.getKey?_insertManyIfNewUnit_list_of_not_mem_of_contains_eq_false + not_mem contains_eq_false -theorem getKey?_insertManyIfNewUnit_list_of_contains_eq_false_of_mem [EquivBEq α] [LawfulHashable α] - {l : List α} {k k' : α} (k_beq : k == k') (contains_eq_false : m.contains k = false) +theorem getKey?_insertManyIfNewUnit_list_of_not_mem_of_mem [EquivBEq α] [LawfulHashable α] + {l : List α} {k k' : α} (k_beq : k == k') (not_mem : ¬ k ∈ m) (distinct : l.Pairwise (fun a b => (a == b) = false)) (mem : k ∈ l) : getKey? (insertManyIfNewUnit m l) k' = some k := - DHashMap.Const.getKey?_insertManyIfNewUnit_list_of_contains_eq_false_of_mem - k_beq contains_eq_false distinct mem + DHashMap.Const.getKey?_insertManyIfNewUnit_list_of_not_mem_of_mem + k_beq not_mem distinct mem theorem getKey?_insertManyIfNewUnit_list_of_mem [EquivBEq α] [LawfulHashable α] {l : List α} {k : α} (mem : k ∈ m) : getKey? (insertManyIfNewUnit m l) k = getKey? m k := DHashMap.Const.getKey?_insertManyIfNewUnit_list_of_mem mem -theorem getKey_insertManyIfNewUnit_list_of_contains_eq_false_of_mem [EquivBEq α] [LawfulHashable α] - {l : List α} {k k' : α} (k_beq : k == k') (contains_eq_false : m.contains k = false) +theorem getKey_insertManyIfNewUnit_list_of_not_mem_of_mem [EquivBEq α] [LawfulHashable α] + {l : List α} {k k' : α} (k_beq : k == k') (not_mem : ¬ k ∈ m) (distinct : l.Pairwise (fun a b => (a == b) = false)) (mem : k ∈ l) {h} : getKey (insertManyIfNewUnit m l) k' h = k := - DHashMap.Const.getKey_insertManyIfNewUnit_list_of_contains_eq_false_of_mem - k_beq contains_eq_false distinct mem + DHashMap.Const.getKey_insertManyIfNewUnit_list_of_not_mem_of_mem + k_beq not_mem distinct mem theorem getKey_insertManyIfNewUnit_list_of_mem [EquivBEq α] [LawfulHashable α] {l : List α} {k : α} (mem : k ∈ m) {h} : getKey (insertManyIfNewUnit m l) k h = getKey m k mem := DHashMap.Const.getKey_insertManyIfNewUnit_list_of_mem mem -theorem getKey!_insertManyIfNewUnit_list_of_contains_eq_false_of_contains_eq_false +theorem getKey!_insertManyIfNewUnit_list_of_not_mem_of_contains_eq_false [EquivBEq α] [LawfulHashable α] [Inhabited α] {l : List α} {k : α} - (contains_eq_false : m.contains k = false) (contains_eq_false' : l.contains k = false) : + (not_mem : ¬ k ∈ m) (contains_eq_false : l.contains k = false) : getKey! (insertManyIfNewUnit m l) k = default := - DHashMap.Const.getKey!_insertManyIfNewUnit_list_of_contains_eq_false_of_contains_eq_false - contains_eq_false contains_eq_false' + DHashMap.Const.getKey!_insertManyIfNewUnit_list_of_not_mem_of_contains_eq_false + not_mem contains_eq_false -theorem getKey!_insertManyIfNewUnit_list_of_contains_eq_false_of_mem [EquivBEq α] [LawfulHashable α] +theorem getKey!_insertManyIfNewUnit_list_of_not_mem_of_mem [EquivBEq α] [LawfulHashable α] [Inhabited α] {l : List α} {k k' : α} (k_beq : k == k') - (contains_eq_false : m.contains k = false) + (not_mem : ¬ k ∈ m) (distinct : l.Pairwise (fun a b => (a == b) = false)) (mem : k ∈ l) : getKey! (insertManyIfNewUnit m l) k' = k := - DHashMap.Const.getKey!_insertManyIfNewUnit_list_of_contains_eq_false_of_mem - k_beq contains_eq_false distinct mem + DHashMap.Const.getKey!_insertManyIfNewUnit_list_of_not_mem_of_mem + k_beq not_mem distinct mem theorem getKey!_insertManyIfNewUnit_list_of_mem [EquivBEq α] [LawfulHashable α] [Inhabited α] {l : List α} {k : α} (mem : k ∈ m) : getKey! (insertManyIfNewUnit m l) k = getKey! m k := DHashMap.Const.getKey!_insertManyIfNewUnit_list_of_mem mem -theorem getKeyD_insertManyIfNewUnit_list_of_contains_eq_false_of_contains_eq_false +theorem getKeyD_insertManyIfNewUnit_list_of_not_mem_of_contains_eq_false [EquivBEq α] [LawfulHashable α] {l : List α} {k fallback : α} - (contains_eq_false : m.contains k = false) (contains_eq_false' : l.contains k = false) : + (not_mem : ¬ k ∈ m) (contains_eq_false : l.contains k = false) : getKeyD (insertManyIfNewUnit m l) k fallback = fallback := - DHashMap.Const.getKeyD_insertManyIfNewUnit_list_of_contains_eq_false_of_contains_eq_false - contains_eq_false contains_eq_false' + DHashMap.Const.getKeyD_insertManyIfNewUnit_list_of_not_mem_of_contains_eq_false + not_mem contains_eq_false -theorem getKeyD_insertManyIfNewUnit_list_of_contains_eq_false_of_mem [EquivBEq α] [LawfulHashable α] +theorem getKeyD_insertManyIfNewUnit_list_of_not_mem_of_mem [EquivBEq α] [LawfulHashable α] {l : List α} {k k' fallback : α} (k_beq : k == k') - (contains_eq_false : m.contains k = false) + (not_mem : ¬ k ∈ m) (distinct : l.Pairwise (fun a b => (a == b) = false)) (mem : k ∈ l ) : getKeyD (insertManyIfNewUnit m l) k' fallback = k := - DHashMap.Const.getKeyD_insertManyIfNewUnit_list_of_contains_eq_false_of_mem - k_beq contains_eq_false distinct mem + DHashMap.Const.getKeyD_insertManyIfNewUnit_list_of_not_mem_of_mem + k_beq not_mem distinct mem theorem getKeyD_insertManyIfNewUnit_list_of_mem [EquivBEq α] [LawfulHashable α] {l : List α} {k fallback : α} (mem : k ∈ m) : diff --git a/src/Std/Data/HashMap/RawLemmas.lean b/src/Std/Data/HashMap/RawLemmas.lean index fffeca95a0c5..bdcc2b0b857c 100644 --- a/src/Std/Data/HashMap/RawLemmas.lean +++ b/src/Std/Data/HashMap/RawLemmas.lean @@ -903,74 +903,74 @@ theorem mem_of_mem_insertManyIfNewUnit_list [EquivBEq α] [LawfulHashable α] (h k ∈ insertManyIfNewUnit m l → k ∈ m := DHashMap.Raw.Const.mem_of_mem_insertManyIfNewUnit_list h.out contains_eq_false -theorem getKey?_insertManyIfNewUnit_list_of_contains_eq_false_of_contains_eq_false +theorem getKey?_insertManyIfNewUnit_list_of_not_mem_of_contains_eq_false [EquivBEq α] [LawfulHashable α] (h : m.WF) {l : List α} {k : α} - (contains_eq_false : m.contains k = false) (contains_eq_false': l.contains k = false) : + (not_mem : ¬ k ∈ m) (contains_eq_false: l.contains k = false) : getKey? (insertManyIfNewUnit m l) k = none := - DHashMap.Raw.Const.getKey?_insertManyIfNewUnit_list_of_contains_eq_false_of_contains_eq_false - h.out contains_eq_false contains_eq_false' + DHashMap.Raw.Const.getKey?_insertManyIfNewUnit_list_of_not_mem_of_contains_eq_false + h.out not_mem contains_eq_false -theorem getKey?_insertManyIfNewUnit_list_of_contains_eq_false_of_mem [EquivBEq α] [LawfulHashable α] +theorem getKey?_insertManyIfNewUnit_list_of_not_mem_of_mem [EquivBEq α] [LawfulHashable α] (h : m.WF) {l : List α} {k k' : α} (k_beq : k == k') - (contains_eq_false : m.contains k = false) + (not_mem : ¬ k ∈ m) (distinct : l.Pairwise (fun a b => (a == b) = false)) (mem : k ∈ l) : getKey? (insertManyIfNewUnit m l) k' = some k := - DHashMap.Raw.Const.getKey?_insertManyIfNewUnit_list_of_contains_eq_false_of_mem - h.out k_beq contains_eq_false distinct mem + DHashMap.Raw.Const.getKey?_insertManyIfNewUnit_list_of_not_mem_of_mem + h.out k_beq not_mem distinct mem theorem getKey?_insertManyIfNewUnit_list_of_mem [EquivBEq α] [LawfulHashable α] (h : m.WF) {l : List α} {k : α} (mem : k ∈ m) : getKey? (insertManyIfNewUnit m l) k = getKey? m k := DHashMap.Raw.Const.getKey?_insertManyIfNewUnit_list_of_mem h.out mem -theorem getKey_insertManyIfNewUnit_list_of_contains_eq_false_of_mem [EquivBEq α] [LawfulHashable α] +theorem getKey_insertManyIfNewUnit_list_of_not_mem_of_mem [EquivBEq α] [LawfulHashable α] (h : m.WF) {l : List α} {k k' : α} (k_beq : k == k') - (contains_eq_false : m.contains k = false) + (not_mem : ¬ k ∈ m) (distinct : l.Pairwise (fun a b => (a == b) = false)) (mem : k ∈ l) {h'} : getKey (insertManyIfNewUnit m l) k' h' = k := - DHashMap.Raw.Const.getKey_insertManyIfNewUnit_list_of_contains_eq_false_of_mem - h.out k_beq contains_eq_false distinct mem + DHashMap.Raw.Const.getKey_insertManyIfNewUnit_list_of_not_mem_of_mem + h.out k_beq not_mem distinct mem theorem getKey_insertManyIfNewUnit_list_of_mem [EquivBEq α] [LawfulHashable α] (h : m.WF) {l : List α} {k : α} (mem: k ∈ m) {h₃} : getKey (insertManyIfNewUnit m l) k h₃ = getKey m k mem := DHashMap.Raw.Const.getKey_insertManyIfNewUnit_list_of_mem h.out mem -theorem getKey!_insertManyIfNewUnit_list_of_contains_eq_false_of_contains_eq_false +theorem getKey!_insertManyIfNewUnit_list_of_not_mem_of_contains_eq_false [EquivBEq α] [LawfulHashable α] [Inhabited α] (h : m.WF) {l : List α} {k : α} - (contains_eq_false : m.contains k = false) (contains_eq_false' : l.contains k = false) : + (not_mem : ¬ k ∈ m) (contains_eq_false : l.contains k = false) : getKey! (insertManyIfNewUnit m l) k = default := - DHashMap.Raw.Const.getKey!_insertManyIfNewUnit_list_of_contains_eq_false_of_contains_eq_false - h.out contains_eq_false contains_eq_false' + DHashMap.Raw.Const.getKey!_insertManyIfNewUnit_list_of_not_mem_of_contains_eq_false + h.out not_mem contains_eq_false -theorem getKey!_insertManyIfNewUnit_list_of_contains_eq_false_of_mem [EquivBEq α] [LawfulHashable α] +theorem getKey!_insertManyIfNewUnit_list_of_not_mem_of_mem [EquivBEq α] [LawfulHashable α] [Inhabited α] (h : m.WF) {l : List α} {k k' : α} (k_beq : k == k') - (contains_eq_false : m.contains k = false) + (not_mem : ¬ k ∈ m) (distinct : l.Pairwise (fun a b => (a == b) = false)) (mem : k ∈ l) : getKey! (insertManyIfNewUnit m l) k' = k := - DHashMap.Raw.Const.getKey!_insertManyIfNewUnit_list_of_contains_eq_false_of_mem - h.out k_beq contains_eq_false distinct mem + DHashMap.Raw.Const.getKey!_insertManyIfNewUnit_list_of_not_mem_of_mem + h.out k_beq not_mem distinct mem theorem getKey!_insertManyIfNewUnit_list_of_mem [EquivBEq α] [LawfulHashable α] [Inhabited α] (h : m.WF) {l : List α} {k : α} (mem : k ∈ m) : getKey! (insertManyIfNewUnit m l) k = getKey! m k := DHashMap.Raw.Const.getKey!_insertManyIfNewUnit_list_of_mem h.out mem -theorem getKeyD_insertManyIfNewUnit_list_of_contains_eq_false_of_contains_eq_false +theorem getKeyD_insertManyIfNewUnit_list_of_not_mem_of_contains_eq_false [EquivBEq α] [LawfulHashable α] (h : m.WF) {l : List α} {k fallback : α} - (contains_eq_false : m.contains k = false) (contains_eq_false' : l.contains k = false) : + (not_mem : ¬ k ∈ m) (contains_eq_false : l.contains k = false) : getKeyD (insertManyIfNewUnit m l) k fallback = fallback := - DHashMap.Raw.Const.getKeyD_insertManyIfNewUnit_list_of_contains_eq_false_of_contains_eq_false - h.out contains_eq_false contains_eq_false' + DHashMap.Raw.Const.getKeyD_insertManyIfNewUnit_list_of_not_mem_of_contains_eq_false + h.out not_mem contains_eq_false -theorem getKeyD_insertManyIfNewUnit_list_of_contains_eq_false_of_mem [EquivBEq α] [LawfulHashable α] +theorem getKeyD_insertManyIfNewUnit_list_of_not_mem_of_mem [EquivBEq α] [LawfulHashable α] (h : m.WF) {l : List α} {k k' fallback : α} (k_beq : k == k') - (contains_eq_false : m.contains k = false) + (not_mem : ¬ k ∈ m) (distinct : l.Pairwise (fun a b => (a == b) = false)) (mem : k ∈ l) : getKeyD (insertManyIfNewUnit m l) k' fallback = k := - DHashMap.Raw.Const.getKeyD_insertManyIfNewUnit_list_of_contains_eq_false_of_mem - h.out k_beq contains_eq_false distinct mem + DHashMap.Raw.Const.getKeyD_insertManyIfNewUnit_list_of_not_mem_of_mem + h.out k_beq not_mem distinct mem theorem getKeyD_insertManyIfNewUnit_list_of_mem [EquivBEq α] [LawfulHashable α] (h : m.WF) {l : List α} {k fallback : α} (mem : k ∈ m) : diff --git a/src/Std/Data/HashSet/Lemmas.lean b/src/Std/Data/HashSet/Lemmas.lean index c5905415a923..b943edb2fec8 100644 --- a/src/Std/Data/HashSet/Lemmas.lean +++ b/src/Std/Data/HashSet/Lemmas.lean @@ -405,74 +405,74 @@ theorem mem_of_mem_insertMany_list [EquivBEq α] [LawfulHashable α] k ∈ insertMany m l → k ∈ m := HashMap.mem_of_mem_insertManyIfNewUnit_list contains_eq_false -theorem get?_insertMany_list_of_contains_eq_false_of_contains_eq_false +theorem get?_insertMany_list_of_not_mem_of_contains_eq_false [EquivBEq α] [LawfulHashable α] {l : List α} {k : α} - (contains_eq_false : m.contains k = false) (contains_eq_false' : l.contains k = false) : + (not_mem : ¬ k ∈ m) (contains_eq_false : l.contains k = false) : get? (insertMany m l) k = none := - HashMap.getKey?_insertManyIfNewUnit_list_of_contains_eq_false_of_contains_eq_false - contains_eq_false contains_eq_false' + HashMap.getKey?_insertManyIfNewUnit_list_of_not_mem_of_contains_eq_false + not_mem contains_eq_false -theorem get?_insertMany_list_of_contains_eq_false_of_mem [EquivBEq α] [LawfulHashable α] +theorem get?_insertMany_list_of_not_mem_false_of_mem [EquivBEq α] [LawfulHashable α] {l : List α} {k k' : α} (k_beq : k == k') - (contains_eq_false : m.contains k = false) + (not_mem : ¬ k ∈ m) (distinct : l.Pairwise (fun a b => (a == b) = false)) (mem : k ∈ l) : get? (insertMany m l) k' = some k := - HashMap.getKey?_insertManyIfNewUnit_list_of_contains_eq_false_of_mem - k_beq contains_eq_false distinct mem + HashMap.getKey?_insertManyIfNewUnit_list_of_not_mem_of_mem + k_beq not_mem distinct mem theorem get?_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] {l : List α} {k : α} (mem : k ∈ m) : get? (insertMany m l) k = get? m k := HashMap.getKey?_insertManyIfNewUnit_list_of_mem mem -theorem get_insertMany_list_of_contains_eq_false_of_mem [EquivBEq α] [LawfulHashable α] +theorem get_insertMany_list_of_not_mem_of_mem [EquivBEq α] [LawfulHashable α] {l : List α} {k k' : α} (k_beq : k == k') - (contains_eq_false : m.contains k = false) + (not_mem : ¬ k ∈ m) (distinct : l.Pairwise (fun a b => (a == b) = false)) (mem : k ∈ l) {h} : get (insertMany m l) k' h = k := - HashMap.getKey_insertManyIfNewUnit_list_of_contains_eq_false_of_mem - k_beq contains_eq_false distinct mem + HashMap.getKey_insertManyIfNewUnit_list_of_not_mem_of_mem + k_beq not_mem distinct mem theorem get_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] {l : List α} {k : α} (mem : k ∈ m) {h} : get (insertMany m l) k h = get m k mem := HashMap.getKey_insertManyIfNewUnit_list_of_mem mem -theorem get!_insertMany_list_of_contains_eq_false_of_contains_eq_false +theorem get!_insertMany_list_of_not_mem_of_contains_eq_false [EquivBEq α] [LawfulHashable α] [Inhabited α] {l : List α} {k : α} - (contains_eq_false : m.contains k = false) (contains_eq_false' : l.contains k = false) : + (not_mem : ¬ k ∈ m) (contains_eq_false : l.contains k = false) : get! (insertMany m l) k = default := - HashMap.getKey!_insertManyIfNewUnit_list_of_contains_eq_false_of_contains_eq_false - contains_eq_false contains_eq_false' + HashMap.getKey!_insertManyIfNewUnit_list_of_not_mem_of_contains_eq_false + not_mem contains_eq_false -theorem get!_insertMany_list_of_contains_eq_false_of_mem [EquivBEq α] [LawfulHashable α] +theorem get!_insertMany_list_of_not_mem_of_mem [EquivBEq α] [LawfulHashable α] [Inhabited α] {l : List α} {k k' : α} (k_beq : k == k') - (contains_eq_false : m.contains k = false) + (not_mem : ¬ k ∈ m) (distinct : l.Pairwise (fun a b => (a == b) = false)) (mem : k ∈ l) : get! (insertMany m l) k' = k := - HashMap.getKey!_insertManyIfNewUnit_list_of_contains_eq_false_of_mem - k_beq contains_eq_false distinct mem + HashMap.getKey!_insertManyIfNewUnit_list_of_not_mem_of_mem + k_beq not_mem distinct mem theorem get!_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] [Inhabited α] {l : List α} {k : α} (mem : k ∈ m) : get! (insertMany m l) k = get! m k := HashMap.getKey!_insertManyIfNewUnit_list_of_mem mem -theorem getD_insertMany_list_of_contains_eq_false_of_contains_eq_false +theorem getD_insertMany_list_of_not_mem_of_contains_eq_false [EquivBEq α] [LawfulHashable α] {l : List α} {k fallback : α} - (contains_eq_false : m.contains k = false) (contains_eq_false' : l.contains k = false) : + (not_mem : ¬ k ∈ m) (contains_eq_false : l.contains k = false) : getD (insertMany m l) k fallback = fallback := - HashMap.getKeyD_insertManyIfNewUnit_list_of_contains_eq_false_of_contains_eq_false - contains_eq_false contains_eq_false' + HashMap.getKeyD_insertManyIfNewUnit_list_of_not_mem_of_contains_eq_false + not_mem contains_eq_false -theorem getD_insertMany_list_of_contains_eq_false_of_mem [EquivBEq α] [LawfulHashable α] +theorem getD_insertMany_list_of_not_mem_of_mem [EquivBEq α] [LawfulHashable α] {l : List α} {k k' fallback : α} (k_beq : k == k') - (contains_eq_false : m.contains k = false) + (not_mem : ¬ k ∈ m) (distinct : l.Pairwise (fun a b => (a == b) = false)) (mem : k ∈ l) : getD (insertMany m l) k' fallback = k := - HashMap.getKeyD_insertManyIfNewUnit_list_of_contains_eq_false_of_mem - k_beq contains_eq_false distinct mem + HashMap.getKeyD_insertManyIfNewUnit_list_of_not_mem_of_mem + k_beq not_mem distinct mem theorem getD_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] {l : List α} {k fallback : α} (mem : k ∈ m) : diff --git a/src/Std/Data/HashSet/RawLemmas.lean b/src/Std/Data/HashSet/RawLemmas.lean index 6ddd2fd99907..b527824e97c9 100644 --- a/src/Std/Data/HashSet/RawLemmas.lean +++ b/src/Std/Data/HashSet/RawLemmas.lean @@ -421,74 +421,74 @@ theorem mem_of_mem_insertMany_list [EquivBEq α] [LawfulHashable α] (h : m.WF) k ∈ insertMany m l → k ∈ m := HashMap.Raw.mem_of_mem_insertManyIfNewUnit_list h.1 contains_eq_false -theorem get?_insertMany_list_of_contains_eq_false_of_contains_eq_false +theorem get?_insertMany_list_of_not_mem_of_contains_eq_false [EquivBEq α] [LawfulHashable α] (h : m.WF) {l : List α} {k : α} - (contains_eq_false : m.contains k = false) (contains_eq_false' : l.contains k = false) : + (not_mem : ¬ k ∈ m) (contains_eq_false : l.contains k = false) : get? (insertMany m l) k = none := - HashMap.Raw.getKey?_insertManyIfNewUnit_list_of_contains_eq_false_of_contains_eq_false - h.1 contains_eq_false contains_eq_false' + HashMap.Raw.getKey?_insertManyIfNewUnit_list_of_not_mem_of_contains_eq_false + h.1 not_mem contains_eq_false -theorem get?_insertMany_list_of_contains_eq_false_of_mem [EquivBEq α] [LawfulHashable α] +theorem get?_insertMany_list_of_not_mem_of_mem [EquivBEq α] [LawfulHashable α] (h : m.WF) {l : List α} {k k' : α} (k_beq : k == k') - (contains_eq_false : m.contains k = false) + (not_mem : ¬ k ∈ m) (distinct : l.Pairwise (fun a b => (a == b) = false)) (mem : k ∈ l) : get? (insertMany m l) k' = some k := - HashMap.Raw.getKey?_insertManyIfNewUnit_list_of_contains_eq_false_of_mem - h.1 k_beq contains_eq_false distinct mem + HashMap.Raw.getKey?_insertManyIfNewUnit_list_of_not_mem_of_mem + h.1 k_beq not_mem distinct mem theorem get?_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] (h : m.WF) {l : List α} {k : α} (mem : k ∈ m) : get? (insertMany m l) k = get? m k := HashMap.Raw.getKey?_insertManyIfNewUnit_list_of_mem h.1 mem -theorem get_insertMany_list_of_contains_eq_false_of_mem [EquivBEq α] [LawfulHashable α] +theorem get_insertMany_list_of_not_mem_of_mem [EquivBEq α] [LawfulHashable α] (h : m.WF) {l : List α} {k k' : α} (k_beq : k == k') - (contains_eq_false : m.contains k = false) + (not_mem : ¬ k ∈ m) (distinct : l.Pairwise (fun a b => (a == b) = false)) (mem : k ∈ l) {h'} : get (insertMany m l) k' h' = k := - HashMap.Raw.getKey_insertManyIfNewUnit_list_of_contains_eq_false_of_mem - h.1 k_beq contains_eq_false distinct mem + HashMap.Raw.getKey_insertManyIfNewUnit_list_of_not_mem_of_mem + h.1 k_beq not_mem distinct mem theorem get_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] (h : m.WF) {l : List α} {k : α} (mem : k ∈ m) {h₃} : get (insertMany m l) k h₃ = get m k mem := HashMap.Raw.getKey_insertManyIfNewUnit_list_of_mem h.1 mem -theorem get!_insertMany_list_of_contains_eq_false_of_contains_eq_false +theorem get!_insertMany_list_of_not_mem_of_contains_eq_false [EquivBEq α] [LawfulHashable α] [Inhabited α] (h : m.WF) {l : List α} {k : α} - (contains_eq_false : m.contains k = false) (contains_eq_false' : l.contains k = false) : + (not_mem : ¬ k ∈ m) (contains_eq_false : l.contains k = false) : get! (insertMany m l) k = default := - HashMap.Raw.getKey!_insertManyIfNewUnit_list_of_contains_eq_false_of_contains_eq_false - h.1 contains_eq_false contains_eq_false' + HashMap.Raw.getKey!_insertManyIfNewUnit_list_of_not_mem_of_contains_eq_false + h.1 not_mem contains_eq_false -theorem get!_insertMany_list_of_contains_eq_false_of_mem [EquivBEq α] [LawfulHashable α] +theorem get!_insertMany_list_of_not_mem_of_mem [EquivBEq α] [LawfulHashable α] [Inhabited α] (h : m.WF) {l : List α} {k k' : α} (k_beq : k == k') - (contains_eq_false : contains m k = false) + (not_mem : ¬ k ∈ m) (distinct : l.Pairwise (fun a b => (a == b) = false)) (mem : k ∈ l) : get! (insertMany m l) k' = k := - HashMap.Raw.getKey!_insertManyIfNewUnit_list_of_contains_eq_false_of_mem - h.1 k_beq contains_eq_false distinct mem + HashMap.Raw.getKey!_insertManyIfNewUnit_list_of_not_mem_of_mem + h.1 k_beq not_mem distinct mem theorem get!_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] [Inhabited α] (h : m.WF) {l : List α} {k : α} (mem : k ∈ m) : get! (insertMany m l) k = get! m k := HashMap.Raw.getKey!_insertManyIfNewUnit_list_of_mem h.1 mem -theorem getD_insertMany_list_of_contains_eq_false_of_contains_eq_false +theorem getD_insertMany_list_of_not_mem_of_contains_eq_false [EquivBEq α] [LawfulHashable α] (h : m.WF) {l : List α} {k fallback : α} - (contains_eq_false : m.contains k = false) (contains_eq_false' : l.contains k = false) : + (not_mem : ¬ k ∈ m) (contains_eq_false : l.contains k = false) : getD (insertMany m l) k fallback = fallback := - HashMap.Raw.getKeyD_insertManyIfNewUnit_list_of_contains_eq_false_of_contains_eq_false - h.1 contains_eq_false contains_eq_false' + HashMap.Raw.getKeyD_insertManyIfNewUnit_list_of_not_mem_of_contains_eq_false + h.1 not_mem contains_eq_false -theorem getD_insertMany_list_of_contains_eq_false_of_mem [EquivBEq α] [LawfulHashable α] +theorem getD_insertMany_list_of_not_mem_of_mem [EquivBEq α] [LawfulHashable α] (h : m.WF) {l : List α} {k k' fallback : α} (k_beq : k == k') - (contains_eq_false : m.contains k = false) + (not_mem : ¬ k ∈ m) (distinct : l.Pairwise (fun a b => (a == b) = false)) (mem : k ∈ l) : getD (insertMany m l) k' fallback = k := - HashMap.Raw.getKeyD_insertManyIfNewUnit_list_of_contains_eq_false_of_mem - h.1 k_beq contains_eq_false distinct mem + HashMap.Raw.getKeyD_insertManyIfNewUnit_list_of_not_mem_of_mem + h.1 k_beq not_mem distinct mem theorem getD_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] (h : m.WF) {l : List α} {k fallback : α} (mem : k ∈ m) : From b43dfca34647df025964d2505f2c62425696550d Mon Sep 17 00:00:00 2001 From: jt0202 Date: Wed, 15 Jan 2025 10:41:21 +0100 Subject: [PATCH 83/83] Style fixes & simplified equality lemmas --- src/Std/Data/DHashMap/Internal/RawLemmas.lean | 20 ++++++------- src/Std/Data/DHashMap/Lemmas.lean | 24 ++++++++-------- src/Std/Data/DHashMap/RawLemmas.lean | 28 +++++++++---------- src/Std/Data/HashMap/Lemmas.lean | 14 +++++----- src/Std/Data/HashMap/RawLemmas.lean | 24 ++++++++-------- src/Std/Data/HashSet/Lemmas.lean | 10 +++---- src/Std/Data/HashSet/RawLemmas.lean | 12 ++++---- 7 files changed, 66 insertions(+), 66 deletions(-) diff --git a/src/Std/Data/DHashMap/Internal/RawLemmas.lean b/src/Std/Data/DHashMap/Internal/RawLemmas.lean index e08aa970e851..f04f9403231d 100644 --- a/src/Std/Data/DHashMap/Internal/RawLemmas.lean +++ b/src/Std/Data/DHashMap/Internal/RawLemmas.lean @@ -1177,7 +1177,7 @@ theorem insertManyIfNewUnit_nil : simp [insertManyIfNewUnit, Id.run] @[simp] -theorem insertManyIfNewUnit_list_singleton [EquivBEq α] [LawfulHashable α] {k : α} : +theorem insertManyIfNewUnit_list_singleton {k : α} : insertManyIfNewUnit m [k] = m.insertIfNew k () := by simp [insertManyIfNewUnit, Id.run] @@ -1319,16 +1319,16 @@ namespace Raw₀ variable [BEq α] [Hashable α] @[simp] -theorem insertMany_empty_list_nil [EquivBEq α] [LawfulHashable α] : +theorem insertMany_empty_list_nil : (insertMany empty ([] : List ((a : α) × (β a)))).1 = empty := by simp @[simp] -theorem insertMany_empty_list_singleton {k : α} {v : β k} [EquivBEq α] [LawfulHashable α] : +theorem insertMany_empty_list_singleton {k : α} {v : β k} : (insertMany empty [⟨k, v⟩]).1 = empty.insert k v := by simp -theorem insertMany_empty_list_cons [EquivBEq α] [LawfulHashable α] {k : α} {v : β k} +theorem insertMany_empty_list_cons {k : α} {v : β k} {tl : List ((a : α) × (β a))} : (insertMany empty (⟨k, v⟩ :: tl)).1 = ((empty.insert k v).insertMany tl).1 := by rw [insertMany_cons] @@ -1464,16 +1464,16 @@ namespace Const variable {β : Type v} @[simp] -theorem insertMany_empty_list_nil [EquivBEq α] [LawfulHashable α] : +theorem insertMany_empty_list_nil : (insertMany empty ([] : List (α × β))).1 = empty := by simp only [insertMany_nil] @[simp] -theorem insertMany_empty_list_singleton {k : α} {v : β} [EquivBEq α] [LawfulHashable α] : +theorem insertMany_empty_list_singleton {k : α} {v : β} : (insertMany empty [⟨k, v⟩]).1 = empty.insert k v := by simp only [insertMany_list_singleton] -theorem insertMany_empty_list_cons [EquivBEq α] [LawfulHashable α] {k : α} {v : β} +theorem insertMany_empty_list_cons {k : α} {v : β} {tl : List (α × β)} : (insertMany empty (⟨k, v⟩ :: tl)) = (insertMany (empty.insert k v) tl).1 := by rw [insertMany_cons] @@ -1606,17 +1606,17 @@ theorem isEmpty_insertMany_empty_list [EquivBEq α] [LawfulHashable α] simp [isEmpty_insertMany_list _ Raw.WF.empty₀] @[simp] -theorem insertManyIfNewUnit_empty_list_nil [EquivBEq α] [LawfulHashable α] : +theorem insertManyIfNewUnit_empty_list_nil : insertManyIfNewUnit (empty : Raw₀ α (fun _ => Unit)) ([] : List α) = (empty : Raw₀ α (fun _ => Unit)) := by simp @[simp] -theorem insertManyIfNewUnit_empty_list_singleton [EquivBEq α] [LawfulHashable α] {k : α} : +theorem insertManyIfNewUnit_empty_list_singleton {k : α} : (insertManyIfNewUnit (empty : Raw₀ α (fun _ => Unit)) [k]).1 = empty.insertIfNew k () := by simp -theorem insertManyIfNewUnit_empty_list_cons [EquivBEq α] [LawfulHashable α] {hd : α} {tl : List α} : +theorem insertManyIfNewUnit_empty_list_cons {hd : α} {tl : List α} : insertManyIfNewUnit (empty : Raw₀ α (fun _ => Unit)) (hd :: tl) = (insertManyIfNewUnit (empty.insertIfNew hd ()) tl).1 := by rw [insertManyIfNewUnit_cons] diff --git a/src/Std/Data/DHashMap/Lemmas.lean b/src/Std/Data/DHashMap/Lemmas.lean index c9e49d1c13be..7d2eebadcef7 100644 --- a/src/Std/Data/DHashMap/Lemmas.lean +++ b/src/Std/Data/DHashMap/Lemmas.lean @@ -1308,7 +1308,7 @@ theorem insertManyIfNewUnit_nil : (Raw₀.Const.insertManyIfNewUnit_nil ⟨m.1, m.2.size_buckets_pos⟩) :) @[simp] -theorem insertManyIfNewUnit_list_singleton [EquivBEq α] [LawfulHashable α] {k : α} : +theorem insertManyIfNewUnit_list_singleton {k : α} : insertManyIfNewUnit m [k] = m.insertIfNew k () := Subtype.eq (congrArg Subtype.val (Raw₀.Const.insertManyIfNewUnit_list_singleton ⟨m.1, m.2.size_buckets_pos⟩) :) @@ -1337,11 +1337,11 @@ theorem mem_of_mem_insertManyIfNewUnit_list [EquivBEq α] [LawfulHashable α] theorem getKey?_insertManyIfNewUnit_list_of_not_mem_of_contains_eq_false [EquivBEq α] [LawfulHashable α] {l : List α} {k : α} - (not_mem : ¬ k ∈ m) (contains_eq_false' : l.contains k = false) : + (not_mem : ¬ k ∈ m) (contains_eq_false : l.contains k = false) : getKey? (insertManyIfNewUnit m l) k = none := by simp only [mem_iff_contains, Bool.not_eq_true] at not_mem exact Raw₀.Const.getKey?_insertManyIfNewUnit_list_of_contains_eq_false_of_contains_eq_false - ⟨m.1, _⟩ m.2 not_mem contains_eq_false' + ⟨m.1, _⟩ m.2 not_mem contains_eq_false theorem getKey?_insertManyIfNewUnit_list_of_not_mem_of_mem [EquivBEq α] [LawfulHashable α] {l : List α} {k k' : α} (k_beq : k == k') @@ -1467,16 +1467,16 @@ end DHashMap namespace DHashMap @[simp] -theorem ofList_nil [EquivBEq α] [LawfulHashable α] : +theorem ofList_nil : ofList ([] : List ((a : α) × (β a))) = ∅ := Subtype.eq (congrArg Subtype.val (Raw₀.insertMany_empty_list_nil (α := α)) :) @[simp] -theorem ofList_singleton {k : α} {v : β k} [EquivBEq α] [LawfulHashable α] : +theorem ofList_singleton {k : α} {v : β k} : ofList [⟨k, v⟩] = (∅: DHashMap α β).insert k v := Subtype.eq (congrArg Subtype.val (Raw₀.insertMany_empty_list_singleton (α := α)) :) -theorem ofList_cons [EquivBEq α] [LawfulHashable α] {k : α} {v : β k} {tl : List ((a : α) × (β a))} : +theorem ofList_cons {k : α} {v : β k} {tl : List ((a : α) × (β a))} : ofList (⟨k, v⟩ :: tl) = ((∅ : DHashMap α β).insert k v).insertMany tl := Subtype.eq (congrArg Subtype.val (Raw₀.insertMany_empty_list_cons (α := α)) :) @@ -1611,16 +1611,16 @@ namespace Const variable {β : Type v} @[simp] -theorem ofList_nil [EquivBEq α] [LawfulHashable α] : +theorem ofList_nil : ofList ([] : List (α × β)) = ∅ := Subtype.eq (congrArg Subtype.val (Raw₀.Const.insertMany_empty_list_nil (α:= α)) :) @[simp] -theorem ofList_singleton {k : α} {v : β} [EquivBEq α] [LawfulHashable α] : +theorem ofList_singleton {k : α} {v : β} : ofList [⟨k, v⟩] = (∅ : DHashMap α (fun _ => β)).insert k v := Subtype.eq (congrArg Subtype.val (Raw₀.Const.insertMany_empty_list_singleton (α:= α)) :) -theorem ofList_cons [EquivBEq α] [LawfulHashable α] {k : α} {v : β} {tl : List (α × β)} : +theorem ofList_cons {k : α} {v : β} {tl : List (α × β)} : ofList (⟨k, v⟩ :: tl) = insertMany ((∅ : DHashMap α (fun _ => β)).insert k v) tl := Subtype.eq (congrArg Subtype.val (Raw₀.Const.insertMany_empty_list_cons (α:= α)) :) @@ -1751,16 +1751,16 @@ theorem isEmpty_ofList [EquivBEq α] [LawfulHashable α] Raw₀.Const.isEmpty_insertMany_empty_list @[simp] -theorem unitOfList_nil [EquivBEq α] [LawfulHashable α] : +theorem unitOfList_nil : unitOfList ([] : List α) = ∅ := Subtype.eq (congrArg Subtype.val (Raw₀.Const.insertManyIfNewUnit_empty_list_nil (α := α)) :) @[simp] -theorem unitOfList_singleton [EquivBEq α] [LawfulHashable α] {k : α} : +theorem unitOfList_singleton {k : α} : unitOfList [k] = (∅ : DHashMap α (fun _ => Unit)).insertIfNew k () := Subtype.eq (congrArg Subtype.val (Raw₀.Const.insertManyIfNewUnit_empty_list_singleton (α := α)) :) -theorem unitOfList_cons [EquivBEq α] [LawfulHashable α] {hd : α} {tl : List α} : +theorem unitOfList_cons {hd : α} {tl : List α} : unitOfList (hd :: tl) = insertManyIfNewUnit ((∅ : DHashMap α (fun _ => Unit)).insertIfNew hd ()) tl := Subtype.eq (congrArg Subtype.val (Raw₀.Const.insertManyIfNewUnit_empty_list_cons (α := α)) :) diff --git a/src/Std/Data/DHashMap/RawLemmas.lean b/src/Std/Data/DHashMap/RawLemmas.lean index 4aa6d3e5774a..147a2ce3bd3d 100644 --- a/src/Std/Data/DHashMap/RawLemmas.lean +++ b/src/Std/Data/DHashMap/RawLemmas.lean @@ -1229,19 +1229,19 @@ namespace Const variable {β : Type v} {m : Raw α (fun _ => β)} @[simp] -theorem insertMany_nil [EquivBEq α] [LawfulHashable α] (h : m.WF) : +theorem insertMany_nil (h : m.WF) : insertMany m [] = m := by simp_to_raw rw [Raw₀.Const.insertMany_nil] @[simp] -theorem insertMany_list_singleton [EquivBEq α] [LawfulHashable α] (h : m.WF) +theorem insertMany_list_singleton (h : m.WF) {k : α} {v : β} : insertMany m [⟨k, v⟩] = m.insert k v := by simp_to_raw rw [Raw₀.Const.insertMany_list_singleton] -theorem insertMany_cons [EquivBEq α] [LawfulHashable α] (h : m.WF) {l : List (α × β)} +theorem insertMany_cons (h : m.WF) {l : List (α × β)} {k : α} {v : β} : insertMany m (⟨k, v⟩ :: l) = insertMany (m.insert k v) l := by simp_to_raw @@ -1401,18 +1401,18 @@ theorem getD_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] (h : m.WF) variable {m : Raw α (fun _ => Unit)} @[simp] -theorem insertManyIfNewUnit_nil [EquivBEq α] [LawfulHashable α] (h : m.WF) : +theorem insertManyIfNewUnit_nil (h : m.WF) : insertManyIfNewUnit m [] = m := by simp_to_raw rw [Raw₀.Const.insertManyIfNewUnit_nil] @[simp] -theorem insertManyIfNewUnit_list_singleton [EquivBEq α] [LawfulHashable α] {k : α} (h : m.WF) : +theorem insertManyIfNewUnit_list_singleton {k : α} (h : m.WF) : insertManyIfNewUnit m [k] = m.insertIfNew k () := by simp_to_raw rw [Raw₀.Const.insertManyIfNewUnit_list_singleton] -theorem insertManyIfNewUnit_cons [EquivBEq α] [LawfulHashable α] (h : m.WF) {l : List α} {k : α} : +theorem insertManyIfNewUnit_cons (h : m.WF) {l : List α} {k : α} : insertManyIfNewUnit m (k :: l) = insertManyIfNewUnit (m.insertIfNew k ()) l := by simp_to_raw rw [Raw₀.Const.insertManyIfNewUnit_cons] @@ -1570,13 +1570,13 @@ variable [BEq α] [Hashable α] open Internal.Raw Internal.Raw₀ @[simp] -theorem ofList_nil [EquivBEq α] [LawfulHashable α] : +theorem ofList_nil : ofList ([] : List ((a : α) × (β a))) = ∅ := by simp_to_raw rw [Raw₀.insertMany_empty_list_nil] @[simp] -theorem ofList_singleton {k : α} {v : β k} [EquivBEq α] [LawfulHashable α] : +theorem ofList_singleton {k : α} {v : β k} : ofList [⟨k, v⟩] = (∅ : Raw α β).insert k v := by simp_to_raw rw [Raw₀.insertMany_empty_list_singleton] @@ -1717,18 +1717,18 @@ namespace Const variable {β : Type v} @[simp] -theorem ofList_nil [EquivBEq α] [LawfulHashable α] : +theorem ofList_nil : ofList ([] : List (α × β)) = ∅ := by simp_to_raw simp @[simp] -theorem ofList_singleton [EquivBEq α] [LawfulHashable α] {k : α} {v : β} : +theorem ofList_singleton {k : α} {v : β} : ofList [⟨k, v⟩] = (∅ : Raw α (fun _ => β)).insert k v := by simp_to_raw simp -theorem ofList_cons [EquivBEq α] [LawfulHashable α] {k : α} {v : β} {tl : List (α × β)} : +theorem ofList_cons {k : α} {v : β} {tl : List (α × β)} : ofList (⟨k, v⟩ :: tl) = insertMany ((∅ : Raw α (fun _ => β)).insert k v) tl := by simp_to_raw rw [Raw₀.Const.insertMany_empty_list_cons] @@ -1860,18 +1860,18 @@ theorem isEmpty_ofList [EquivBEq α] [LawfulHashable α] simp_to_raw using Raw₀.Const.isEmpty_insertMany_empty_list @[simp] -theorem unitOfList_nil [EquivBEq α] [LawfulHashable α] : +theorem unitOfList_nil : unitOfList ([] : List α) = ∅ := by simp_to_raw simp @[simp] -theorem unitOfList_singleton [EquivBEq α] [LawfulHashable α] {k : α} : +theorem unitOfList_singleton {k : α} : unitOfList [k] = (∅ : Raw α (fun _ => Unit)).insertIfNew k () := by simp_to_raw simp -theorem unitOfList_cons [EquivBEq α] [LawfulHashable α] {hd : α} {tl : List α} : +theorem unitOfList_cons {hd : α} {tl : List α} : unitOfList (hd :: tl) = insertManyIfNewUnit ((∅ : Raw α (fun _ => Unit)).insertIfNew hd ()) tl := by simp_to_raw rw [Raw₀.Const.insertManyIfNewUnit_empty_list_cons] diff --git a/src/Std/Data/HashMap/Lemmas.lean b/src/Std/Data/HashMap/Lemmas.lean index 19e2d073bc57..688be7b35a72 100644 --- a/src/Std/Data/HashMap/Lemmas.lean +++ b/src/Std/Data/HashMap/Lemmas.lean @@ -869,7 +869,7 @@ theorem insertManyIfNewUnit_nil : ext DHashMap.Const.insertManyIfNewUnit_nil @[simp] -theorem insertManyIfNewUnit_list_singleton [EquivBEq α] [LawfulHashable α] {k : α} : +theorem insertManyIfNewUnit_list_singleton {k : α} : insertManyIfNewUnit m [k] = m.insertIfNew k () := ext DHashMap.Const.insertManyIfNewUnit_list_singleton @@ -1014,16 +1014,16 @@ end section @[simp] -theorem ofList_nil [EquivBEq α] [LawfulHashable α] : +theorem ofList_nil : ofList ([] : List (α × β)) = ∅ := ext DHashMap.Const.ofList_nil @[simp] -theorem ofList_singleton {k : α} {v : β} [EquivBEq α] [LawfulHashable α] : +theorem ofList_singleton {k : α} {v : β} : ofList [⟨k, v⟩] = (∅ : HashMap α β).insert k v := ext DHashMap.Const.ofList_singleton -theorem ofList_cons [EquivBEq α] [LawfulHashable α] {k : α} {v : β} {tl : List (α × β)} : +theorem ofList_cons {k : α} {v : β} {tl : List (α × β)} : ofList (⟨k, v⟩ :: tl) = insertMany ((∅ : HashMap α β).insert k v) tl := ext DHashMap.Const.ofList_cons @@ -1154,16 +1154,16 @@ theorem isEmpty_ofList [EquivBEq α] [LawfulHashable α] DHashMap.Const.isEmpty_ofList @[simp] -theorem unitOfList_nil [EquivBEq α] [LawfulHashable α] : +theorem unitOfList_nil : unitOfList ([] : List α) = ∅ := ext DHashMap.Const.unitOfList_nil @[simp] -theorem unitOfList_singleton [EquivBEq α] [LawfulHashable α] {k : α} : +theorem unitOfList_singleton {k : α} : unitOfList [k] = (∅ : HashMap α Unit).insertIfNew k () := ext DHashMap.Const.unitOfList_singleton -theorem unitOfList_cons [EquivBEq α] [LawfulHashable α] {hd : α} {tl : List α} : +theorem unitOfList_cons {hd : α} {tl : List α} : unitOfList (hd :: tl) = insertManyIfNewUnit ((∅ : HashMap α Unit).insertIfNew hd ()) tl := ext DHashMap.Const.unitOfList_cons diff --git a/src/Std/Data/HashMap/RawLemmas.lean b/src/Std/Data/HashMap/RawLemmas.lean index bdcc2b0b857c..02d8442c622c 100644 --- a/src/Std/Data/HashMap/RawLemmas.lean +++ b/src/Std/Data/HashMap/RawLemmas.lean @@ -706,17 +706,17 @@ theorem distinct_keys [EquivBEq α] [LawfulHashable α] (h : m.WF) : DHashMap.Raw.distinct_keys h.out @[simp] -theorem insertMany_nil [EquivBEq α] [LawfulHashable α] (h : m.WF) : +theorem insertMany_nil (h : m.WF) : insertMany m [] = m := ext (DHashMap.Raw.Const.insertMany_nil h.out) @[simp] -theorem insertMany_list_singleton [EquivBEq α] [LawfulHashable α] (h : m.WF) +theorem insertMany_list_singleton (h : m.WF) {k : α} {v : β} : insertMany m [⟨k, v⟩] = m.insert k v := ext (DHashMap.Raw.Const.insertMany_list_singleton h.out) -theorem insertMany_cons [EquivBEq α] [LawfulHashable α] (h : m.WF) {l : List (α × β)} +theorem insertMany_cons (h : m.WF) {l : List (α × β)} {k : α} {v : β} : insertMany m (⟨k, v⟩ :: l) = insertMany (m.insert k v) l := ext (DHashMap.Raw.Const.insertMany_cons h.out) @@ -873,16 +873,16 @@ theorem getD_insertMany_list_of_mem [EquivBEq α] [LawfulHashable α] variable {m : Raw α Unit} @[simp] -theorem insertManyIfNewUnit_nil [EquivBEq α] [LawfulHashable α] (h : m.WF) : +theorem insertManyIfNewUnit_nil (h : m.WF) : insertManyIfNewUnit m [] = m := ext (DHashMap.Raw.Const.insertManyIfNewUnit_nil h.out) @[simp] -theorem insertManyIfNewUnit_list_singleton [EquivBEq α] [LawfulHashable α] (h : m.WF) {k : α} : +theorem insertManyIfNewUnit_list_singleton (h : m.WF) {k : α} : insertManyIfNewUnit m [k] = m.insertIfNew k () := ext (DHashMap.Raw.Const.insertManyIfNewUnit_list_singleton h.out) -theorem insertManyIfNewUnit_cons [EquivBEq α] [LawfulHashable α] (h : m.WF) {l : List α} {k : α} : +theorem insertManyIfNewUnit_cons (h : m.WF) {l : List α} {k : α} : insertManyIfNewUnit m (k :: l) = insertManyIfNewUnit (m.insertIfNew k ()) l := ext (DHashMap.Raw.Const.insertManyIfNewUnit_cons h.out) @@ -1032,16 +1032,16 @@ namespace Raw variable [BEq α] [Hashable α] @[simp] -theorem ofList_nil [EquivBEq α] [LawfulHashable α] : +theorem ofList_nil : ofList ([] : List (α × β)) = ∅ := ext DHashMap.Raw.Const.ofList_nil @[simp] -theorem ofList_singleton [EquivBEq α] [LawfulHashable α] {k : α} {v : β} : +theorem ofList_singleton {k : α} {v : β} : ofList [⟨k, v⟩] = (∅ : Raw α β).insert k v := ext DHashMap.Raw.Const.ofList_singleton -theorem ofList_cons [EquivBEq α] [LawfulHashable α] {k : α} {v : β} {tl : List (α × β)} : +theorem ofList_cons {k : α} {v : β} {tl : List (α × β)} : ofList (⟨k, v⟩ :: tl) = insertMany ((∅ : Raw α β).insert k v) tl := ext DHashMap.Raw.Const.ofList_cons @@ -1172,16 +1172,16 @@ theorem isEmpty_ofList [EquivBEq α] [LawfulHashable α] DHashMap.Raw.Const.isEmpty_ofList @[simp] -theorem unitOfList_nil [EquivBEq α] [LawfulHashable α] : +theorem unitOfList_nil : unitOfList ([] : List α) = ∅ := ext DHashMap.Raw.Const.unitOfList_nil @[simp] -theorem unitOfList_singleton [EquivBEq α] [LawfulHashable α] {k : α} : +theorem unitOfList_singleton {k : α} : unitOfList [k] = (∅ : Raw α Unit).insertIfNew k () := ext DHashMap.Raw.Const.unitOfList_singleton -theorem unitOfList_cons [EquivBEq α] [LawfulHashable α] {hd : α} {tl : List α} : +theorem unitOfList_cons {hd : α} {tl : List α} : unitOfList (hd :: tl) = insertManyIfNewUnit ((∅ : Raw α Unit).insertIfNew hd ()) tl := ext DHashMap.Raw.Const.unitOfList_cons diff --git a/src/Std/Data/HashSet/Lemmas.lean b/src/Std/Data/HashSet/Lemmas.lean index b943edb2fec8..b3c9d8b8858e 100644 --- a/src/Std/Data/HashSet/Lemmas.lean +++ b/src/Std/Data/HashSet/Lemmas.lean @@ -380,7 +380,7 @@ theorem insertMany_nil : ext HashMap.insertManyIfNewUnit_nil @[simp] -theorem insertMany_list_singleton [EquivBEq α] [LawfulHashable α] {k : α} : +theorem insertMany_list_singleton {k : α} : insertMany m [k] = m.insert k := ext HashMap.insertManyIfNewUnit_list_singleton @@ -412,7 +412,7 @@ theorem get?_insertMany_list_of_not_mem_of_contains_eq_false HashMap.getKey?_insertManyIfNewUnit_list_of_not_mem_of_contains_eq_false not_mem contains_eq_false -theorem get?_insertMany_list_of_not_mem_false_of_mem [EquivBEq α] [LawfulHashable α] +theorem get?_insertMany_list_of_not_mem_of_mem [EquivBEq α] [LawfulHashable α] {l : List α} {k k' : α} (k_beq : k == k') (not_mem : ¬ k ∈ m) (distinct : l.Pairwise (fun a b => (a == b) = false)) (mem : k ∈ l) : @@ -507,16 +507,16 @@ end section @[simp] -theorem ofList_nil [EquivBEq α] [LawfulHashable α] : +theorem ofList_nil : ofList ([] : List α) = ∅ := ext HashMap.unitOfList_nil @[simp] -theorem ofList_singleton [EquivBEq α] [LawfulHashable α] {k : α} : +theorem ofList_singleton {k : α} : ofList [k] = (∅ : HashSet α).insert k := ext HashMap.unitOfList_singleton -theorem ofList_cons [EquivBEq α] [LawfulHashable α] {hd : α} {tl : List α} : +theorem ofList_cons {hd : α} {tl : List α} : ofList (hd :: tl) = insertMany ((∅ : HashSet α).insert hd) tl := ext HashMap.unitOfList_cons diff --git a/src/Std/Data/HashSet/RawLemmas.lean b/src/Std/Data/HashSet/RawLemmas.lean index b527824e97c9..344cd8e2667f 100644 --- a/src/Std/Data/HashSet/RawLemmas.lean +++ b/src/Std/Data/HashSet/RawLemmas.lean @@ -391,16 +391,16 @@ theorem distinct_toList [EquivBEq α] [LawfulHashable α] (h : m.WF) : HashMap.Raw.distinct_keys h.1 @[simp] -theorem insertMany_nil [EquivBEq α] [LawfulHashable α] (h : m.WF) : +theorem insertMany_nil (h : m.WF) : insertMany m [] = m := ext (HashMap.Raw.insertManyIfNewUnit_nil h.1) @[simp] -theorem insertMany_list_singleton [EquivBEq α] [LawfulHashable α] (h : m.WF) {k : α} : +theorem insertMany_list_singleton (h : m.WF) {k : α} : insertMany m [k] = m.insert k := ext (HashMap.Raw.insertManyIfNewUnit_list_singleton h.1) -theorem insertMany_cons [EquivBEq α] [LawfulHashable α] (h : m.WF) {l : List α} {k : α} : +theorem insertMany_cons (h : m.WF) {l : List α} {k : α} : insertMany m (k :: l) = insertMany (m.insert k) l := ext (HashMap.Raw.insertManyIfNewUnit_cons h.1) @@ -519,16 +519,16 @@ theorem isEmpty_insertMany_list [EquivBEq α] [LawfulHashable α] (h : m.WF) HashMap.Raw.isEmpty_insertManyIfNewUnit_list h.1 @[simp] -theorem ofList_nil [EquivBEq α] [LawfulHashable α] : +theorem ofList_nil : ofList ([] : List α) = ∅ := ext HashMap.Raw.unitOfList_nil @[simp] -theorem ofList_singleton [EquivBEq α] [LawfulHashable α] {k : α} : +theorem ofList_singleton {k : α} : ofList [k] = (∅ : Raw α).insert k := ext HashMap.Raw.unitOfList_singleton -theorem ofList_cons [EquivBEq α] [LawfulHashable α] {hd : α} {tl : List α} : +theorem ofList_cons {hd : α} {tl : List α} : ofList (hd :: tl) = insertMany ((∅ : Raw α).insert hd) tl := ext HashMap.Raw.unitOfList_cons