From f59fa7dcaa6decce8ab7e2d2ffb3b9487ac29d42 Mon Sep 17 00:00:00 2001 From: Wojciech Rozowski Date: Fri, 31 Oct 2025 16:43:11 +0100 Subject: [PATCH 01/88] Initial `getEntry` work --- .../DHashMap/Internal/AssocList/Basic.lean | 5 +++ .../DHashMap/Internal/AssocList/Lemmas.lean | 8 ++++ src/Std/Data/DHashMap/Internal/Defs.lean | 7 ++++ src/Std/Data/DHashMap/Internal/Model.lean | 7 ++++ src/Std/Data/DHashMap/Internal/RawLemmas.lean | 10 +++++ src/Std/Data/DHashMap/Internal/WF.lean | 16 ++++++++ src/Std/Data/Internal/List/Associative.lean | 41 +++++++++++++++++++ 7 files changed, 94 insertions(+) diff --git a/src/Std/Data/DHashMap/Internal/AssocList/Basic.lean b/src/Std/Data/DHashMap/Internal/AssocList/Basic.lean index 5f400d679493..ab286e2a938c 100644 --- a/src/Std/Data/DHashMap/Internal/AssocList/Basic.lean +++ b/src/Std/Data/DHashMap/Internal/AssocList/Basic.lean @@ -122,6 +122,11 @@ def getCast [BEq α] [LawfulBEq α] (a : α) : (l : AssocList α β) → l.conta | cons k v es, h => if hka : k == a then cast (congrArg β (eq_of_beq hka)) v else es.getCast a (by rw [← h, contains, Bool.of_not_eq_true hka, Bool.false_or]) +/-- Internal implementation detail of the hash map -/ +def getEntry [BEq α] (a : α) : (l : AssocList α β) → l.contains a → (a : α) × β a + | cons k v es, h => if hka : k == a then ⟨k, v⟩ + else es.getEntry a (by rw [← h, contains, Bool.of_not_eq_true hka, Bool.false_or]) + /-- Internal implementation detail of the hash map -/ def getKey [BEq α] (a : α) : (l : AssocList α β) → l.contains a → α | cons k _ es, h => if hka : k == a then k diff --git a/src/Std/Data/DHashMap/Internal/AssocList/Lemmas.lean b/src/Std/Data/DHashMap/Internal/AssocList/Lemmas.lean index 6a62e63cd431..ce9aa1556565 100644 --- a/src/Std/Data/DHashMap/Internal/AssocList/Lemmas.lean +++ b/src/Std/Data/DHashMap/Internal/AssocList/Lemmas.lean @@ -82,6 +82,14 @@ theorem get_eq {β : Type v} [BEq α] {l : AssocList α (fun _ => β)} {a : α} · simp [contains] at h next k v t ih => simp only [get, toList_cons, List.getValue_cons, ih] +@[simp] +theorem getEntry_eq [BEq α] {l : AssocList α β} {a : α} {h} : + l.getEntry a h = List.getEntry a l.toList (contains_eq.symm.trans h) := by + induction l + · simp [contains] at h + next k v t ih => + simp only [getEntry, toList_cons, List.getEntry_cons, ih] + @[simp] theorem getCastD_eq [BEq α] [LawfulBEq α] {l : AssocList α β} {a : α} {fallback : β a} : l.getCastD a fallback = getValueCastD a l.toList fallback := by diff --git a/src/Std/Data/DHashMap/Internal/Defs.lean b/src/Std/Data/DHashMap/Internal/Defs.lean index 4b197702ce3b..9ba8368602fe 100644 --- a/src/Std/Data/DHashMap/Internal/Defs.lean +++ b/src/Std/Data/DHashMap/Internal/Defs.lean @@ -368,6 +368,13 @@ def get [BEq α] [LawfulBEq α] [Hashable α] (m : Raw₀ α β) (a : α) (hma : let idx := mkIdx buckets.size h (hash a) buckets[idx.1].getCast a hma +/-- Internal implementation detail of the hash map -/ +def getEntry [BEq α] [Hashable α] (m : Raw₀ α β) (a : α) (hma : m.contains a) : + (a : α) × β a := + let ⟨⟨_, buckets⟩, h⟩ := m + let idx := mkIdx buckets.size h (hash a) + buckets[idx.1].getEntry a hma + /-- Internal implementation detail of the hash map -/ def getD [BEq α] [LawfulBEq α] [Hashable α] (m : Raw₀ α β) (a : α) (fallback : β a) : β a := diff --git a/src/Std/Data/DHashMap/Internal/Model.lean b/src/Std/Data/DHashMap/Internal/Model.lean index de50c3d65ada..34ae8eabac43 100644 --- a/src/Std/Data/DHashMap/Internal/Model.lean +++ b/src/Std/Data/DHashMap/Internal/Model.lean @@ -294,6 +294,10 @@ def containsₘ [BEq α] [Hashable α] (m : Raw₀ α β) (a : α) : Bool := def getₘ [BEq α] [LawfulBEq α] [Hashable α] (m : Raw₀ α β) (a : α) (h : m.containsₘ a) : β a := (bucket m.1.buckets m.2 a).getCast a h +/-- Internal implementation detail of the hash map -/ +def getEntryₘ [BEq α] [Hashable α] (m : Raw₀ α β) (a : α) (h : m.containsₘ a) : (a : α) × β a := + (bucket m.1.buckets m.2 a).getEntry a h + /-- Internal implementation detail of the hash map -/ def getDₘ [BEq α] [LawfulBEq α] [Hashable α] (m : Raw₀ α β) (a : α) (fallback : β a) : β a := (m.get?ₘ a).getD fallback @@ -452,6 +456,9 @@ theorem get?_eq_get?ₘ [BEq α] [LawfulBEq α] [Hashable α] (m : Raw₀ α β) theorem get_eq_getₘ [BEq α] [LawfulBEq α] [Hashable α] (m : Raw₀ α β) (a : α) (h : m.contains a) : get m a h = getₘ m a (by exact h) := (rfl) +theorem getEntry_eq_getEntryₘ [BEq α] [Hashable α] (m : Raw₀ α β) (a : α) (h : m.contains a) : + getEntry m a h = getEntryₘ m a (by exact h) := (rfl) + theorem getD_eq_getDₘ [BEq α] [LawfulBEq α] [Hashable α] (m : Raw₀ α β) (a : α) (fallback : β a) : getD m a fallback = getDₘ m a fallback := by simp [getD, getDₘ, get?ₘ, List.getValueCastD_eq_getValueCast?, bucket] diff --git a/src/Std/Data/DHashMap/Internal/RawLemmas.lean b/src/Std/Data/DHashMap/Internal/RawLemmas.lean index d291e2946a34..4d11e260983e 100644 --- a/src/Std/Data/DHashMap/Internal/RawLemmas.lean +++ b/src/Std/Data/DHashMap/Internal/RawLemmas.lean @@ -140,6 +140,7 @@ private meta def queryMap : Std.DHashMap Name (fun _ => Name × Array (MacroM (T ⟨`getKey, (``getKey_eq_getKey, #[`(getKey_of_perm _)])⟩, ⟨`getKeyD, (``getKeyD_eq_getKeyD, #[`(getKeyD_of_perm _)])⟩, ⟨`getKey!, (``getKey!_eq_getKey!, #[`(getKey!_of_perm _)])⟩, + ⟨`getEntry, (``getEntry_eq_getEntry, #[`(getEntry_of_perm _)])⟩, ⟨`toList, (``Raw.toList_eq_toListModel, #[])⟩, ⟨`keys, (``Raw.keys_eq_keys_toListModel, #[`(perm_keys_congr_left)])⟩, ⟨`Const.toList, (``Raw.Const.toList_eq_toListModel_map, #[`(perm_map_congr_left)])⟩, @@ -376,6 +377,15 @@ theorem get_insert [LawfulBEq α] (h : m.1.WF) {k a : α} {v : β k} {h₁} : m.get a (contains_of_contains_insert _ h h₁ (Bool.eq_false_iff.2 h₂)) := by simp_to_model [insert, get] using List.getValueCast_insertEntry +theorem getEntry_insert [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {k a : α} {v : β k} {h₁} : + (m.insert k v).getEntry a h₁ = + if h₂ : k == a then + ⟨k, v⟩ + else + m.getEntry a (contains_of_contains_insert _ h h₁ (Bool.eq_false_iff.2 h₂)) := by + simp_to_model [insert, getEntry] using List.getEntry_insertEntry + + theorem get_insert_self [LawfulBEq α] (h : m.1.WF) {k : α} {v : β k} : (m.insert k v).get k (contains_insert_self _ h) = v := by simp_to_model [insert, get] using List.getValueCast_insertEntry_self diff --git a/src/Std/Data/DHashMap/Internal/WF.lean b/src/Std/Data/DHashMap/Internal/WF.lean index bdfbe5407c4a..d246c603a592 100644 --- a/src/Std/Data/DHashMap/Internal/WF.lean +++ b/src/Std/Data/DHashMap/Internal/WF.lean @@ -405,11 +405,27 @@ theorem getₘ_eq_getValue [BEq α] [Hashable α] [LawfulBEq α] {m : Raw₀ α apply_bucket_with_proof hm a AssocList.getCast List.getValueCast AssocList.getCast_eq List.getValueCast_of_perm List.getValueCast_append_of_containsKey_eq_false +theorem getEntryₘ_eq_getEntry [BEq α] [EquivBEq α] [Hashable α] [LawfulHashable α] {m : Raw₀ α β} (hm : Raw.WFImp m.1) + {a : α} {h : m.containsₘ a} : + m.getEntryₘ a h = List.getEntry a (toListModel m.1.buckets) (containsₘ_eq_containsKey hm ▸ h) := by + apply apply_bucket_with_proof hm a AssocList.getEntry List.getEntry + · intro a' l h + apply AssocList.getEntry_eq + · intro l l' a' h hd hp + apply getEntry_of_perm hd hp + · intro l l' a' h₁ h₂ + apply getEntry_append_of_containsKey_eq_false h₂ + theorem get_eq_getValueCast [BEq α] [Hashable α] [LawfulBEq α] {m : Raw₀ α β} (hm : Raw.WFImp m.1) {a : α} {h : m.contains a} : m.get a h = getValueCast a (toListModel m.1.buckets) (contains_eq_containsKey hm ▸ h) := by rw [get_eq_getₘ, getₘ_eq_getValue hm] +theorem getEntry_eq_getEntry [BEq α] [Hashable α] [EquivBEq α] [LawfulHashable α] {m : Raw₀ α β} (hm : Raw.WFImp m.1) + {a : α} {h : m.contains a} : + m.getEntry a h = List.getEntry a (toListModel m.1.buckets) (contains_eq_containsKey hm ▸ h) := by + rw [getEntry_eq_getEntryₘ, getEntryₘ_eq_getEntry hm] + theorem get!ₘ_eq_getValueCast! [BEq α] [Hashable α] [LawfulBEq α] {m : Raw₀ α β} (hm : Raw.WFImp m.1) {a : α} [Inhabited (β a)] : m.get!ₘ a = getValueCast! a (toListModel m.1.buckets) := by diff --git a/src/Std/Data/Internal/List/Associative.lean b/src/Std/Data/Internal/List/Associative.lean index 6bc4cfcc7b76..43cbc53d8950 100644 --- a/src/Std/Data/Internal/List/Associative.lean +++ b/src/Std/Data/Internal/List/Associative.lean @@ -50,9 +50,11 @@ def getEntry? [BEq α] (a : α) : List ((a : α) × β a) → Option ((a : α) @[simp] theorem getEntry?_nil [BEq α] {a : α} : getEntry? a ([] : List ((a : α) × β a)) = none := (rfl) + theorem getEntry?_cons [BEq α] {l : List ((a : α) × β a)} {k a : α} {v : β k} : getEntry? a (⟨k, v⟩ :: l) = bif k == a then some ⟨k, v⟩ else getEntry? a l := (rfl) + theorem getEntry?_eq_find [BEq α] {k : α} {l : List ((a : α) × β a)} : getEntry? k l = l.find? (·.1 == k) := by induction l using assoc_induction with @@ -483,6 +485,23 @@ theorem getEntry_eq_of_getEntry?_eq_some [BEq α] {l : List ((a : α) × β a)} (h : getEntry? a l = some ⟨k, v⟩) {h'} : getEntry a l h' = ⟨k, v⟩ := by simp [getEntry, h] +theorem getEntry_cons [BEq α] {l : List ((a : α) × β a)} {k a : α} {v : β k} + (h : containsKey a (⟨k, v⟩ :: l)) : + getEntry a (⟨k, v⟩ :: l) h = + if h' : k == a then + ⟨k, v⟩ + else + getEntry a l (containsKey_of_containsKey_cons (k := k) h (Bool.eq_false_iff.2 h')) := by + rw [getEntry, Option.get_congr getEntry?_cons] + split + case _ => + rename_i h + simp [h] + case _ => + rename_i h + simp [h] + rw [getEntry] + theorem getEntry_cons_of_beq [BEq α] {l : List ((a : α) × β a)} {k a : α} {v : β k} (h : k == a) : getEntry a (⟨k, v⟩ :: l) (containsKey_cons_of_beq (v := v) h) = ⟨k, v⟩ := by simp [getEntry, getEntry?_cons_of_true h] @@ -1620,6 +1639,14 @@ theorem getValueCast_insertEntry [BEq α] [LawfulBEq α] {l : List ((a : α) × getValueCast?_insertEntry] simp only [← getValueCast?_eq_some_getValueCast] +theorem getEntry_insertEntry [BEq α] [PartialEquivBEq α] {l : List ((a : α) × β a)} {k a : α} + {v : β k} {h} : + getEntry a (insertEntry k v l) h = if h' : k == a then ⟨k, v⟩ else getEntry a l (containsKey_of_containsKey_insertEntry h (Bool.eq_false_iff.2 h')) := by + suffices some (getEntry a (insertEntry k v l) h) = if h' : (k == a) = true then some ⟨k, v⟩ else some (getEntry a l (containsKey_of_containsKey_insertEntry h (Bool.eq_false_iff.2 h'))) by + split <;> (rename_i h; simpa [h] using this) + simp only [← getEntry?_eq_some_getEntry] + apply getEntry?_insertEntry + theorem getValueCast_insertEntry_self [BEq α] [LawfulBEq α] {l : List ((a : α) × β a)} {k : α} {v : β k} : getValueCast k (insertEntry k v l) containsKey_insertEntry_self = v := by simp [getValueCast_insertEntry] @@ -2131,6 +2158,7 @@ theorem getEntry?_of_perm [BEq α] [PartialEquivBEq α] {l l' : List ((a : α) exact ((Bool.eq_false_iff.1 hl.2.1).elim (BEq.trans h₁ (BEq.symm h₂))).elim next l₁ l₂ l₃ hl₁₂ _ ih₁ ih₂ => exact (ih₁ hl).trans (ih₂ (hl.perm (hl₁₂.symm))) + theorem containsKey_of_perm [BEq α] [PartialEquivBEq α] {l l' : List ((a : α) × β a)} {k : α} (h : Perm l l') : containsKey k l = containsKey k l' := by induction h @@ -2151,6 +2179,12 @@ theorem getValueCast_of_perm [BEq α] [LawfulBEq α] {l l' : List ((a : α) × rw [← Option.some_inj, ← getValueCast?_eq_some_getValueCast, ← getValueCast?_eq_some_getValueCast, getValueCast?_of_perm hl h] +theorem getEntry_of_perm [BEq α] [PartialEquivBEq α] {l l' : List ((a : α) × β a)} {a : α} {h'} (hl : DistinctKeys l) (h : Perm l l') : getEntry a l h' = getEntry a l' ((containsKey_of_perm h).symm.trans h') := by + suffices some (getEntry a l h') = some (getEntry a l' ((containsKey_of_perm h).symm.trans h')) by injections + rw [← getEntry?_eq_some_getEntry, ← getEntry?_eq_some_getEntry] + apply getEntry?_of_perm hl h + + theorem getValueCast!_of_perm [BEq α] [LawfulBEq α] {l l' : List ((a : α) × β a)} {a : α} [Inhabited (β a)] (hl : DistinctKeys l) (h : Perm l l') : getValueCast! a l = getValueCast! a l' := by @@ -2326,6 +2360,13 @@ theorem getEntry?_append_of_containsKey_eq_false [BEq α] {l l' : List ((a : α) (h : containsKey a l' = false) : getEntry? a (l ++ l') = getEntry? a l := by rw [getEntry?_append, getEntry?_eq_none.2 h, Option.or_none] +theorem getEntry_append_of_containsKey_eq_false [BEq α] {l l' : List ((a : α) × β a)} {a : α} + (h : containsKey a l' = false) {h₁ h₂}: getEntry a (l ++ l') h₁ = getEntry a l h₂ := by + suffices some (getEntry a (l ++ l') h₁) = some (getEntry a l h₂) by + injections + simp only [← getEntry?_eq_some_getEntry] + apply getEntry?_append_of_containsKey_eq_false h + @[simp] theorem containsKey_append [BEq α] {l l' : List ((a : α) × β a)} {a : α} : containsKey a (l ++ l') = (containsKey a l || containsKey a l') := by From 5911740da4e329654528d078c4602915cb264139 Mon Sep 17 00:00:00 2001 From: Wojciech Rozowski Date: Fri, 31 Oct 2025 18:38:47 +0100 Subject: [PATCH 02/88] end of day push --- src/Std/Data/DHashMap/Internal/AssocList/Basic.lean | 6 ++++++ src/Std/Data/DHashMap/Internal/Defs.lean | 7 +++++++ src/Std/Data/DHashMap/Internal/Model.lean | 7 +++++++ src/Std/Data/DHashMap/Internal/WF.lean | 10 ++-------- src/Std/Data/Internal/List/Associative.lean | 13 ++++++------- 5 files changed, 28 insertions(+), 15 deletions(-) diff --git a/src/Std/Data/DHashMap/Internal/AssocList/Basic.lean b/src/Std/Data/DHashMap/Internal/AssocList/Basic.lean index ab286e2a938c..00001a4a3f06 100644 --- a/src/Std/Data/DHashMap/Internal/AssocList/Basic.lean +++ b/src/Std/Data/DHashMap/Internal/AssocList/Basic.lean @@ -107,6 +107,12 @@ def getCast? [BEq α] [LawfulBEq α] (a : α) : AssocList α β → Option (β a | cons k v es => if h : k == a then some (cast (congrArg β (eq_of_beq h)) v) else es.getCast? a +/-- Internal implementation detail of the hash map -/ +def getEntry? [BEq α] (a : α) : (l : AssocList α β) → Option ((a : α) × β a) + | nil => none + | cons k v es => if k == a then some ⟨k, v⟩ + else es.getEntry? a + /-- Internal implementation detail of the hash map -/ def contains [BEq α] (a : α) : AssocList α β → Bool | nil => false diff --git a/src/Std/Data/DHashMap/Internal/Defs.lean b/src/Std/Data/DHashMap/Internal/Defs.lean index 9ba8368602fe..15eb9089058d 100644 --- a/src/Std/Data/DHashMap/Internal/Defs.lean +++ b/src/Std/Data/DHashMap/Internal/Defs.lean @@ -375,6 +375,13 @@ def getEntry [BEq α] [Hashable α] (m : Raw₀ α β) (a : α) (hma : m.contain let idx := mkIdx buckets.size h (hash a) buckets[idx.1].getEntry a hma +/-- Internal implementation detail of the hash map -/ +def getEntry? [BEq α] [Hashable α] (m : Raw₀ α β) (a : α) : + Option ((a : α) × β a) := + let ⟨⟨_, buckets⟩, h⟩ := m + let idx := mkIdx buckets.size h (hash a) + buckets[idx.1].getEntry? a + /-- Internal implementation detail of the hash map -/ def getD [BEq α] [LawfulBEq α] [Hashable α] (m : Raw₀ α β) (a : α) (fallback : β a) : β a := diff --git a/src/Std/Data/DHashMap/Internal/Model.lean b/src/Std/Data/DHashMap/Internal/Model.lean index 34ae8eabac43..a65660f96646 100644 --- a/src/Std/Data/DHashMap/Internal/Model.lean +++ b/src/Std/Data/DHashMap/Internal/Model.lean @@ -298,6 +298,10 @@ def getₘ [BEq α] [LawfulBEq α] [Hashable α] (m : Raw₀ α β) (a : α) (h def getEntryₘ [BEq α] [Hashable α] (m : Raw₀ α β) (a : α) (h : m.containsₘ a) : (a : α) × β a := (bucket m.1.buckets m.2 a).getEntry a h +/-- Internal implementation detail of the hash map -/ +def getEntry?ₘ [BEq α] [Hashable α] (m : Raw₀ α β) (a : α) : Option ((a : α) × β a) := + (bucket m.1.buckets m.2 a).getEntry? a + /-- Internal implementation detail of the hash map -/ def getDₘ [BEq α] [LawfulBEq α] [Hashable α] (m : Raw₀ α β) (a : α) (fallback : β a) : β a := (m.get?ₘ a).getD fallback @@ -459,6 +463,9 @@ theorem get_eq_getₘ [BEq α] [LawfulBEq α] [Hashable α] (m : Raw₀ α β) ( theorem getEntry_eq_getEntryₘ [BEq α] [Hashable α] (m : Raw₀ α β) (a : α) (h : m.contains a) : getEntry m a h = getEntryₘ m a (by exact h) := (rfl) +theorem getEntry?_eq_getEntry?ₘ [BEq α] [Hashable α] (m : Raw₀ α β) (a : α) : + getEntry? m a = getEntry?ₘ m a := (rfl) + theorem getD_eq_getDₘ [BEq α] [LawfulBEq α] [Hashable α] (m : Raw₀ α β) (a : α) (fallback : β a) : getD m a fallback = getDₘ m a fallback := by simp [getD, getDₘ, get?ₘ, List.getValueCastD_eq_getValueCast?, bucket] diff --git a/src/Std/Data/DHashMap/Internal/WF.lean b/src/Std/Data/DHashMap/Internal/WF.lean index d246c603a592..3154e4ab0fbe 100644 --- a/src/Std/Data/DHashMap/Internal/WF.lean +++ b/src/Std/Data/DHashMap/Internal/WF.lean @@ -407,14 +407,8 @@ theorem getₘ_eq_getValue [BEq α] [Hashable α] [LawfulBEq α] {m : Raw₀ α theorem getEntryₘ_eq_getEntry [BEq α] [EquivBEq α] [Hashable α] [LawfulHashable α] {m : Raw₀ α β} (hm : Raw.WFImp m.1) {a : α} {h : m.containsₘ a} : - m.getEntryₘ a h = List.getEntry a (toListModel m.1.buckets) (containsₘ_eq_containsKey hm ▸ h) := by - apply apply_bucket_with_proof hm a AssocList.getEntry List.getEntry - · intro a' l h - apply AssocList.getEntry_eq - · intro l l' a' h hd hp - apply getEntry_of_perm hd hp - · intro l l' a' h₁ h₂ - apply getEntry_append_of_containsKey_eq_false h₂ + m.getEntryₘ a h = List.getEntry a (toListModel m.1.buckets) (containsₘ_eq_containsKey hm ▸ h) := + apply_bucket_with_proof hm a AssocList.getEntry List.getEntry AssocList.getEntry_eq getEntry_of_perm getEntry_append_of_containsKey_eq_false theorem get_eq_getValueCast [BEq α] [Hashable α] [LawfulBEq α] {m : Raw₀ α β} (hm : Raw.WFImp m.1) {a : α} {h : m.contains a} : diff --git a/src/Std/Data/Internal/List/Associative.lean b/src/Std/Data/Internal/List/Associative.lean index 43cbc53d8950..bb0749edd06b 100644 --- a/src/Std/Data/Internal/List/Associative.lean +++ b/src/Std/Data/Internal/List/Associative.lean @@ -2360,13 +2360,6 @@ theorem getEntry?_append_of_containsKey_eq_false [BEq α] {l l' : List ((a : α) (h : containsKey a l' = false) : getEntry? a (l ++ l') = getEntry? a l := by rw [getEntry?_append, getEntry?_eq_none.2 h, Option.or_none] -theorem getEntry_append_of_containsKey_eq_false [BEq α] {l l' : List ((a : α) × β a)} {a : α} - (h : containsKey a l' = false) {h₁ h₂}: getEntry a (l ++ l') h₁ = getEntry a l h₂ := by - suffices some (getEntry a (l ++ l') h₁) = some (getEntry a l h₂) by - injections - simp only [← getEntry?_eq_some_getEntry] - apply getEntry?_append_of_containsKey_eq_false h - @[simp] theorem containsKey_append [BEq α] {l l' : List ((a : α) × β a)} {a : α} : containsKey a (l ++ l') = (containsKey a l || containsKey a l') := by @@ -2416,6 +2409,12 @@ theorem getValueCast_append_of_containsKey_eq_false [BEq α] [LawfulBEq α] rw [← Option.some_inj, ← getValueCast?_eq_some_getValueCast, ← getValueCast?_eq_some_getValueCast, getValueCast?_append_of_containsKey_eq_false hl'] +theorem getEntry_append_of_containsKey_eq_false [BEq α] {l l' : List ((a : α) × β a)} {a : α} {h} (hl' : containsKey a l' = false) : getEntry a (l ++ l') h = getEntry a l ((containsKey_append_of_not_contains_right hl').symm.trans h) := by + suffices some (getEntry a (l ++ l') h) = some (getEntry a l ((containsKey_append_of_not_contains_right hl').symm.trans h)) by + injections + simp only [← getEntry?_eq_some_getEntry] + apply getEntry?_append_of_containsKey_eq_false hl' + theorem getKey?_append_of_containsKey_eq_false [BEq α] [PartialEquivBEq α] {l l' : List ((a : α) × β a)} {a : α} (hl' : containsKey a l' = false) : getKey? a (l ++ l') = getKey? a l := by From e44545c63c06273eb2b1b3978ad722750e77c6f4 Mon Sep 17 00:00:00 2001 From: Wojciech Rozowski Date: Mon, 3 Nov 2025 15:16:29 +0100 Subject: [PATCH 03/88] push intermediate changes --- .../Data/DHashMap/Internal/AssocList/Basic.lean | 1 - .../Data/DHashMap/Internal/AssocList/Lemmas.lean | 8 ++++++++ src/Std/Data/DHashMap/Internal/RawLemmas.lean | 4 ++++ src/Std/Data/DHashMap/Internal/WF.lean | 14 ++++++++++++-- src/Std/Data/Internal/List/Associative.lean | 8 ++++++++ 5 files changed, 32 insertions(+), 3 deletions(-) diff --git a/src/Std/Data/DHashMap/Internal/AssocList/Basic.lean b/src/Std/Data/DHashMap/Internal/AssocList/Basic.lean index 00001a4a3f06..b73ded470597 100644 --- a/src/Std/Data/DHashMap/Internal/AssocList/Basic.lean +++ b/src/Std/Data/DHashMap/Internal/AssocList/Basic.lean @@ -98,7 +98,6 @@ variable {β : Type v} def get? [BEq α] (a : α) : AssocList α (fun _ => β) → Option β | nil => none | cons k v es => bif k == a then some v else get? a es - end /-- Internal implementation detail of the hash map -/ diff --git a/src/Std/Data/DHashMap/Internal/AssocList/Lemmas.lean b/src/Std/Data/DHashMap/Internal/AssocList/Lemmas.lean index ce9aa1556565..a8abd66ec54f 100644 --- a/src/Std/Data/DHashMap/Internal/AssocList/Lemmas.lean +++ b/src/Std/Data/DHashMap/Internal/AssocList/Lemmas.lean @@ -90,6 +90,14 @@ theorem getEntry_eq [BEq α] {l : AssocList α β} {a : α} {h} : next k v t ih => simp only [getEntry, toList_cons, List.getEntry_cons, ih] +@[simp] +theorem getEntry?_eq [BEq α] {l : AssocList α β} {a : α} : + l.getEntry? a = List.getEntry? a l.toList := by + induction l + · simp only [getEntry?, toList_nil, getEntry?_nil] + next k v t ih => + simp only [getEntry?, ih, toList_cons, getEntry?_cons, Bool.ite_eq_cond_iff] + @[simp] theorem getCastD_eq [BEq α] [LawfulBEq α] {l : AssocList α β} {a : α} {fallback : β a} : l.getCastD a fallback = getValueCastD a l.toList fallback := by diff --git a/src/Std/Data/DHashMap/Internal/RawLemmas.lean b/src/Std/Data/DHashMap/Internal/RawLemmas.lean index 4d11e260983e..40b12a0b7d59 100644 --- a/src/Std/Data/DHashMap/Internal/RawLemmas.lean +++ b/src/Std/Data/DHashMap/Internal/RawLemmas.lean @@ -141,6 +141,7 @@ private meta def queryMap : Std.DHashMap Name (fun _ => Name × Array (MacroM (T ⟨`getKeyD, (``getKeyD_eq_getKeyD, #[`(getKeyD_of_perm _)])⟩, ⟨`getKey!, (``getKey!_eq_getKey!, #[`(getKey!_of_perm _)])⟩, ⟨`getEntry, (``getEntry_eq_getEntry, #[`(getEntry_of_perm _)])⟩, + ⟨`getEntry?, (``getEntry?_eq_getEntry?, #[`(getEntry?_of_perm _)])⟩, ⟨`toList, (``Raw.toList_eq_toListModel, #[])⟩, ⟨`keys, (``Raw.keys_eq_keys_toListModel, #[`(perm_keys_congr_left)])⟩, ⟨`Const.toList, (``Raw.Const.toList_eq_toListModel_map, #[`(perm_map_congr_left)])⟩, @@ -385,6 +386,9 @@ theorem getEntry_insert [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {k a : α m.getEntry a (contains_of_contains_insert _ h h₁ (Bool.eq_false_iff.2 h₂)) := by simp_to_model [insert, getEntry] using List.getEntry_insertEntry +theorem getEntry_congr [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {a b : α} (hab : a == b) {h₁ h₂} : + getEntry m a h₁ = getEntry m b h₂ := by + simp_to_model [getEntry, contains] using List.getEntry_congr theorem get_insert_self [LawfulBEq α] (h : m.1.WF) {k : α} {v : β k} : (m.insert k v).get k (contains_insert_self _ h) = v := by diff --git a/src/Std/Data/DHashMap/Internal/WF.lean b/src/Std/Data/DHashMap/Internal/WF.lean index 3154e4ab0fbe..c1a88f7b07bc 100644 --- a/src/Std/Data/DHashMap/Internal/WF.lean +++ b/src/Std/Data/DHashMap/Internal/WF.lean @@ -405,21 +405,31 @@ theorem getₘ_eq_getValue [BEq α] [Hashable α] [LawfulBEq α] {m : Raw₀ α apply_bucket_with_proof hm a AssocList.getCast List.getValueCast AssocList.getCast_eq List.getValueCast_of_perm List.getValueCast_append_of_containsKey_eq_false -theorem getEntryₘ_eq_getEntry [BEq α] [EquivBEq α] [Hashable α] [LawfulHashable α] {m : Raw₀ α β} (hm : Raw.WFImp m.1) +theorem getEntryₘ_eq_getEntry [BEq α] [PartialEquivBEq α] [Hashable α] [LawfulHashable α] {m : Raw₀ α β} (hm : Raw.WFImp m.1) {a : α} {h : m.containsₘ a} : m.getEntryₘ a h = List.getEntry a (toListModel m.1.buckets) (containsₘ_eq_containsKey hm ▸ h) := apply_bucket_with_proof hm a AssocList.getEntry List.getEntry AssocList.getEntry_eq getEntry_of_perm getEntry_append_of_containsKey_eq_false +theorem getEntry?ₘ_eq_getEntry? [BEq α] [PartialEquivBEq α] [Hashable α] [LawfulHashable α] {m : Raw₀ α β} (hm : Raw.WFImp m.1) + {a : α} : + m.getEntry?ₘ a = List.getEntry? a (toListModel m.1.buckets) := + apply_bucket hm AssocList.getEntry?_eq getEntry?_of_perm getEntry?_append_of_containsKey_eq_false + theorem get_eq_getValueCast [BEq α] [Hashable α] [LawfulBEq α] {m : Raw₀ α β} (hm : Raw.WFImp m.1) {a : α} {h : m.contains a} : m.get a h = getValueCast a (toListModel m.1.buckets) (contains_eq_containsKey hm ▸ h) := by rw [get_eq_getₘ, getₘ_eq_getValue hm] -theorem getEntry_eq_getEntry [BEq α] [Hashable α] [EquivBEq α] [LawfulHashable α] {m : Raw₀ α β} (hm : Raw.WFImp m.1) +theorem getEntry_eq_getEntry [BEq α] [Hashable α] [PartialEquivBEq α] [LawfulHashable α] {m : Raw₀ α β} (hm : Raw.WFImp m.1) {a : α} {h : m.contains a} : m.getEntry a h = List.getEntry a (toListModel m.1.buckets) (contains_eq_containsKey hm ▸ h) := by rw [getEntry_eq_getEntryₘ, getEntryₘ_eq_getEntry hm] +theorem getEntry?_eq_getEntry? [BEq α] [Hashable α] [PartialEquivBEq α] [LawfulHashable α] {m : Raw₀ α β} (hm : Raw.WFImp m.1) + {a : α} : + m.getEntry? a = List.getEntry? a (toListModel m.1.buckets) := by + rw [getEntry?_eq_getEntry?ₘ, getEntry?ₘ_eq_getEntry? hm] + theorem get!ₘ_eq_getValueCast! [BEq α] [Hashable α] [LawfulBEq α] {m : Raw₀ α β} (hm : Raw.WFImp m.1) {a : α} [Inhabited (β a)] : m.get!ₘ a = getValueCast! a (toListModel m.1.buckets) := by diff --git a/src/Std/Data/Internal/List/Associative.lean b/src/Std/Data/Internal/List/Associative.lean index bb0749edd06b..16582ad87f9a 100644 --- a/src/Std/Data/Internal/List/Associative.lean +++ b/src/Std/Data/Internal/List/Associative.lean @@ -502,6 +502,13 @@ theorem getEntry_cons [BEq α] {l : List ((a : α) × β a)} {k a : α} {v : β simp [h] rw [getEntry] +theorem getEntry_congr [BEq α] [PartialEquivBEq α] {l : List ((a : α) × β a)} {a b : α} + (h : a == b) {h₁ h₂} : getEntry a l h₁ = getEntry b l h₂ := by + suffices some (getEntry a l h₁) = some (getEntry b l h₂) by + injections + simp only [← getEntry?_eq_some_getEntry] + apply getEntry?_congr h + theorem getEntry_cons_of_beq [BEq α] {l : List ((a : α) × β a)} {k a : α} {v : β k} (h : k == a) : getEntry a (⟨k, v⟩ :: l) (containsKey_cons_of_beq (v := v) h) = ⟨k, v⟩ := by simp [getEntry, getEntry?_cons_of_true h] @@ -618,6 +625,7 @@ def getValueCast [BEq α] [LawfulBEq α] (a : α) (l : List ((a : α) × β a)) β a := (getValueCast? a l).get <| containsKey_eq_isSome_getValueCast?.symm.trans h + theorem getValueCast?_eq_some_getValueCast [BEq α] [LawfulBEq α] {l : List ((a : α) × β a)} {a : α} (h : containsKey a l) : getValueCast? a l = some (getValueCast a l h) := by simp [getValueCast] From cb7ac1d0e42b2fe7505a77db143e9ad3196a5d38 Mon Sep 17 00:00:00 2001 From: Wojciech Rozowski Date: Mon, 3 Nov 2025 20:09:50 +0100 Subject: [PATCH 04/88] end of day push --- src/Std/Data/DHashMap/Internal/RawLemmas.lean | 4 ++++ src/Std/Data/Internal/List/Associative.lean | 23 +++++++++++++++++++ 2 files changed, 27 insertions(+) diff --git a/src/Std/Data/DHashMap/Internal/RawLemmas.lean b/src/Std/Data/DHashMap/Internal/RawLemmas.lean index 40b12a0b7d59..855f1689fd0e 100644 --- a/src/Std/Data/DHashMap/Internal/RawLemmas.lean +++ b/src/Std/Data/DHashMap/Internal/RawLemmas.lean @@ -386,6 +386,10 @@ theorem getEntry_insert [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {k a : α m.getEntry a (contains_of_contains_insert _ h h₁ (Bool.eq_false_iff.2 h₂)) := by simp_to_model [insert, getEntry] using List.getEntry_insertEntry +theorem getEntry_eq [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {k k': α} {v : β k'} {h} : + m.getEntry k h = ⟨k', v⟩ → k == k' := by + simp_to_model [getEntry, contains] using List.getEntry_eq + theorem getEntry_congr [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {a b : α} (hab : a == b) {h₁ h₂} : getEntry m a h₁ = getEntry m b h₂ := by simp_to_model [getEntry, contains] using List.getEntry_congr diff --git a/src/Std/Data/Internal/List/Associative.lean b/src/Std/Data/Internal/List/Associative.lean index 16582ad87f9a..cedfd9e15547 100644 --- a/src/Std/Data/Internal/List/Associative.lean +++ b/src/Std/Data/Internal/List/Associative.lean @@ -1655,6 +1655,29 @@ theorem getEntry_insertEntry [BEq α] [PartialEquivBEq α] {l : List ((a : α) simp only [← getEntry?_eq_some_getEntry] apply getEntry?_insertEntry +theorem getEntry_eq [BEq α] [PartialEquivBEq α] {l : List ((a : α) × β a)} {k k' : α} + {v : β k'} {h} : + getEntry k l h = ⟨k', v⟩ → k == k' := by + intro w + induction l using assoc_induction + case nil => simp only [containsKey_nil, Bool.false_eq_true] at h + case cons k'' v' t ih => + have w' : some (getEntry k (⟨k'', v'⟩ :: t) h) = some ⟨k', v⟩ := by simp only [w] + rw [← getEntry?_eq_some_getEntry] at w' + simp [getEntry?_cons] at w' + apply Or.elim <| Classical.em <| k'' == k + . intro heq + simp [heq] at w' + rw [w'.1] at heq + rw [BEq.symm] + exact heq + . intro heq + simp [heq] at w' + simp [heq] at h + rw [getEntry?_eq_some_getEntry h] at w' + injections heq₂ + exact @ih h heq₂ + theorem getValueCast_insertEntry_self [BEq α] [LawfulBEq α] {l : List ((a : α) × β a)} {k : α} {v : β k} : getValueCast k (insertEntry k v l) containsKey_insertEntry_self = v := by simp [getValueCast_insertEntry] From b54f6838e9fd68f799846b12b86eb29b61229b28 Mon Sep 17 00:00:00 2001 From: Wojciech Rozowski Date: Tue, 4 Nov 2025 12:36:46 +0100 Subject: [PATCH 05/88] chore: push lemma about `getEntry` --- src/Std/Data/DHashMap/Internal/RawLemmas.lean | 6 ++--- src/Std/Data/Internal/List/Associative.lean | 27 ++++--------------- 2 files changed, 8 insertions(+), 25 deletions(-) diff --git a/src/Std/Data/DHashMap/Internal/RawLemmas.lean b/src/Std/Data/DHashMap/Internal/RawLemmas.lean index 855f1689fd0e..b4ce6c6dbb1d 100644 --- a/src/Std/Data/DHashMap/Internal/RawLemmas.lean +++ b/src/Std/Data/DHashMap/Internal/RawLemmas.lean @@ -386,9 +386,9 @@ theorem getEntry_insert [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {k a : α m.getEntry a (contains_of_contains_insert _ h h₁ (Bool.eq_false_iff.2 h₂)) := by simp_to_model [insert, getEntry] using List.getEntry_insertEntry -theorem getEntry_eq [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {k k': α} {v : β k'} {h} : - m.getEntry k h = ⟨k', v⟩ → k == k' := by - simp_to_model [getEntry, contains] using List.getEntry_eq +theorem beq_of_getEntry_eq [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {p : ((a : α) × β a)} {k : α} {h₁} : + (w : m.getEntry k h₁ = p) → p.fst == k := by + simp_to_model [getEntry, contains] using List.beq_of_getEntry_eq theorem getEntry_congr [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {a b : α} (hab : a == b) {h₁ h₂} : getEntry m a h₁ = getEntry m b h₂ := by diff --git a/src/Std/Data/Internal/List/Associative.lean b/src/Std/Data/Internal/List/Associative.lean index cedfd9e15547..28feb9d97e15 100644 --- a/src/Std/Data/Internal/List/Associative.lean +++ b/src/Std/Data/Internal/List/Associative.lean @@ -1655,28 +1655,11 @@ theorem getEntry_insertEntry [BEq α] [PartialEquivBEq α] {l : List ((a : α) simp only [← getEntry?_eq_some_getEntry] apply getEntry?_insertEntry -theorem getEntry_eq [BEq α] [PartialEquivBEq α] {l : List ((a : α) × β a)} {k k' : α} - {v : β k'} {h} : - getEntry k l h = ⟨k', v⟩ → k == k' := by - intro w - induction l using assoc_induction - case nil => simp only [containsKey_nil, Bool.false_eq_true] at h - case cons k'' v' t ih => - have w' : some (getEntry k (⟨k'', v'⟩ :: t) h) = some ⟨k', v⟩ := by simp only [w] - rw [← getEntry?_eq_some_getEntry] at w' - simp [getEntry?_cons] at w' - apply Or.elim <| Classical.em <| k'' == k - . intro heq - simp [heq] at w' - rw [w'.1] at heq - rw [BEq.symm] - exact heq - . intro heq - simp [heq] at w' - simp [heq] at h - rw [getEntry?_eq_some_getEntry h] at w' - injections heq₂ - exact @ih h heq₂ +theorem beq_of_getEntry_eq [BEq α] [PartialEquivBEq α] {l : List ((a : α) × β a)} {p : ((a : α) × β a)} + {k : α} {h₁} (h : getEntry k l h₁ = p) : p.fst == k := by + have heq1 : some (getEntry k l h₁) = some p := by simp only [h] + rw [← getEntry?_eq_some_getEntry] at heq1 + exact @beq_of_getEntry?_eq_some α β _ l k p heq1 theorem getValueCast_insertEntry_self [BEq α] [LawfulBEq α] {l : List ((a : α) × β a)} {k : α} {v : β k} : getValueCast k (insertEntry k v l) containsKey_insertEntry_self = v := by From 037c5b8ff609eee44281689e2822d8643139c5ca Mon Sep 17 00:00:00 2001 From: Wojciech Rozowski Date: Tue, 4 Nov 2025 17:58:55 +0100 Subject: [PATCH 06/88] foldl model --- src/Std/Data/DHashMap/Basic.lean | 4 ++++ src/Std/Data/DHashMap/Internal/Defs.lean | 18 ++++++++++++++++++ src/Std/Data/DHashMap/Internal/Model.lean | 4 ++++ src/Std/Data/DHashMap/Internal/WF.lean | 21 +++++++++++++++++++++ 4 files changed, 47 insertions(+) diff --git a/src/Std/Data/DHashMap/Basic.lean b/src/Std/Data/DHashMap/Basic.lean index 1316a4a28e06..952fab9ffb13 100644 --- a/src/Std/Data/DHashMap/Basic.lean +++ b/src/Std/Data/DHashMap/Basic.lean @@ -319,6 +319,10 @@ This function always merges the smaller map into the larger map, so the expected inner := Raw₀.union ⟨m₁.1, m₁.2.size_buckets_pos⟩ ⟨m₂.1, m₂.2.size_buckets_pos⟩ wf := Std.DHashMap.Raw.WF.union₀ m₁.2 m₂.2 +@[inline] def inter [BEq α] [Hashable α] (m₁ m₂ : DHashMap α β) : DHashMap α β where + inner := Raw₀.inter ⟨m₁.1, m₁.2.size_buckets_pos⟩ ⟨m₂.1, m₂.2.size_buckets_pos⟩ + wf := sorry + instance [BEq α] [Hashable α] : Union (DHashMap α β) := ⟨union⟩ section Unverified diff --git a/src/Std/Data/DHashMap/Internal/Defs.lean b/src/Std/Data/DHashMap/Internal/Defs.lean index 15eb9089058d..ce0d057fe41f 100644 --- a/src/Std/Data/DHashMap/Internal/Defs.lean +++ b/src/Std/Data/DHashMap/Internal/Defs.lean @@ -447,10 +447,28 @@ def insertMany {ρ : Type w} [ForIn Id ρ ((a : α) × β a)] [BEq α] [Hashable r := ⟨r.1.insertIfNew a b, fun _ h hm => h (r.2 _ h hm)⟩ return r +/-- Internal implementation detail of the hash map -/ +@[inline] def foldl {δ : Type w} (f : δ → (a : α) → β a → δ) (init : δ) + (m : Raw₀ α β) : { d : δ // ∀ (P : δ → Prop), + (∀ {d' a b}, P d' → P (f d' a b)) → P init → P d } := Id.run do + let mut r : { d : δ // ∀ (P : δ → Prop), + (∀ {d' a b}, P d' → P (f d' a b)) → P init → P d } := ⟨init, fun _ _ => id⟩ + for ⟨a, b⟩ in m.1 do + r := ⟨f r.1 a b, fun _ h hm => h (r.2 _ h hm)⟩ + return r + +/-- Internal implementation detail of the hash map -/ +def interSmaller [BEq α] [Hashable α] (m₁ m₂ : Raw₀ α β) : Raw₀ α β := + (m₂.foldl (fun sofar k _ => match m₁.getEntry? k with | some kv' => sofar.insert kv'.1 kv'.2 | none => sofar) emptyWithCapacity).1 + /-- Internal implementation detail of the hash map -/ @[inline] def union [BEq α] [Hashable α] (m₁ m₂ : Raw₀ α β) : Raw₀ α β := if m₁.1.size ≤ m₂.1.size then (m₂.insertManyIfNew m₁.1).1 else (m₁.insertMany m₂.1).1 +/-- Internal implementation detail of the hash map -/ +def inter [BEq α] [Hashable α] (m₁ m₂ : Raw₀ α β) : Raw₀ α β := + if m₁.1.size ≤ m₂.1.size then m₁.filter fun k _ => m₂.contains k else interSmaller m₁ m₂ + section variable {β : Type v} diff --git a/src/Std/Data/DHashMap/Internal/Model.lean b/src/Std/Data/DHashMap/Internal/Model.lean index a65660f96646..a87b0e95ecb8 100644 --- a/src/Std/Data/DHashMap/Internal/Model.lean +++ b/src/Std/Data/DHashMap/Internal/Model.lean @@ -412,6 +412,10 @@ def unionₘ [BEq α] [Hashable α] (m₁ m₂ : Raw₀ α β) : Raw₀ α β := else insertListₘ m₁ (toListModel m₂.1.buckets) +/-- Internal implementation detail of the hash map -/ +def foldlₘ {δ} (f : δ → (a : α) → β a → δ) (init : δ) (m : List ((a : α) × β a)) : δ := + m.foldl (fun acc ⟨k,v⟩ => f acc k v) init + section variable {β : Type v} diff --git a/src/Std/Data/DHashMap/Internal/WF.lean b/src/Std/Data/DHashMap/Internal/WF.lean index c1a88f7b07bc..87d875f59b26 100644 --- a/src/Std/Data/DHashMap/Internal/WF.lean +++ b/src/Std/Data/DHashMap/Internal/WF.lean @@ -1290,6 +1290,27 @@ theorem toListModel_union [BEq α] [Hashable α] [EquivBEq α] [LawfulHashable rw [union_eq_unionₘ] exact toListModel_unionₘ h₁ h₂ +/-! # `foldl`-/ +theorem foldl_eq_foldlₘ {δ} (f : δ → (a : α) → β a → δ) (init : δ) (m : Raw₀ α β) : + foldl f init m = foldlₘ f init (m.1.toList) := by + simp [foldl,foldlₘ] + simp [ForIn.forIn] + rw [Raw.forIn_eq_forIn_toListModel] + rw [Raw.toList_eq_toListModel] + generalize (toListModel m.val.buckets) = l + simp only [forIn_pure_yield_eq_foldl, Id.run_pure] + suffices ∀ (t : { d : δ // ∀ (P : δ → Prop), + (∀ {d' : δ} {a : α} {b : β a}, P d' → P (f d' a b)) → P init → P d }), + (List.foldl (fun d' p => ⟨f d'.val p.1 p.2, fun P h₁ h₂ => h₁ (d'.2 _ h₁ h₂)⟩) t l).val = + foldlₘ f t.val l from this _ + intro t + induction l generalizing t with + | nil => simp [foldlₘ] + | cons hd tl ih => + simp only [List.foldl_cons, foldlₘ] + simp [foldlₘ] at ih + apply ih + /-! # `Const.insertListₘ` -/ theorem Const.toListModel_insertListₘ {β : Type v} [BEq α] [Hashable α] [EquivBEq α] From cd12c0a384494b51b385f87eae72a8143ac4e9df Mon Sep 17 00:00:00 2001 From: Wojciech Rozowski Date: Tue, 4 Nov 2025 21:40:37 +0100 Subject: [PATCH 07/88] chore: intermediate work --- src/Std/Data/DHashMap/Internal/Defs.lean | 8 +- src/Std/Data/DHashMap/Internal/Model.lean | 14 ++ src/Std/Data/DHashMap/Internal/WF.lean | 178 +++++++++++++++++++++- 3 files changed, 194 insertions(+), 6 deletions(-) diff --git a/src/Std/Data/DHashMap/Internal/Defs.lean b/src/Std/Data/DHashMap/Internal/Defs.lean index ce0d057fe41f..a7a0d14181e4 100644 --- a/src/Std/Data/DHashMap/Internal/Defs.lean +++ b/src/Std/Data/DHashMap/Internal/Defs.lean @@ -457,9 +457,15 @@ def insertMany {ρ : Type w} [ForIn Id ρ ((a : α) × β a)] [BEq α] [Hashable r := ⟨f r.1 a b, fun _ h hm => h (r.2 _ h hm)⟩ return r +/-- Internal implementation detail of the hash map -/ +def interSmallerFn [BEq α] [Hashable α] (m sofar : Raw₀ α β) (k : α) : Raw₀ α β := + match m.getEntry? k with + | some kv' => sofar.insert kv'.1 kv'.2 + | none => sofar + /-- Internal implementation detail of the hash map -/ def interSmaller [BEq α] [Hashable α] (m₁ m₂ : Raw₀ α β) : Raw₀ α β := - (m₂.foldl (fun sofar k _ => match m₁.getEntry? k with | some kv' => sofar.insert kv'.1 kv'.2 | none => sofar) emptyWithCapacity).1 + (m₂.foldl (fun sofar k _ => interSmallerFn m₁ sofar k) emptyWithCapacity).1 /-- Internal implementation detail of the hash map -/ @[inline] def union [BEq α] [Hashable α] (m₁ m₂ : Raw₀ α β) : Raw₀ α β := diff --git a/src/Std/Data/DHashMap/Internal/Model.lean b/src/Std/Data/DHashMap/Internal/Model.lean index a87b0e95ecb8..e3a4d1cb91d9 100644 --- a/src/Std/Data/DHashMap/Internal/Model.lean +++ b/src/Std/Data/DHashMap/Internal/Model.lean @@ -416,6 +416,12 @@ def unionₘ [BEq α] [Hashable α] (m₁ m₂ : Raw₀ α β) : Raw₀ α β := def foldlₘ {δ} (f : δ → (a : α) → β a → δ) (init : δ) (m : List ((a : α) × β a)) : δ := m.foldl (fun acc ⟨k,v⟩ => f acc k v) init +/-- Internal implementation detail of the hash map -/ +def interSmallerFnₘ [BEq α] [Hashable α] (m sofar : Raw₀ α β) (k : α) : Raw₀ α β := + match m.getEntry?ₘ k with + | some kv' => sofar.insertₘ kv'.1 kv'.2 + | none => sofar + section variable {β : Type v} @@ -654,6 +660,14 @@ theorem insertManyIfNew_eq_insertListIfNewₘ [BEq α] [Hashable α] (m : Raw₀ simp only [List.foldl_cons, insertListIfNewₘ] apply ih +theorem interSmallerFn_eq_interSmallerFnₘ [BEq α] [Hashable α] (m sofar : Raw₀ α β) (k : α) : + interSmallerFn m sofar k = interSmallerFnₘ m sofar k := by + rw [interSmallerFn, interSmallerFnₘ] + rw [getEntry?_eq_getEntry?ₘ] + congr + ext + rw [insert_eq_insertₘ] + section variable {β : Type v} diff --git a/src/Std/Data/DHashMap/Internal/WF.lean b/src/Std/Data/DHashMap/Internal/WF.lean index 87d875f59b26..db82f9b78829 100644 --- a/src/Std/Data/DHashMap/Internal/WF.lean +++ b/src/Std/Data/DHashMap/Internal/WF.lean @@ -1293,10 +1293,8 @@ theorem toListModel_union [BEq α] [Hashable α] [EquivBEq α] [LawfulHashable /-! # `foldl`-/ theorem foldl_eq_foldlₘ {δ} (f : δ → (a : α) → β a → δ) (init : δ) (m : Raw₀ α β) : foldl f init m = foldlₘ f init (m.1.toList) := by - simp [foldl,foldlₘ] - simp [ForIn.forIn] - rw [Raw.forIn_eq_forIn_toListModel] - rw [Raw.toList_eq_toListModel] + simp only [foldl, ForIn.forIn, bind_pure_comp, map_pure, bind_pure, foldlₘ] + rw [Raw.forIn_eq_forIn_toListModel, Raw.toList_eq_toListModel] generalize (toListModel m.val.buckets) = l simp only [forIn_pure_yield_eq_foldl, Id.run_pure] suffices ∀ (t : { d : δ // ∀ (P : δ → Prop), @@ -1308,9 +1306,179 @@ theorem foldl_eq_foldlₘ {δ} (f : δ → (a : α) → β a → δ) (init : δ) | nil => simp [foldlₘ] | cons hd tl ih => simp only [List.foldl_cons, foldlₘ] - simp [foldlₘ] at ih + simp only [foldlₘ, Subtype.forall] at ih apply ih +-- /-- Internal implementation detail of the hash map -/ +-- def interSmaller [BEq α] [Hashable α] (m₁ m₂ : Raw₀ α β) : Raw₀ α β := +-- (m₂.foldl (fun sofar k _ => match m₁.getEntry? k with | some kv' => sofar.insert kv'.1 kv'.2 | none => sofar) emptyWithCapacity).1 + +-- /-- Internal implementation detail of the hash map -/ +-- @[inline] def union [BEq α] [Hashable α] (m₁ m₂ : Raw₀ α β) : Raw₀ α β := +-- if m₁.1.size ≤ m₂.1.size then (m₂.insertManyIfNew m₁.1).1 else (m₁.insertMany m₂.1).1 + +-- /-- Internal implementation detail of the hash map -/ +-- def inter [BEq α] [Hashable α] (m₁ m₂ : Raw₀ α β) : Raw₀ α β := +-- if m₁.1.size ≤ m₂.1.size then m₁.filter fun k _ => m₂.contains k else interSmaller m₁ m₂ + +/-- Internal implementation detail of the hash map -/ +def interSmallerₘ [BEq α] [Hashable α] (m₁ : Raw₀ α β) (l : List ((a : α) × β a)) : Raw₀ α β := + foldlₘ (fun sofar k _ => interSmallerFnₘ m₁ sofar k) emptyWithCapacity l + +theorem foldl_perm_cong [BEq α] {init₁ init₂ : List ((a : α) × β a)} {l : List ((a : α) × β a)} + {f : List ((a : α) × β a) → ((a : α) × β a) → List ((a : α) × β a)} (h₁ : Perm init₁ init₂) + (h₂ : ∀ h l₁ l₂, (w : DistinctKeys l₁) → Perm l₁ l₂ → Perm (f l₁ h) (f l₂ h) ∧ DistinctKeys (f l₁ h)) + (h₃ : DistinctKeys init₁) + : Perm (List.foldl f init₁ l) (List.foldl f init₂ l) := by + induction l generalizing init₁ init₂ + case nil => + simp only [foldl_nil, h₁] + case cons h t ih => + simp only [foldl_cons] + apply ih + . exact (h₂ h init₁ init₂ h₃ h₁).1 + . exact (h₂ h init₁ init₂ h₃ h₁).2 + +theorem toListModel_interSmallerFnₘ [BEq α] [EquivBEq α] [Hashable α] [LawfulHashable α] (m sofar : Raw₀ α β) (hm₁ : Raw.WFImp m.1) (hm₂ : Raw.WFImp sofar.1) (k : α) : + Perm (toListModel ((interSmallerFnₘ m sofar k).1.buckets)) (match List.getEntry? k (toListModel m.1.buckets) with + | some kv' => List.insertEntry kv'.1 kv'.2 (toListModel sofar.1.buckets) + | none => (toListModel sofar.1.buckets)) := by + rw [interSmallerFnₘ] + rw [←getEntry?_eq_getEntry?ₘ, getEntry?_eq_getEntry?] + . split + case h_1 _ kv' heq => + simp only [heq] + apply toListModel_insertₘ + . exact hm₂ + case h_2 _ heq => + simp [heq] + . exact hm₁ + +theorem toListModel_interSmallerₘ [BEq α] [EquivBEq α] [Hashable α] [LawfulHashable α] (m₁ m₂ : Raw₀ α β) (hm₁ : Raw.WFImp m₁.1) (hm₂ : Raw.WFImp m₂.1) (l : List ((a : α) × β a)) : + toListModel + (List.foldl + (fun sofar x => + match x with + | ⟨k, _⟩ => m₁.interSmallerFnₘ sofar k) + m₂ l).val.buckets ~ + List.foldl + (fun sofar x => + match x with + | ⟨k, _⟩ => + match List.getEntry? k (toListModel m₁.val.buckets) with + | some kv' => insertEntry kv'.fst kv'.snd sofar + | none => sofar) + (toListModel m₂.val.buckets) l := by + induction l generalizing m₂ + case nil => + simp + case cons kv t ih => + simp + apply Perm.trans + . apply ih + . unfold interSmallerFnₘ + split + case _ kv' heq => + rw [← insert_eq_insertₘ] + apply wfImp_insert hm₂ + . exact hm₂ + . rw [interSmallerFnₘ] + apply foldl_perm_cong + . split + case h_1 _ kv' heq => + rw [← getEntry?_eq_getEntry?ₘ, getEntry?_eq_getEntry?] at heq + simp only [heq] + apply toListModel_insertₘ + . exact hm₂ + . exact hm₁ + case h_2 _ heq => + rw [← getEntry?_eq_getEntry?ₘ, getEntry?_eq_getEntry?] at heq + simp only [heq] + apply Perm.refl + . exact hm₁ + . intro h l₁ l₂ hd hp + simp + apply And.intro + . split + case h_1 _ kv' heq => + apply insertEntry_of_perm + . exact hd + . exact hp + case h_2 _ heq => + exact hp + . split + case h_1 _ kv' heq => + apply DistinctKeys.insertEntry + . exact hd + case h_2 _ heq => + . exact hd + . split + case h_1 _ kv' heq => + apply DistinctKeys.perm + . apply toListModel_insertₘ + . exact hm₂ + . apply DistinctKeys.insertEntry + . exact hm₂.distinct + case h_2 _ => + exact hm₂.distinct + +theorem interSmaller_eq_interSmallerₘ [BEq α] [Hashable α] (m₁ m₂ : Raw₀ α β) : + interSmaller m₁ m₂ = interSmallerₘ m₁ (toListModel m₂.1.buckets) := by + rw [interSmaller, foldl_eq_foldlₘ] + rw [interSmallerₘ] + rw [Raw.toList_eq_toListModel] + simp only [interSmallerFn_eq_interSmallerFnₘ] + +theorem toListModel_inter [BEq α] [EquivBEq α] [Hashable α] [LawfulHashable α] (m₁ m₂ : Raw₀ α β) (hm₁ : Raw.WFImp m₁.1) (hm₂ : Raw.WFImp m₂.1) : + Perm (toListModel (m₁.inter m₂).1.buckets) ((toListModel m₁.1.buckets).filter fun p => containsKey p.1 (toListModel m₂.1.buckets) ) := by + simp [inter] + split + . rw [filter_eq_filterₘ] + apply Perm.trans + . apply toListModel_filterₘ + . suffices (fun p : (a : α) × β a => m₂.contains p.fst) = (fun p : (a : α) × β a => containsKey p.fst (toListModel m₂.val.buckets)) by rw [this] + ext x + rw [contains_eq_containsKey] + exact hm₂ + . apply Perm.trans + . rw [interSmaller_eq_interSmallerₘ] + . rw [interSmallerₘ, foldlₘ] + apply Perm.trans + . apply toListModel_interSmallerₘ + . exact hm₁ + . apply wfImp_emptyWithCapacity + . generalize heq1 : (toListModel m₁.val.buckets) = l₁ + generalize heq2 : (toListModel m₂.val.buckets) = l₂ + induction l₁ + case nil => + simp + induction l₂ + case nil => simp + case cons h t ih => + rename_i heq + rw [Raw.WFImp.size_eq, Raw.WFImp.size_eq, heq1, heq2] at heq + simp [length] at heq + . exact hm₂ + . exact hm₁ + case cons h t t_ih => + simp + rw [List.filter_cons] + sorry + + + + + + + + + + + + + + + /-! # `Const.insertListₘ` -/ theorem Const.toListModel_insertListₘ {β : Type v} [BEq α] [Hashable α] [EquivBEq α] From 2239c23f757d23209f41f629d16fe5e4a5e675f7 Mon Sep 17 00:00:00 2001 From: Wojciech Rozowski Date: Wed, 5 Nov 2025 14:54:12 +0100 Subject: [PATCH 08/88] Massive lemma is done --- src/Std/Data/DHashMap/Internal/WF.lean | 150 ++++++++++++++++++++----- 1 file changed, 124 insertions(+), 26 deletions(-) diff --git a/src/Std/Data/DHashMap/Internal/WF.lean b/src/Std/Data/DHashMap/Internal/WF.lean index db82f9b78829..12f51aa9dde2 100644 --- a/src/Std/Data/DHashMap/Internal/WF.lean +++ b/src/Std/Data/DHashMap/Internal/WF.lean @@ -1429,6 +1429,122 @@ theorem interSmaller_eq_interSmallerₘ [BEq α] [Hashable α] (m₁ m₂ : Raw rw [Raw.toList_eq_toListModel] simp only [interSmallerFn_eq_interSmallerFnₘ] +theorem getEntry_foldl [BEq α] [EquivBEq α] (a : α) {acc l₁ l₂ : List ((a : α) × β a)} (hd : DistinctKeys l₁) (hyp : ∀ (a : α) (kv''), List.getEntry? a l₁ = some kv'' → (List.getEntry? a acc = none) ∨ (List.getEntry? a acc = some kv'')): List.getEntry? a + (List.foldl + (fun sofar x => + match List.getEntry? x.fst l₁ with + | some kv' => insertEntry kv'.fst kv'.snd sofar + | none => sofar) + acc l₂) = List.getEntry? a (acc ++ List.filter (fun p => containsKey p.fst l₂) l₁) := by + induction l₂ generalizing acc + case nil => + simp + suffices (List.getEntry? a (List.filter (fun p => false) l₁)) = .none by + simp [this] + have : (List.filter (fun p => false) l₁) = [] := by simp + simp [this] + case cons h t ih => + simp only [foldl_cons, getEntry?_append] + simp at ih + specialize @ih ((match List.getEntry? h.fst l₁ with + | some kv' => insertEntry kv'.fst kv'.snd acc + | none => acc)) + specialize ih ?goal + . intro a kv hkv + split + case h_1 _ kv' heq => + by_cases a == kv'.fst + case pos isTrue => + apply Or.inr + simp [getEntry?_insertEntry] + simp [PartialEquivBEq.symm isTrue] + rw [@getEntry?_congr α β _ _ l₁ _ _ (PartialEquivBEq.trans (PartialEquivBEq.symm <| beq_of_getEntry?_eq_some heq) ( PartialEquivBEq.symm <| isTrue))] at heq + rw [heq] at hkv + injections + case neg isFalse => + specialize hyp a kv hkv + apply Or.elim hyp + . intro hyp2 + apply Or.inl + simp [containsKey_insertEntry] + simp at isFalse + simp [BEq.symm_false isFalse] + rw [containsKey_eq_isSome_getEntry?, hyp2] + simp + . intro isSome + apply Or.inr + simp at isFalse + simp [getEntry?_insertEntry, BEq.symm_false isFalse, isSome] + case h_2 _ heq => + apply Or.elim <| hyp a kv hkv + . rw [containsKey_eq_isSome_getEntry?] + intro hyp + simp [hyp] + . intro hyp + exact Or.inr hyp + simp [ih] + split + case h_1 _ kv' heq => + by_cases ha : a == kv'.fst + case pos => + simp [getEntry?_insertEntry, PartialEquivBEq.symm ha] + have heq2 := beq_of_getEntry?_eq_some heq + specialize hyp a kv' (by rw [getEntry?_congr ha, getEntry?_congr heq2]; exact heq) + apply Or.elim hyp + . intro isNone + simp [isNone] + rw [getEntry?_filter] + rw [getEntry?_congr ha, getEntry?_congr heq2, heq] + simp [Option.filter] + simp [containsKey, PartialEquivBEq.symm heq2] + . exact hd + . intro hyp + simp [hyp] + case neg => + simp at ha + simp [getEntry?_insertEntry, BEq.symm_false ha] + rw [getEntry?_filter, getEntry?_filter] + generalize heq3 : List.getEntry? a l₁ = x + cases x + case none => + simp + case some kv'' => + simp [containsKey, Option.filter] + split + case isTrue heq4 => + simp [heq4] + case isFalse heq4 => + simp [heq4, BEq.symm_false (BEq.neq_of_beq_of_neq (beq_of_getEntry?_eq_some heq3) (BEq.neq_of_neq_of_beq ha (beq_of_getEntry?_eq_some heq)))] + . exact hd + . exact hd + case h_2 _ heq => + congr 1 + rw [getEntry?_filter, getEntry?_filter] + generalize heq2 : List.getEntry? a l₁ = x + cases x + case e_a.none => + simp [Option.filter] + case e_a.some kv' => + simp [Option.filter, containsKey] + congr + simp + intro hyp + have heq3 := beq_of_getEntry?_eq_some heq2 + have heq4 : h.fst == a := by + apply PartialEquivBEq.trans + . exact hyp + . exact heq3 + rw [@getEntry?_congr α β _ _ l₁ h.fst a heq4] at heq + rw [heq] at heq2 + simp_all + . exact hd + . exact hd + + + + + + theorem toListModel_inter [BEq α] [EquivBEq α] [Hashable α] [LawfulHashable α] (m₁ m₂ : Raw₀ α β) (hm₁ : Raw.WFImp m₁.1) (hm₂ : Raw.WFImp m₂.1) : Perm (toListModel (m₁.inter m₂).1.buckets) ((toListModel m₁.1.buckets).filter fun p => containsKey p.1 (toListModel m₂.1.buckets) ) := by simp [inter] @@ -1447,34 +1563,16 @@ theorem toListModel_inter [BEq α] [EquivBEq α] [Hashable α] [LawfulHashable . apply toListModel_interSmallerₘ . exact hm₁ . apply wfImp_emptyWithCapacity - . generalize heq1 : (toListModel m₁.val.buckets) = l₁ - generalize heq2 : (toListModel m₂.val.buckets) = l₂ - induction l₁ - case nil => - simp - induction l₂ - case nil => simp - case cons h t ih => - rename_i heq - rw [Raw.WFImp.size_eq, Raw.WFImp.size_eq, heq1, heq2] at heq - simp [length] at heq - . exact hm₂ - . exact hm₁ - case cons h t t_ih => + . generalize (toListModel m₁.val.buckets) = l₁ + generalize (toListModel m₂.val.buckets) = l₂ + apply getEntry?_ext + . sorry + . sorry + . intro a simp - rw [List.filter_cons] + apply getEntry_foldl + sorry sorry - - - - - - - - - - - From 8b8f3bb44471a9f767f5d6f535f1d68e42884f48 Mon Sep 17 00:00:00 2001 From: Wojciech Rozowski Date: Wed, 5 Nov 2025 15:10:13 +0100 Subject: [PATCH 09/88] refactor --- src/Std/Data/DHashMap/Internal/WF.lean | 105 ++++++++++++++++--------- 1 file changed, 68 insertions(+), 37 deletions(-) diff --git a/src/Std/Data/DHashMap/Internal/WF.lean b/src/Std/Data/DHashMap/Internal/WF.lean index 12f51aa9dde2..694e14ca6483 100644 --- a/src/Std/Data/DHashMap/Internal/WF.lean +++ b/src/Std/Data/DHashMap/Internal/WF.lean @@ -1439,10 +1439,10 @@ theorem getEntry_foldl [BEq α] [EquivBEq α] (a : α) {acc l₁ l₂ : List ((a induction l₂ generalizing acc case nil => simp - suffices (List.getEntry? a (List.filter (fun p => false) l₁)) = .none by - simp [this] - have : (List.filter (fun p => false) l₁) = [] := by simp - simp [this] + suffices (List.getEntry? a (List.filter (fun p => false) l₁)) = .none by simp [this] + have : (List.filter (fun p => false) l₁) = [] := by simp only [filter_eq_nil_iff, + Bool.false_eq_true, not_false_eq_true, implies_true] + simp only [getEntry?_eq_none, this, containsKey_nil] case cons h t ih => simp only [foldl_cons, getEntry?_append] simp at ih @@ -1456,8 +1456,7 @@ theorem getEntry_foldl [BEq α] [EquivBEq α] (a : α) {acc l₁ l₂ : List ((a by_cases a == kv'.fst case pos isTrue => apply Or.inr - simp [getEntry?_insertEntry] - simp [PartialEquivBEq.symm isTrue] + simp [getEntry?_insertEntry, PartialEquivBEq.symm isTrue] rw [@getEntry?_congr α β _ _ l₁ _ _ (PartialEquivBEq.trans (PartialEquivBEq.symm <| beq_of_getEntry?_eq_some heq) ( PartialEquivBEq.symm <| isTrue))] at heq rw [heq] at hkv injections @@ -1466,37 +1465,37 @@ theorem getEntry_foldl [BEq α] [EquivBEq α] (a : α) {acc l₁ l₂ : List ((a apply Or.elim hyp . intro hyp2 apply Or.inl - simp [containsKey_insertEntry] - simp at isFalse + simp only [containsKey_insertEntry, Bool.or_eq_false_iff] + simp only [Bool.not_eq_true] at isFalse simp [BEq.symm_false isFalse] rw [containsKey_eq_isSome_getEntry?, hyp2] - simp + simp only [Option.isSome_none] . intro isSome apply Or.inr - simp at isFalse - simp [getEntry?_insertEntry, BEq.symm_false isFalse, isSome] + simp only [Bool.not_eq_true] at isFalse + simp only [getEntry?_insertEntry, BEq.symm_false isFalse, Bool.false_eq_true, + ↓reduceIte, isSome] case h_2 _ heq => apply Or.elim <| hyp a kv hkv . rw [containsKey_eq_isSome_getEntry?] intro hyp - simp [hyp] + simp only [hyp, Option.isSome_none, reduceCtorEq, or_false] . intro hyp exact Or.inr hyp - simp [ih] + rw [ih] split case h_1 _ kv' heq => by_cases ha : a == kv'.fst case pos => - simp [getEntry?_insertEntry, PartialEquivBEq.symm ha] + simp only [getEntry?_insertEntry, PartialEquivBEq.symm ha, ↓reduceIte, Option.some_or] have heq2 := beq_of_getEntry?_eq_some heq specialize hyp a kv' (by rw [getEntry?_congr ha, getEntry?_congr heq2]; exact heq) apply Or.elim hyp . intro isNone simp [isNone] - rw [getEntry?_filter] - rw [getEntry?_congr ha, getEntry?_congr heq2, heq] - simp [Option.filter] - simp [containsKey, PartialEquivBEq.symm heq2] + rw [getEntry?_filter, getEntry?_congr ha, getEntry?_congr heq2, heq] + simp only [Option.filter, containsKey, PartialEquivBEq.symm heq2, Bool.true_or, + ↓reduceIte] . exact hd . intro hyp simp [hyp] @@ -1509,12 +1508,16 @@ theorem getEntry_foldl [BEq α] [EquivBEq α] (a : α) {acc l₁ l₂ : List ((a case none => simp case some kv'' => - simp [containsKey, Option.filter] + simp only [Option.filter, containsKey, Bool.or_eq_true] split case isTrue heq4 => - simp [heq4] + simp only [Option.or_some, heq4, or_true, ↓reduceIte] case isFalse heq4 => - simp [heq4, BEq.symm_false (BEq.neq_of_beq_of_neq (beq_of_getEntry?_eq_some heq3) (BEq.neq_of_neq_of_beq ha (beq_of_getEntry?_eq_some heq)))] + simp only [Option.or_none, + BEq.symm_false + (BEq.neq_of_beq_of_neq (beq_of_getEntry?_eq_some heq3) + (BEq.neq_of_neq_of_beq ha (beq_of_getEntry?_eq_some heq))), + Bool.false_eq_true, heq4, or_self, ↓reduceIte] . exact hd . exact hd case h_2 _ heq => @@ -1525,9 +1528,9 @@ theorem getEntry_foldl [BEq α] [EquivBEq α] (a : α) {acc l₁ l₂ : List ((a case e_a.none => simp [Option.filter] case e_a.some kv' => - simp [Option.filter, containsKey] + simp only [Option.filter, containsKey, Bool.or_eq_true] congr - simp + simp only [eq_iff_iff, iff_or_self] intro hyp have heq3 := beq_of_getEntry?_eq_some heq2 have heq4 : h.fst == a := by @@ -1536,14 +1539,38 @@ theorem getEntry_foldl [BEq α] [EquivBEq α] (a : α) {acc l₁ l₂ : List ((a . exact heq3 rw [@getEntry?_congr α β _ _ l₁ h.fst a heq4] at heq rw [heq] at heq2 - simp_all + simp_all only [reduceCtorEq] . exact hd . exact hd - - - - +theorem foldl_insertEntry_distinctKeys [BEq α] [PartialEquivBEq α] (l₁ l₂ : List ((a : α) × β a)) (acc : List ((a : α) × β a)) + (h_acc : DistinctKeys acc) : + DistinctKeys + (List.foldl + (fun sofar x => + match x with + | ⟨k, _⟩ => + match List.getEntry? k l₁ with + | some kv' => insertEntry kv'.fst kv'.snd sofar + | none => sofar) + acc l₂) := by + induction l₂ generalizing acc with + | nil => + simp [List.foldl] + exact h_acc + | cons h t ih => + simp only [List.foldl_cons] + cases heq : List.getEntry? h.fst l₁ with + | none => + apply ih + exact h_acc + | some kv' => + apply ih + split + case h_1 _ kv' heq => + apply DistinctKeys.insertEntry + exact h_acc + case h_2 => exact h_acc theorem toListModel_inter [BEq α] [EquivBEq α] [Hashable α] [LawfulHashable α] (m₁ m₂ : Raw₀ α β) (hm₁ : Raw.WFImp m₁.1) (hm₂ : Raw.WFImp m₂.1) : Perm (toListModel (m₁.inter m₂).1.buckets) ((toListModel m₁.1.buckets).filter fun p => containsKey p.1 (toListModel m₂.1.buckets) ) := by @@ -1563,19 +1590,23 @@ theorem toListModel_inter [BEq α] [EquivBEq α] [Hashable α] [LawfulHashable . apply toListModel_interSmallerₘ . exact hm₁ . apply wfImp_emptyWithCapacity - . generalize (toListModel m₁.val.buckets) = l₁ - generalize (toListModel m₂.val.buckets) = l₂ + . generalize heq1 : (toListModel m₁.val.buckets) = l₁ + generalize heq2 : (toListModel m₂.val.buckets) = l₂ apply getEntry?_ext - . sorry - . sorry + . apply foldl_insertEntry_distinctKeys + rw [toListModel_buckets_emptyWithCapacity] + simp + . apply @DistinctKeys.filter _ _ _ l₁ (fun k v => containsKey k l₂) + rw [← heq1] + exact hm₁.distinct . intro a simp apply getEntry_foldl - sorry - sorry - - - + . rw [← heq1] + exact hm₁.distinct + . intro a _ _ + apply Or.inl + simp /-! # `Const.insertListₘ` -/ From 02226fc1b8a34608bb9f4e988cd386ca1b109dc4 Mon Sep 17 00:00:00 2001 From: Wojciech Rozowski Date: Wed, 5 Nov 2025 16:08:50 +0100 Subject: [PATCH 10/88] WF lemma is now working --- src/Std/Data/DHashMap/Internal/WF.lean | 22 +++++++-------- src/Std/Data/DHashMap/Raw.lean | 37 ++++++++++++++++++++++++++ 2 files changed, 47 insertions(+), 12 deletions(-) diff --git a/src/Std/Data/DHashMap/Internal/WF.lean b/src/Std/Data/DHashMap/Internal/WF.lean index 694e14ca6483..68578f235d05 100644 --- a/src/Std/Data/DHashMap/Internal/WF.lean +++ b/src/Std/Data/DHashMap/Internal/WF.lean @@ -1309,18 +1309,6 @@ theorem foldl_eq_foldlₘ {δ} (f : δ → (a : α) → β a → δ) (init : δ) simp only [foldlₘ, Subtype.forall] at ih apply ih --- /-- Internal implementation detail of the hash map -/ --- def interSmaller [BEq α] [Hashable α] (m₁ m₂ : Raw₀ α β) : Raw₀ α β := --- (m₂.foldl (fun sofar k _ => match m₁.getEntry? k with | some kv' => sofar.insert kv'.1 kv'.2 | none => sofar) emptyWithCapacity).1 - --- /-- Internal implementation detail of the hash map -/ --- @[inline] def union [BEq α] [Hashable α] (m₁ m₂ : Raw₀ α β) : Raw₀ α β := --- if m₁.1.size ≤ m₂.1.size then (m₂.insertManyIfNew m₁.1).1 else (m₁.insertMany m₂.1).1 - --- /-- Internal implementation detail of the hash map -/ --- def inter [BEq α] [Hashable α] (m₁ m₂ : Raw₀ α β) : Raw₀ α β := --- if m₁.1.size ≤ m₂.1.size then m₁.filter fun k _ => m₂.contains k else interSmaller m₁ m₂ - /-- Internal implementation detail of the hash map -/ def interSmallerₘ [BEq α] [Hashable α] (m₁ : Raw₀ α β) (l : List ((a : α) × β a)) : Raw₀ α β := foldlₘ (fun sofar k _ => interSmallerFnₘ m₁ sofar k) emptyWithCapacity l @@ -1608,6 +1596,16 @@ theorem toListModel_inter [BEq α] [EquivBEq α] [Hashable α] [LawfulHashable apply Or.inl simp +theorem wf_inter₀ [BEq α] [Hashable α] [LawfulHashable α] + {m₁ m₂ : Raw α β} {h₁ : 0 < m₁.buckets.size} {h₂ : 0 < m₂.buckets.size} (h'₁ : m₁.WF) + (h'₂ : m₂.WF) : + (Raw₀.inter ⟨m₁, h₁⟩ ⟨m₂, h₂⟩).1.WF := by + rw [inter] + split + . sorry + . sorry + + /-! # `Const.insertListₘ` -/ theorem Const.toListModel_insertListₘ {β : Type v} [BEq α] [Hashable α] [EquivBEq α] diff --git a/src/Std/Data/DHashMap/Raw.lean b/src/Std/Data/DHashMap/Raw.lean index b0018b935cdf..5166f9642065 100644 --- a/src/Std/Data/DHashMap/Raw.lean +++ b/src/Std/Data/DHashMap/Raw.lean @@ -470,6 +470,23 @@ This function always merges the smaller map into the larger map, so the expected instance [BEq α] [Hashable α] : Union (Raw α β) := ⟨union⟩ +/-- +Computes the intersection of the given hash maps. The result will only contain entries from the first map. + +This function always merges the smaller map into the larger map, so the expected runtime is +`O(min(m₁.size, m₂.size))`. +-/ +@[inline] def inter [BEq α] [Hashable α] (m₁ m₂ : Raw α β) : Raw α β := + if h₁ : 0 < m₁.buckets.size then + if h₂ : 0 < m₂.buckets.size then + Raw₀.inter ⟨m₁, h₁⟩ ⟨m₂, h₂⟩ + else + m₁ + else + m₂ + +instance [BEq α] [Hashable α] : Inter (Raw α β) := ⟨inter⟩ + section Unverified @@ -711,10 +728,30 @@ theorem WF.union₀ [BEq α] [Hashable α] {m₁ m₂ : Raw α β} (h₁ : m₁. . exact (Raw₀.insertManyIfNew ⟨m₂, h₂.size_buckets_pos⟩ m₁).2 _ WF.insertIfNew₀ h₂ . exact (Raw₀.insertMany ⟨m₁, h₁.size_buckets_pos⟩ m₂).2 _ WF.insert₀ h₁ +theorem WF.inter₀ [BEq α] [Hashable α] {m₁ m₂ : Raw α β} (h₁ : m₁.WF) (h₂ : m₂.WF) : (Raw₀.inter ⟨m₁, h₁.size_buckets_pos⟩ ⟨m₂, h₂.size_buckets_pos⟩).val.WF := by + simp only [Raw₀.inter] + split + . apply WF.filter₀ h₁ + . rw [Raw₀.interSmaller] + have := (@Raw₀.foldl α β _ (fun sofar k _ => Raw₀.interSmallerFn ⟨m₁, h₁.size_buckets_pos⟩ sofar k) Raw₀.emptyWithCapacity ⟨m₂, h₂.size_buckets_pos⟩).2 (fun x => x.val.WF) + apply this + . intro ⟨d, hd⟩ k v hw + rw [Raw₀.interSmallerFn] + split + . apply insert₀ + simp [hw] + . exact hw + . apply emptyWithCapacity₀ + + theorem WF.union [BEq α] [Hashable α] {m₁ m₂ : Raw α β} (h₁ : m₁.WF) (h₂ : m₂.WF) : (m₁.union m₂ : Raw α β).WF := by simp [Std.DHashMap.Raw.union, h₁.size_buckets_pos, h₂.size_buckets_pos] exact WF.union₀ h₁ h₂ +theorem WF.inter [BEq α] [Hashable α] {m₁ m₂ : Raw α β} (h₁ : m₁.WF) (h₂ : m₂.WF) : (m₁.inter m₂ : Raw α β).WF := by + simp [Std.DHashMap.Raw.inter, h₁.size_buckets_pos, h₂.size_buckets_pos] + exact WF.inter₀ h₁ h₂ + end WF end Raw From 69669c691139a265c1b094277bcc1c6594d30177 Mon Sep 17 00:00:00 2001 From: Wojciech Rozowski Date: Wed, 5 Nov 2025 16:10:54 +0100 Subject: [PATCH 11/88] chore : the wf variant works --- src/Std/Data/DHashMap/Basic.lean | 8 +++++++- src/Std/Data/DHashMap/Raw.lean | 1 - 2 files changed, 7 insertions(+), 2 deletions(-) diff --git a/src/Std/Data/DHashMap/Basic.lean b/src/Std/Data/DHashMap/Basic.lean index 952fab9ffb13..88dc4ac067ee 100644 --- a/src/Std/Data/DHashMap/Basic.lean +++ b/src/Std/Data/DHashMap/Basic.lean @@ -319,9 +319,15 @@ This function always merges the smaller map into the larger map, so the expected inner := Raw₀.union ⟨m₁.1, m₁.2.size_buckets_pos⟩ ⟨m₂.1, m₂.2.size_buckets_pos⟩ wf := Std.DHashMap.Raw.WF.union₀ m₁.2 m₂.2 +/-- +Computes the intersection of the given hash maps. The result will only contain entries from the first map. + +This function always merges the smaller map into the larger map, so the expected runtime is +`O(min(m₁.size, m₂.size))`. +-/ @[inline] def inter [BEq α] [Hashable α] (m₁ m₂ : DHashMap α β) : DHashMap α β where inner := Raw₀.inter ⟨m₁.1, m₁.2.size_buckets_pos⟩ ⟨m₂.1, m₂.2.size_buckets_pos⟩ - wf := sorry + wf := Std.DHashMap.Raw.WF.inter₀ m₁.2 m₂.2 instance [BEq α] [Hashable α] : Union (DHashMap α β) := ⟨union⟩ diff --git a/src/Std/Data/DHashMap/Raw.lean b/src/Std/Data/DHashMap/Raw.lean index 5166f9642065..e3668f8da573 100644 --- a/src/Std/Data/DHashMap/Raw.lean +++ b/src/Std/Data/DHashMap/Raw.lean @@ -743,7 +743,6 @@ theorem WF.inter₀ [BEq α] [Hashable α] {m₁ m₂ : Raw α β} (h₁ : m₁. . exact hw . apply emptyWithCapacity₀ - theorem WF.union [BEq α] [Hashable α] {m₁ m₂ : Raw α β} (h₁ : m₁.WF) (h₂ : m₂.WF) : (m₁.union m₂ : Raw α β).WF := by simp [Std.DHashMap.Raw.union, h₁.size_buckets_pos, h₂.size_buckets_pos] exact WF.union₀ h₁ h₂ From 439e24212570598c3d45be772a1971dbf30733cd Mon Sep 17 00:00:00 2001 From: Wojciech Rozowski Date: Wed, 5 Nov 2025 16:22:59 +0100 Subject: [PATCH 12/88] chore add lemmas --- src/Std/Data/DHashMap/Internal/RawLemmas.lean | 176 +++++++++++++++++- 1 file changed, 175 insertions(+), 1 deletion(-) diff --git a/src/Std/Data/DHashMap/Internal/RawLemmas.lean b/src/Std/Data/DHashMap/Internal/RawLemmas.lean index b4ce6c6dbb1d..fbd1b1242cff 100644 --- a/src/Std/Data/DHashMap/Internal/RawLemmas.lean +++ b/src/Std/Data/DHashMap/Internal/RawLemmas.lean @@ -91,7 +91,7 @@ macro_rules | apply Raw₀.wf_insertMany₀ | apply Raw₀.Const.wf_insertMany₀ | apply Raw₀.Const.wf_insertManyIfNewUnit₀ | apply Raw₀.wf_union₀ | apply Raw.WF.filter₀ | apply Raw₀.wf_map₀ | apply Raw₀.wf_filterMap₀ - | apply Raw.WF.emptyWithCapacity₀) <;> wf_trivial) + | apply Raw.WF.emptyWithCapacity₀ | apply Raw₀.wf_inter₀) <;> wf_trivial) /-- Internal implementation detail of the hash map -/ scoped macro "empty" : tactic => `(tactic| { intros; simp_all [List.isEmpty_iff] } ) @@ -105,6 +105,7 @@ private meta def modifyMap : Std.DHashMap Name (fun _ => Name) := ⟨`insertIfNew, ``toListModel_insertIfNew⟩, ⟨`insertMany, ``toListModel_insertMany_list⟩, ⟨`union, ``toListModel_union⟩, + ⟨`inter, ``toListModel_inter⟩, ⟨`Const.insertMany, ``Const.toListModel_insertMany_list⟩, ⟨`Const.insertManyIfNewUnit, ``Const.toListModel_insertManyIfNewUnit_list⟩, ⟨`alter, ``toListModel_alter⟩, @@ -2775,6 +2776,179 @@ theorem get!_union_of_contains_eq_false_right [EquivBEq α] [LawfulHashable α] end Const +section Inter + +variable (m₁ m₂ : Raw₀ α β) +variable {m₁ m₂} + +/- contains -/ +theorem contains_inter [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) + (h₂ : m₂.val.WF) {k : α} : + (m₁.inter m₂).contains k = (m₁.contains k && m₂.contains k) := by + simp_to_model [contains, inter] + sorry + +theorem contains_inter_iff [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) + (h₂ : m₂.val.WF) {k : α} : + (m₁.inter m₂).contains k ↔ m₁.contains k ∧ m₂.contains k := by + simp_to_model [inter, contains] + sorry -- List.contains_filter_iff + +theorem not_contains_inter_of_not_contains_left [EquivBEq α] [LawfulHashable α] + (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} + (h : m₁.contains k = false) : + (m₁.inter m₂).contains k = false := by + simp_to_model [inter, contains] -- using List.not_contains_filter_of_not_contains + sorry + +theorem not_contains_inter_of_not_contains_right [EquivBEq α] [LawfulHashable α] + (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} + (h : m₂.contains k = false) : + (m₁.inter m₂).contains k = false := by + simp_to_model [inter, contains] -- using List.not_contains_filter_of_not_containsKey_pred + sorry + +/- get? -/ +theorem get?_inter [LawfulBEq α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} : + (m₁.inter m₂).get? k = + if m₂.contains k then m₁.get? k else none := by + simp_to_model [inter, get?, contains] -- using List.getValueCast?_filter + sorry + +-- Alternative formulation -- using Option.filter +theorem get?_inter' [LawfulBEq α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} : + (m₁.inter m₂).get? k = + (m₁.get? k).filter (fun _ => m₂.contains k) := by + simp_to_model [inter, get?, contains] -- using List.getValueCast?_filter_option + sorry + +theorem get?_inter_of_contains_right [LawfulBEq α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) + {k : α} (h : m₂.contains k) : + (m₁.inter m₂).get? k = m₁.get? k := by + simp_to_model [inter, get?, contains] -- using List.getValueCast?_filter_of_pred_true + sorry + +theorem get?_inter_of_not_contains_left [LawfulBEq α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) + {k : α} (h : m₁.contains k = false) : + (m₁.inter m₂).get? k = none := by + simp_to_model [inter, get?, contains] -- using List.getValueCast?_filter_of_not_contains + sorry + +theorem get?_inter_of_not_contains_right [LawfulBEq α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) + {k : α} (h : m₂.contains k = false) : + (m₁.inter m₂).get? k = none := by + simp_to_model [inter, get?, contains] -- using List.getValueCast?_filter_of_pred_false + sorry + +/- get -/ +theorem get_inter [LawfulBEq α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) + {k : α} {h_contains : (m₁.inter m₂).contains k} : + (m₁.inter m₂).get k h_contains = + m₁.get k (by simp [contains_inter_iff h₁ h₂] at h_contains; exact h_contains.1) := by + simp_to_model [inter, get, contains] -- using List.getValueCast_filter + sorry + +/- getD -/ +theorem getD_inter [LawfulBEq α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) + {k : α} {fallback : β k} : + (m₁.inter m₂).getD k fallback = + if m₂.contains k then m₁.getD k fallback else fallback := by + simp_to_model [inter, getD, contains] -- using List.getValueCastD_filter + sorry + +theorem getD_inter_of_contains_right [LawfulBEq α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) + {k : α} {fallback : β k} (h : m₂.contains k) : + (m₁.inter m₂).getD k fallback = m₁.getD k fallback := by + simp_to_model [inter, getD, contains] -- using List.getValueCastD_filter_of_pred_true + sorry + +theorem getD_inter_of_not_contains_right [LawfulBEq α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) + {k : α} {fallback : β k} (h : m₂.contains k = false) : + (m₁.inter m₂).getD k fallback = fallback := by + simp_to_model [inter, getD, contains] -- using List.getValueCastD_filter_of_pred_false + sorry + +/- get! -/ +theorem get!_inter [LawfulBEq α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) + {k : α} [Inhabited (β k)] : + (m₁.inter m₂).get! k = + if m₂.contains k then m₁.get! k else default := by + simp_to_model [inter, get!, contains] -- using List.getValueCastD_filter + sorry + +/- getKey? -/ +theorem getKey?_inter [EquivBEq α] [LawfulHashable α] + (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} : + (m₁.inter m₂).getKey? k = + if m₂.contains k then m₁.getKey? k else none := by + simp_to_model [inter, contains, getKey?] -- using List.getKey?_filter + sorry + +theorem getKey?_inter_of_contains_right [EquivBEq α] [LawfulHashable α] + (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} (h : m₂.contains k) : + (m₁.inter m₂).getKey? k = m₁.getKey? k := by + simp_to_model [contains, getKey?, inter] -- using List.getKey?_filter_of_pred_true + sorry + +theorem getKey?_inter_of_not_contains_right [EquivBEq α] [LawfulHashable α] + (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} (h : m₂.contains k = false) : + (m₁.inter m₂).getKey? k = none := by + simp_to_model [contains, getKey?, inter] -- using List.getKey?_filter_of_pred_false + sorry + +/- getKey -/ +theorem getKey_inter [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) + {k : α} {h_contains : (m₁.inter m₂).contains k} : + (m₁.inter m₂).getKey k h_contains = + m₁.getKey k (by simp [contains_inter_iff h₁ h₂] at h_contains; exact h_contains.1) := by + simp_to_model [inter, contains, getKey] -- using List.getKey_filter + sorry + +/- getKeyD -/ +theorem getKeyD_inter [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) + (h₂ : m₂.val.WF) {k fallback : α} : + (m₁.inter m₂).getKeyD k fallback = + if m₂.contains k then m₁.getKeyD k fallback else fallback := by + simp_to_model [inter, getKeyD, contains] -- using List.getKeyD_filter + sorry + +/- getKey! -/ +theorem getKey!_inter [EquivBEq α] [LawfulHashable α] [Inhabited α] + (h₁ : m₁.1.WF) (h₂ : m₂.1.WF) {k : α} : + (m₁.inter m₂).getKey! k = + if m₂.contains k then m₁.getKey! k else default := by + simp_to_model [inter, getKey!, contains] -- using List.getKeyD_filter + sorry + +/- size -/ +theorem size_inter_le_size_left [EquivBEq α] [LawfulHashable α] + (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) : + (m₁.inter m₂).1.size ≤ m₁.1.size := by + simp_to_model [inter, size] -- using List.length_filter_le + sorry + +theorem size_inter_le_size_right [EquivBEq α] [LawfulHashable α] + (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) : + (m₁.inter m₂).1.size ≤ m₂.1.size := by + simp_to_model [inter, size, contains] -- using List.length_filter_containsKey_le + sorry + +theorem size_inter_eq_size_left_of_subset [EquivBEq α] [LawfulHashable α] + (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) + (h : ∀ (a : α), m₁.contains a → m₂.contains a) : + (m₁.inter m₂).1.size = m₁.1.size := by + simp_to_model [inter, size, contains] -- using List.length_filter_of_forall_pred + sorry + +/- isEmpty -/ +@[simp] +theorem isEmpty_inter [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) : + (m₁.inter m₂).1.isEmpty = (m₁.1.isEmpty || m₂.1.isEmpty) := by + simp_to_model [isEmpty, inter, contains] -- using List.isEmpty_filter_containsKey + sorry + +end Inter + section Alter theorem isEmpty_alter_eq_isEmpty_erase [LawfulBEq α] (h : m.1.WF) {k : α} From a7c71fab615c406d784eefbe7da4ecf506aa3190 Mon Sep 17 00:00:00 2001 From: Wojciech Rozowski Date: Wed, 5 Nov 2025 18:13:37 +0100 Subject: [PATCH 13/88] more lemmas --- src/Std/Data/DHashMap/Internal/RawLemmas.lean | 27 +++----- src/Std/Data/Internal/List/Associative.lean | 69 +++++++++++++++++++ 2 files changed, 77 insertions(+), 19 deletions(-) diff --git a/src/Std/Data/DHashMap/Internal/RawLemmas.lean b/src/Std/Data/DHashMap/Internal/RawLemmas.lean index fbd1b1242cff..2ea4f4062ff1 100644 --- a/src/Std/Data/DHashMap/Internal/RawLemmas.lean +++ b/src/Std/Data/DHashMap/Internal/RawLemmas.lean @@ -2785,42 +2785,32 @@ variable {m₁ m₂} theorem contains_inter [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} : (m₁.inter m₂).contains k = (m₁.contains k && m₂.contains k) := by - simp_to_model [contains, inter] - sorry + simp_to_model [contains, inter] using List.containsKey_filter_containsKey theorem contains_inter_iff [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} : (m₁.inter m₂).contains k ↔ m₁.contains k ∧ m₂.contains k := by - simp_to_model [inter, contains] - sorry -- List.contains_filter_iff + simp_to_model [inter, contains] using List.contains_filter_iff theorem not_contains_inter_of_not_contains_left [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} (h : m₁.contains k = false) : (m₁.inter m₂).contains k = false := by - simp_to_model [inter, contains] -- using List.not_contains_filter_of_not_contains - sorry + revert h + simp_to_model [inter, contains] using List.not_contains_filter_of_not_contains theorem not_contains_inter_of_not_contains_right [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} (h : m₂.contains k = false) : (m₁.inter m₂).contains k = false := by - simp_to_model [inter, contains] -- using List.not_contains_filter_of_not_containsKey_pred - sorry + revert h + simp_to_model [inter, contains] using List.not_contains_filter_of_not_containsKey_pred /- get? -/ theorem get?_inter [LawfulBEq α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} : (m₁.inter m₂).get? k = if m₂.contains k then m₁.get? k else none := by - simp_to_model [inter, get?, contains] -- using List.getValueCast?_filter - sorry - --- Alternative formulation -- using Option.filter -theorem get?_inter' [LawfulBEq α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} : - (m₁.inter m₂).get? k = - (m₁.get? k).filter (fun _ => m₂.contains k) := by - simp_to_model [inter, get?, contains] -- using List.getValueCast?_filter_option - sorry + simp_to_model [inter, get?, contains] using List.getValueCast?_filter_of_containsKey theorem get?_inter_of_contains_right [LawfulBEq α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} (h : m₂.contains k) : @@ -2924,8 +2914,7 @@ theorem getKey!_inter [EquivBEq α] [LawfulHashable α] [Inhabited α] theorem size_inter_le_size_left [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) : (m₁.inter m₂).1.size ≤ m₁.1.size := by - simp_to_model [inter, size] -- using List.length_filter_le - sorry + simp_to_model [inter, size] using List.length_filter_le theorem size_inter_le_size_right [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) : diff --git a/src/Std/Data/Internal/List/Associative.lean b/src/Std/Data/Internal/List/Associative.lean index 28feb9d97e15..b98c6d091e59 100644 --- a/src/Std/Data/Internal/List/Associative.lean +++ b/src/Std/Data/Internal/List/Associative.lean @@ -5128,6 +5128,38 @@ theorem getEntry?_filter [BEq α] [EquivBEq α] intro p simp only [Option.all_guard, BEq.rfl, Bool.or_true] +theorem getValueCast?_filter_of_containsKey [BEq α] [LawfulBEq α] + {l₁ l₂ : List ((a : α) × β a)} {k : α} + (dl₁ : DistinctKeys l₁) : + getValueCast? k (List.filter (fun p => containsKey p.fst l₂) l₁) = + if containsKey k l₂ = true then getValueCast? k l₁ else none := by + simp only [getValueCast?_eq_getEntry?] + split + case isTrue h => + apply Option.dmap_congr + . simp + . rw [getEntry?_filter] + . simp [Option.filter] + split + case h_1 _ val heq => + rw [@containsKey_congr _ _ _ _ l₂ _ _ (PartialEquivBEq.symm <| @beq_of_getEntry?_eq_some α β _ l₁ k val heq)] at h + simp [h, heq] + case h_2 _ heq => + simp [heq] + . exact dl₁ + case isFalse h => + suffices (getEntry? k (List.filter (fun p => containsKey p.fst l₂) l₁)) = none by simp [this] + rw [getEntry?_filter] + simp [Option.filter] + split + case h_1 _ val heq => + have := PartialEquivBEq.symm <| @beq_of_getEntry?_eq_some α β _ l₁ k val heq + simp at h + rw [@containsKey_congr _ _ _ _ l₂ _ _ (PartialEquivBEq.symm <| @beq_of_getEntry?_eq_some α β _ l₁ k val heq)] at h + simp [h] + case h_2 => rfl + exact dl₁ + theorem getEntry?_filterMap [BEq α] [EquivBEq α] {f : (a : α) → β a → Option (γ a)} {l : List ((a : α) × β a)} {k : α} (hl : DistinctKeys l) : @@ -5146,6 +5178,42 @@ theorem getEntry?_map [BEq α] [EquivBEq α] intro p exact BEq.rfl +theorem containsKey_filter_containsKey [BEq α] [EquivBEq α] {l₁ l₂ : List ((a : α) × β a)} {hl₁ : DistinctKeys l₁} {k : α} : + (containsKey k (List.filter (fun p => containsKey p.fst (l₂)) l₁)) = + (containsKey k l₁ && containsKey k l₂) := by + rw [containsKey_eq_isSome_getEntry?, containsKey_eq_isSome_getEntry?] + rw [getEntry?_filter] + generalize heq : (getEntry? k l₁) = x + cases x + case none => + simp + case some kv => + simp only [Option.isSome_filter, Option.any_some, Option.isSome_some, Bool.true_and] + rw [containsKey_congr] + apply List.beq_of_getEntry?_eq_some heq + . exact hl₁ + +theorem contains_filter_iff [BEq α] [EquivBEq α] {l₁ l₂ : List ((a : α) × β a)} {hl₁ : DistinctKeys l₁} {k : α} : + (containsKey k (List.filter (fun p => containsKey p.fst (l₂)) l₁)) ↔ + (containsKey k l₁ ∧ containsKey k l₂) := by + rw [containsKey_filter_containsKey] + . simp + . exact hl₁ + +theorem not_contains_filter_of_not_contains [BEq α] [EquivBEq α] {l₁ l₂ : List ((a : α) × β a)} {hl₁ : DistinctKeys l₁} {k : α} : + containsKey k l₁ = false → containsKey k (List.filter (fun p => containsKey p.fst l₂) l₁) = false := by + intro h + rw [containsKey_filter_containsKey] + . simp [h] + . exact hl₁ + +theorem not_contains_filter_of_not_containsKey_pred [BEq α] [EquivBEq α] {l₁ l₂ : List ((a : α) × β a)} {hl₁ : DistinctKeys l₁} {k : α} : + containsKey k l₂ = false → containsKey k (List.filter (fun p => containsKey p.fst l₂) l₁) = false := by + intro h + rw [containsKey_filter_containsKey] + . simp [h] + . exact hl₁ + theorem containsKey_of_containsKey_filterMap' [BEq α] [EquivBEq α] {f : ((a : α) × β a) → Option ((a : α) × γ a)} (hf : ∀ p, (f p).all (·.1 == p.1)) @@ -5740,6 +5808,7 @@ theorem isEmpty_filter_key_iff [BEq α] [EquivBEq α] {f : α → Bool} simp only [getKey, getKey?_eq_getEntry?, this] at h exact h + namespace Const theorem getKey_getValue_mem [BEq α] [EquivBEq α] {β : Type v} {l : List ((_ : α) × β)} {k : α} {h} : From 725689c33461231b17a4c24b8ccb019373879518 Mon Sep 17 00:00:00 2001 From: Wojciech Rozowski Date: Thu, 6 Nov 2025 14:29:01 +0100 Subject: [PATCH 14/88] all get lemmas are done --- src/Std/Data/DHashMap/Internal/RawLemmas.lean | 57 ++++++++---- src/Std/Data/Internal/List/Associative.lean | 93 ++++++++++++++++++- 2 files changed, 129 insertions(+), 21 deletions(-) diff --git a/src/Std/Data/DHashMap/Internal/RawLemmas.lean b/src/Std/Data/DHashMap/Internal/RawLemmas.lean index 2ea4f4062ff1..4db16b4ca8bc 100644 --- a/src/Std/Data/DHashMap/Internal/RawLemmas.lean +++ b/src/Std/Data/DHashMap/Internal/RawLemmas.lean @@ -2810,61 +2810,82 @@ theorem not_contains_inter_of_not_contains_right [EquivBEq α] [LawfulHashable theorem get?_inter [LawfulBEq α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} : (m₁.inter m₂).get? k = if m₂.contains k then m₁.get? k else none := by - simp_to_model [inter, get?, contains] using List.getValueCast?_filter_of_containsKey + simp_to_model [inter, get?, contains] using List.getValueCast?_filter_containsKey_eq_if_containsKey theorem get?_inter_of_contains_right [LawfulBEq α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} (h : m₂.contains k) : (m₁.inter m₂).get? k = m₁.get? k := by - simp_to_model [inter, get?, contains] -- using List.getValueCast?_filter_of_pred_true - sorry + revert h + simp_to_model [inter, get?, contains] using List.getValueCast?_filter_containsKey_eq_containsKey theorem get?_inter_of_not_contains_left [LawfulBEq α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} (h : m₁.contains k = false) : (m₁.inter m₂).get? k = none := by - simp_to_model [inter, get?, contains] -- using List.getValueCast?_filter_of_not_contains - sorry + revert h + simp_to_model [inter, get?, contains] using List.getValueCast?_filter_of_not_contains theorem get?_inter_of_not_contains_right [LawfulBEq α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} (h : m₂.contains k = false) : (m₁.inter m₂).get? k = none := by - simp_to_model [inter, get?, contains] -- using List.getValueCast?_filter_of_pred_false - sorry + revert h + simp_to_model [inter, get?, contains] using List.getValueCast?_filter_of_not_contains_right /- get -/ theorem get_inter [LawfulBEq α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} {h_contains : (m₁.inter m₂).contains k} : (m₁.inter m₂).get k h_contains = - m₁.get k (by simp [contains_inter_iff h₁ h₂] at h_contains; exact h_contains.1) := by - simp_to_model [inter, get, contains] -- using List.getValueCast_filter - sorry + m₁.get k (((@contains_inter_iff α β _ _ m₁ m₂ _ _ h₁ h₂ k).1 h_contains).1) := by + simp_to_model [inter, get, contains] using List.getValueCast_filter_containsKey_eq_containsKey /- getD -/ theorem getD_inter [LawfulBEq α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} {fallback : β k} : (m₁.inter m₂).getD k fallback = if m₂.contains k then m₁.getD k fallback else fallback := by - simp_to_model [inter, getD, contains] -- using List.getValueCastD_filter - sorry + simp_to_model [inter, getD, contains] using List.getValueCastD_filter_containsKey_eq_if_containsKey theorem getD_inter_of_contains_right [LawfulBEq α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} {fallback : β k} (h : m₂.contains k) : (m₁.inter m₂).getD k fallback = m₁.getD k fallback := by - simp_to_model [inter, getD, contains] -- using List.getValueCastD_filter_of_pred_true - sorry + revert h + simp_to_model [inter, getD, contains] using List.getValueCastD_filter_containsKey_eq_containsKey theorem getD_inter_of_not_contains_right [LawfulBEq α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} {fallback : β k} (h : m₂.contains k = false) : (m₁.inter m₂).getD k fallback = fallback := by - simp_to_model [inter, getD, contains] -- using List.getValueCastD_filter_of_pred_false - sorry + revert h + simp_to_model [inter, getD, contains] using List.getValueCastD_filter_of_not_contains_right + +theorem getD_inter_of_not_contains_left [LawfulBEq α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) + {k : α} {fallback : β k} (h : m₁.contains k = false) : + (m₁.inter m₂).getD k fallback = fallback := by + revert h + simp_to_model [inter, getD, contains] using List.getValueCastD_filter_of_not_contains_left /- get! -/ theorem get!_inter [LawfulBEq α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} [Inhabited (β k)] : (m₁.inter m₂).get! k = if m₂.contains k then m₁.get! k else default := by - simp_to_model [inter, get!, contains] -- using List.getValueCastD_filter - sorry + simp_to_model [inter, get!, contains] using List.getValueCastD_filter_containsKey_eq_if_containsKey + +theorem get!_inter_of_contains_right [LawfulBEq α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) + {k : α} [Inhabited (β k)] (h : m₂.contains k) : + (m₁.inter m₂).get! k = m₁.get! k := by + revert h + simp_to_model [inter, get!, contains] using List.getValueCastD_filter_containsKey_eq_containsKey + +theorem get!_inter_of_not_contains_right [LawfulBEq α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) + {k : α} [Inhabited (β k)] (h : m₂.contains k = false) : + (m₁.inter m₂).get! k = default := by + revert h + simp_to_model [inter, get!, contains] using List.getValueCastD_filter_of_not_contains_right + +theorem get!_inter_of_not_contains_left [LawfulBEq α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) + {k : α} [Inhabited (β k)] (h : m₁.contains k = false) : + (m₁.inter m₂).get! k = default := by + revert h + simp_to_model [inter, get!, contains] using List.getValueCastD_filter_of_not_contains_left /- getKey? -/ theorem getKey?_inter [EquivBEq α] [LawfulHashable α] diff --git a/src/Std/Data/Internal/List/Associative.lean b/src/Std/Data/Internal/List/Associative.lean index b98c6d091e59..0d62a91ba12f 100644 --- a/src/Std/Data/Internal/List/Associative.lean +++ b/src/Std/Data/Internal/List/Associative.lean @@ -5128,7 +5128,7 @@ theorem getEntry?_filter [BEq α] [EquivBEq α] intro p simp only [Option.all_guard, BEq.rfl, Bool.or_true] -theorem getValueCast?_filter_of_containsKey [BEq α] [LawfulBEq α] +theorem getValueCast?_filter_containsKey_eq_if_containsKey [BEq α] [LawfulBEq α] {l₁ l₂ : List ((a : α) × β a)} {k : α} (dl₁ : DistinctKeys l₁) : getValueCast? k (List.filter (fun p => containsKey p.fst l₂) l₁) = @@ -5160,6 +5160,36 @@ theorem getValueCast?_filter_of_containsKey [BEq α] [LawfulBEq α] case h_2 => rfl exact dl₁ +theorem getValueCast?_filter_containsKey_eq_containsKey [BEq α] [LawfulBEq α] + {l₁ l₂ : List ((a : α) × β a)} {k : α} + (dl₁ : DistinctKeys l₁) : + containsKey k l₂ → getValueCast? k (List.filter (fun p => containsKey p.fst l₂) l₁) = + getValueCast? k l₁ := by + intro h + rw [getValueCast?_filter_containsKey_eq_if_containsKey, h] + . simp only [↓reduceIte] + . exact dl₁ + +theorem getValueCast?_filter_of_not_contains [BEq α] [LawfulBEq α] + {l₁ l₂ : List ((a : α) × β a)} {k : α} + (dl₁ : DistinctKeys l₁) : + containsKey k l₁ = false → getValueCast? k (List.filter (fun p => containsKey p.fst l₂) l₁) = none := by + intro h + rw [getValueCast?_filter_containsKey_eq_if_containsKey] + . split + . rw [getValueCast?_eq_none h] + . rfl + . exact dl₁ + +theorem getValueCast?_filter_of_not_contains_right [BEq α] [LawfulBEq α] + {l₁ l₂ : List ((a : α) × β a)} {k : α} + (dl₁ : DistinctKeys l₁) : + containsKey k l₂ = false → getValueCast? k (List.filter (fun p => containsKey p.fst l₂) l₁) = none := by + intro h + rw [getValueCast?_filter_containsKey_eq_if_containsKey, h] + . simp only [Bool.false_eq_true, ↓reduceIte] + . exact dl₁ + theorem getEntry?_filterMap [BEq α] [EquivBEq α] {f : (a : α) → β a → Option (γ a)} {l : List ((a : α) × β a)} {k : α} (hl : DistinctKeys l) : @@ -5194,12 +5224,23 @@ theorem containsKey_filter_containsKey [BEq α] [EquivBEq α] {l₁ l₂ : List . exact hl₁ theorem contains_filter_iff [BEq α] [EquivBEq α] {l₁ l₂ : List ((a : α) × β a)} {hl₁ : DistinctKeys l₁} {k : α} : - (containsKey k (List.filter (fun p => containsKey p.fst (l₂)) l₁)) ↔ - (containsKey k l₁ ∧ containsKey k l₂) := by + (containsKey k (List.filter (fun p => containsKey p.fst l₂) l₁)) ↔ + (containsKey k l₁ ∧ containsKey k l₂) := by rw [containsKey_filter_containsKey] . simp . exact hl₁ +theorem getValueCast_filter_containsKey_eq_containsKey [BEq α] [LawfulBEq α] + {l₁ l₂ : List ((a : α) × β a)} {k : α} + (dl₁ : DistinctKeys l₁) {h₁ h₂} : + getValueCast k (List.filter (fun p => containsKey p.fst l₂) l₁) h₁ = getValueCast k l₁ h₂ := by + suffices some (getValueCast k (List.filter (fun p => containsKey p.fst l₂) l₁) h₁) = some (getValueCast k l₁ h₂) by injections + simp only [← getValueCast?_eq_some_getValueCast] + apply getValueCast?_filter_containsKey_eq_containsKey dl₁ + . rw [contains_filter_iff] at h₁ + . exact h₁.2 + . exact dl₁ + theorem not_contains_filter_of_not_contains [BEq α] [EquivBEq α] {l₁ l₂ : List ((a : α) × β a)} {hl₁ : DistinctKeys l₁} {k : α} : containsKey k l₁ = false → containsKey k (List.filter (fun p => containsKey p.fst l₂) l₁) = false := by intro h @@ -5214,6 +5255,52 @@ theorem not_contains_filter_of_not_containsKey_pred [BEq α] [EquivBEq α] {l₁ . simp [h] . exact hl₁ +theorem getValueCastD_filter_containsKey_eq_if_containsKey [BEq α] [LawfulBEq α] + {l₁ l₂ : List ((a : α) × β a)} {k : α} {fallback : β k} + (dl₁ : DistinctKeys l₁) : + getValueCastD k (List.filter (fun p => containsKey p.fst l₂) l₁) fallback = + if containsKey k l₂ = true then getValueCastD k l₁ fallback else fallback := by + split + case isTrue h => + rw [getValueCastD_eq_getValueCast?, getValueCastD_eq_getValueCast?] + congr 1 + apply getValueCast?_filter_containsKey_eq_containsKey dl₁ h + case isFalse h => + apply getValueCastD_eq_fallback + apply not_contains_filter_of_not_containsKey_pred + . exact dl₁ + . simp only [h] + +theorem getValueCastD_filter_containsKey_eq_containsKey [BEq α] [LawfulBEq α] + {l₁ l₂ : List ((a : α) × β a)} {k : α} {fallback: β k} + (dl₁ : DistinctKeys l₁) : + containsKey k l₂ → getValueCastD k (List.filter (fun p => containsKey p.fst l₂) l₁) fallback = + getValueCastD k l₁ fallback:= by + intro h + rw [getValueCastD_eq_getValueCast?, getValueCastD_eq_getValueCast?] + congr 1 + apply getValueCast?_filter_containsKey_eq_containsKey dl₁ h + +theorem getValueCastD_filter_of_not_contains_left [BEq α] [LawfulBEq α] + {l₁ l₂ : List ((a : α) × β a)} {k : α} {fallback} + (dl₁ : DistinctKeys l₁) : + containsKey k l₁ = false → getValueCastD k (List.filter (fun p => containsKey p.fst l₂) l₁) fallback = fallback := by + intro h + apply getValueCastD_eq_fallback + apply not_contains_filter_of_not_contains + . exact dl₁ + . exact h + +theorem getValueCastD_filter_of_not_contains_right [BEq α] [LawfulBEq α] + {l₁ l₂ : List ((a : α) × β a)} {k : α} {fallback} + (dl₁ : DistinctKeys l₁) : + containsKey k l₂ = false → getValueCastD k (List.filter (fun p => containsKey p.fst l₂) l₁) fallback = fallback := by + intro h + apply getValueCastD_eq_fallback + apply not_contains_filter_of_not_containsKey_pred + . exact dl₁ + . exact h + theorem containsKey_of_containsKey_filterMap' [BEq α] [EquivBEq α] {f : ((a : α) × β a) → Option ((a : α) × γ a)} (hf : ∀ p, (f p).all (·.1 == p.1)) From 64a78917fe2269e9fb6b2481e6a952f28fcbf0dd Mon Sep 17 00:00:00 2001 From: Wojciech Rozowski Date: Thu, 6 Nov 2025 15:00:40 +0100 Subject: [PATCH 15/88] getKey? lemmas are done --- src/Std/Data/DHashMap/Internal/RawLemmas.lean | 17 +++-- src/Std/Data/Internal/List/Associative.lean | 64 +++++++++++++++++++ 2 files changed, 75 insertions(+), 6 deletions(-) diff --git a/src/Std/Data/DHashMap/Internal/RawLemmas.lean b/src/Std/Data/DHashMap/Internal/RawLemmas.lean index 4db16b4ca8bc..0d12a16c2a48 100644 --- a/src/Std/Data/DHashMap/Internal/RawLemmas.lean +++ b/src/Std/Data/DHashMap/Internal/RawLemmas.lean @@ -2892,20 +2892,25 @@ theorem getKey?_inter [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} : (m₁.inter m₂).getKey? k = if m₂.contains k then m₁.getKey? k else none := by - simp_to_model [inter, contains, getKey?] -- using List.getKey?_filter - sorry + simp_to_model [inter, contains, getKey?] using List.getKey?_filter_containsKey_eq_if_containsKey theorem getKey?_inter_of_contains_right [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} (h : m₂.contains k) : (m₁.inter m₂).getKey? k = m₁.getKey? k := by - simp_to_model [contains, getKey?, inter] -- using List.getKey?_filter_of_pred_true - sorry + revert h + simp_to_model [contains, getKey?, inter] using List.getKey?_filter_containsKey_eq_containsKey theorem getKey?_inter_of_not_contains_right [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} (h : m₂.contains k = false) : (m₁.inter m₂).getKey? k = none := by - simp_to_model [contains, getKey?, inter] -- using List.getKey?_filter_of_pred_false - sorry + revert h + simp_to_model [contains, getKey?, inter] using List.getKey?_filter_of_not_contains_right + +theorem getKey?_inter_of_not_contains_left [EquivBEq α] [LawfulHashable α] + (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} (h : m₁.contains k = false) : + (m₁.inter m₂).getKey? k = none := by + revert h + simp_to_model [contains, getKey?, inter] using List.getKey?_filter_of_not_contains /- getKey -/ theorem getKey_inter [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) diff --git a/src/Std/Data/Internal/List/Associative.lean b/src/Std/Data/Internal/List/Associative.lean index 0d62a91ba12f..b89fd14f4b13 100644 --- a/src/Std/Data/Internal/List/Associative.lean +++ b/src/Std/Data/Internal/List/Associative.lean @@ -5160,6 +5160,70 @@ theorem getValueCast?_filter_containsKey_eq_if_containsKey [BEq α] [LawfulBEq case h_2 => rfl exact dl₁ +theorem getKey?_filter_containsKey_eq_if_containsKey [BEq α] [EquivBEq α] + {l₁ l₂ : List ((a : α) × β a)} {k : α} + (dl₁ : DistinctKeys l₁) : + getKey? k (List.filter (fun p => containsKey p.fst l₂) l₁) = + if containsKey k l₂ = true then getKey? k l₁ else none := by + simp only [getKey?_eq_getEntry?] + split + case isTrue h => + rw [getEntry?_filter] + · congr 1 + simp [Option.filter] + split + case h_1 _ val heq => + rw [heq] + have key_eq := beq_of_getEntry?_eq_some heq + rw [@containsKey_congr _ _ _ _ l₂ _ _ (PartialEquivBEq.symm key_eq)] at h + simp [h] + case h_2 _ heq => + simp only [heq] + · exact dl₁ + case isFalse h => + suffices (getEntry? k (List.filter (fun p => containsKey p.fst l₂) l₁)) = none by + simp [this] + rw [getEntry?_filter] + · simp [Option.filter] + split + case h_1 _ kv heq => + have key_eq := beq_of_getEntry?_eq_some heq + simp at h + rw [@containsKey_congr _ _ _ _ l₂ _ _ (PartialEquivBEq.symm key_eq)] at h + simp [h] + case h_2 => rfl + · exact dl₁ + +theorem getKey?_filter_containsKey_eq_containsKey [BEq α] [EquivBEq α] + {l₁ l₂ : List ((a : α) × β a)} {k : α} + (dl₁ : DistinctKeys l₁) : + containsKey k l₂ → getKey? k (List.filter (fun p => containsKey p.fst l₂) l₁) = + getKey? k l₁ := by + intro h + rw [getKey?_filter_containsKey_eq_if_containsKey, h] + · simp only [↓reduceIte] + · exact dl₁ + +theorem getKey?_filter_of_not_contains [BEq α] [EquivBEq α] + {l₁ l₂ : List ((a : α) × β a)} {k : α} + (dl₁ : DistinctKeys l₁) : + containsKey k l₁ = false → getKey? k (List.filter (fun p => containsKey p.fst l₂) l₁) = none := by + intro h + rw [getKey?_filter_containsKey_eq_if_containsKey] + · split + · rw [getKey?_eq_none h] + · rfl + · exact dl₁ + +theorem getKey?_filter_of_not_contains_right [BEq α] [EquivBEq α] + {l₁ l₂ : List ((a : α) × β a)} {k : α} + (dl₁ : DistinctKeys l₁) : + containsKey k l₂ = false → getKey? k (List.filter (fun p => containsKey p.fst l₂) l₁) = none := by + intro h + rw [getKey?_filter_containsKey_eq_if_containsKey, h] + · simp only [Bool.false_eq_true, ↓reduceIte] + · exact dl₁ + theorem getValueCast?_filter_containsKey_eq_containsKey [BEq α] [LawfulBEq α] {l₁ l₂ : List ((a : α) × β a)} {k : α} (dl₁ : DistinctKeys l₁) : From 929860892d70e751004997e6c45cba95640b6898 Mon Sep 17 00:00:00 2001 From: Wojciech Rozowski Date: Thu, 6 Nov 2025 19:26:04 +0100 Subject: [PATCH 16/88] Add const lemmas --- src/Std/Data/DHashMap/Internal/RawLemmas.lean | 162 ++++++++++-- src/Std/Data/Internal/List/Associative.lean | 238 ++++++++++++++++++ 2 files changed, 386 insertions(+), 14 deletions(-) diff --git a/src/Std/Data/DHashMap/Internal/RawLemmas.lean b/src/Std/Data/DHashMap/Internal/RawLemmas.lean index 0d12a16c2a48..b64cf9312088 100644 --- a/src/Std/Data/DHashMap/Internal/RawLemmas.lean +++ b/src/Std/Data/DHashMap/Internal/RawLemmas.lean @@ -2917,24 +2917,57 @@ theorem getKey_inter [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h {k : α} {h_contains : (m₁.inter m₂).contains k} : (m₁.inter m₂).getKey k h_contains = m₁.getKey k (by simp [contains_inter_iff h₁ h₂] at h_contains; exact h_contains.1) := by - simp_to_model [inter, contains, getKey] -- using List.getKey_filter - sorry + simp_to_model [inter, contains, getKey] using List.getKey_filter_containsKey_eq_containsKey /- getKeyD -/ theorem getKeyD_inter [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k fallback : α} : (m₁.inter m₂).getKeyD k fallback = if m₂.contains k then m₁.getKeyD k fallback else fallback := by - simp_to_model [inter, getKeyD, contains] -- using List.getKeyD_filter - sorry + simp_to_model [inter, getKeyD, contains] using List.getKeyD_filter_containsKey_eq_if_containsKey + +theorem getKeyD_inter_of_contains_right [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) + (h₂ : m₂.val.WF) {k fallback : α} (h : m₂.contains k) : + (m₁.inter m₂).getKeyD k fallback = m₁.getKeyD k fallback := by + revert h + simp_to_model [inter, getKeyD, contains] using List.getKeyD_filter_containsKey_eq_containsKey + +theorem getKeyD_inter_of_not_contains_right [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) + (h₂ : m₂.val.WF) {k fallback : α} (h : m₂.contains k = false) : + (m₁.inter m₂).getKeyD k fallback = fallback := by + revert h + simp_to_model [inter, getKeyD, contains] using List.getKeyD_filter_of_not_contains_right + +theorem getKeyD_inter_of_not_contains_left [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) + (h₂ : m₂.val.WF) {k fallback : α} (h : m₁.contains k = false) : + (m₁.inter m₂).getKeyD k fallback = fallback := by + revert h + simp_to_model [inter, getKeyD, contains] using List.getKeyD_filter_of_not_contains_left /- getKey! -/ -theorem getKey!_inter [EquivBEq α] [LawfulHashable α] [Inhabited α] - (h₁ : m₁.1.WF) (h₂ : m₂.1.WF) {k : α} : +theorem getKey!_inter [EquivBEq α] [LawfulHashable α] [Inhabited α] (h₁ : m₁.val.WF) + (h₂ : m₂.val.WF) {k : α} [Inhabited α] : (m₁.inter m₂).getKey! k = if m₂.contains k then m₁.getKey! k else default := by - simp_to_model [inter, getKey!, contains] -- using List.getKeyD_filter - sorry + simp_to_model [inter, getKey!, contains] using List.getKeyD_filter_containsKey_eq_if_containsKey + +theorem getKey!_inter_of_contains_right [EquivBEq α] [LawfulHashable α] [Inhabited α] (h₁ : m₁.val.WF) + (h₂ : m₂.val.WF) {k : α} (h : m₂.contains k) : + (m₁.inter m₂).getKey! k = m₁.getKey! k := by + revert h + simp_to_model [inter, getKey!, contains] using List.getKeyD_filter_containsKey_eq_containsKey + +theorem getKey!_inter_of_not_contains_right [EquivBEq α] [LawfulHashable α] [Inhabited α] (h₁ : m₁.val.WF) + (h₂ : m₂.val.WF) {k : α} (h : m₂.contains k = false) : + (m₁.inter m₂).getKey! k = default := by + revert h + simp_to_model [inter, getKey!, contains] using List.getKeyD_filter_of_not_contains_right + +theorem getKey!_inter_of_not_contains_left [EquivBEq α] [LawfulHashable α] [Inhabited α] (h₁ : m₁.val.WF) + (h₂ : m₂.val.WF) {k : α} (h : m₁.contains k = false) : + (m₁.inter m₂).getKey! k = default := by + revert h + simp_to_model [inter, getKey!, contains] using List.getKeyD_filter_of_not_contains_left /- size -/ theorem size_inter_le_size_left [EquivBEq α] [LawfulHashable α] @@ -2948,22 +2981,123 @@ theorem size_inter_le_size_right [EquivBEq α] [LawfulHashable α] simp_to_model [inter, size, contains] -- using List.length_filter_containsKey_le sorry -theorem size_inter_eq_size_left_of_subset [EquivBEq α] [LawfulHashable α] +theorem size_inter_eq_size_left [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) (h : ∀ (a : α), m₁.contains a → m₂.contains a) : (m₁.inter m₂).1.size = m₁.1.size := by - simp_to_model [inter, size, contains] -- using List.length_filter_of_forall_pred + revert h + simp_to_model [inter, size, contains] using List.length_filter_containsKey_of_forall + +theorem size_inter_eq_size_right [EquivBEq α] [LawfulHashable α] + (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) + (h : ∀ (a : α), m₂.contains a → m₁.contains a) : + (m₁.inter m₂).1.size = m₂.1.size := by + revert h + simp_to_model [inter, size, contains] sorry /- isEmpty -/ @[simp] -theorem isEmpty_inter [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) : - (m₁.inter m₂).1.isEmpty = (m₁.1.isEmpty || m₂.1.isEmpty) := by - simp_to_model [isEmpty, inter, contains] -- using List.isEmpty_filter_containsKey - sorry +theorem isEmpty_inter_left [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) (h : m₁.1.isEmpty): + (m₁.inter m₂).1.isEmpty = true := by + revert h + simp_to_model [isEmpty, inter, contains] using List.isEmpty_filter_containsKey_left + +/- isEmpty -/ +@[simp] +theorem isEmpty_inter_right [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) (h : m₁.1.isEmpty): + (m₁.inter m₂).1.isEmpty = true := by + revert h + simp_to_model [isEmpty, inter, contains] using List.isEmpty_filter_containsKey_left end Inter +namespace Const + +variable {β : Type v} {m₁ m₂ : Raw₀ α (fun _ => β)} + +/- get? -/ +theorem get?_inter [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} : + Const.get? (m₁.inter m₂) k = + if m₂.contains k then Const.get? m₁ k else none := by + simp_to_model [inter, Const.get?, contains] using List.getValue?_filter_containsKey_eq_if_containsKey + +theorem get?_inter_of_contains_right [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) + {k : α} (h : m₂.contains k) : + Const.get? (m₁.inter m₂) k = Const.get? m₁ k := by + revert h + simp_to_model [inter, Const.get?, contains] using List.getValue?_filter_containsKey_eq_containsKey + +theorem get?_inter_of_not_contains_left [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) + {k : α} (h : m₁.contains k = false) : + Const.get? (m₁.inter m₂) k = none := by + revert h + simp_to_model [inter, Const.get?, contains] using List.getValue?_filter_of_not_contains + +theorem get?_inter_of_not_contains_right [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) + {k : α} (h : m₂.contains k = false) : + Const.get? (m₁.inter m₂) k = none := by + revert h + simp_to_model [inter, Const.get?, contains] using List.getValue?_filter_of_not_contains_right + +theorem get_inter [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) + {k : α} {h_contains : (m₁.inter m₂).contains k} : + Const.get (m₁.inter m₂) k h_contains = + Const.get m₁ k ((contains_inter_iff h₁ h₂).1 h_contains).1 := by + simp_to_model [inter, Const.get, contains] using List.getValue_filter_containsKey_eq_containsKey + +/- getD -/ +theorem getD_inter [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) + {k : α} {fallback : β} : + Const.getD (m₁.inter m₂) k fallback = + if m₂.contains k then Const.getD m₁ k fallback else fallback := by + simp_to_model [inter, Const.getD, contains] using List.getValueD_filter_containsKey_eq_if_containsKey + +theorem getD_inter_of_contains_right [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) + {k : α} {fallback : β} (h : m₂.contains k) : + Const.getD (m₁.inter m₂) k fallback = Const.getD m₁ k fallback := by + revert h + simp_to_model [inter, Const.getD, contains] using List.getValueD_filter_containsKey_eq_containsKey + +theorem getD_inter_of_not_contains_right [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) + {k : α} {fallback : β} (h : m₂.contains k = false) : + Const.getD (m₁.inter m₂) k fallback = fallback := by + revert h + simp_to_model [inter, Const.getD, contains] using List.getValueD_filter_of_not_contains_right + +theorem getD_inter_of_not_contains_left [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) + {k : α} {fallback : β} (h : m₁.contains k = false) : + Const.getD (m₁.inter m₂) k fallback = fallback := by + revert h + simp_to_model [inter, Const.getD, contains] using List.getValueD_filter_of_not_contains_left + +/- get! -/ +theorem get!_inter [EquivBEq α] [LawfulHashable α] [Inhabited β] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) + {k : α} : + Const.get! (m₁.inter m₂) k = + if m₂.contains k then Const.get! m₁ k else default := by + simp_to_model [inter, Const.get!, contains] using List.getValueD_filter_containsKey_eq_if_containsKey + +theorem get!_inter_of_contains_right [EquivBEq α] [LawfulHashable α] [Inhabited β] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) + {k : α} (h : m₂.contains k) : + Const.get! (m₁.inter m₂) k = Const.get! m₁ k := by + revert h + simp_to_model [inter, Const.get!, contains] using List.getValueD_filter_containsKey_eq_containsKey + +theorem get!_inter_of_not_contains_right [EquivBEq α] [LawfulHashable α] [Inhabited β] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) + {k : α} (h : m₂.contains k = false) : + Const.get! (m₁.inter m₂) k = default := by + revert h + simp_to_model [inter, Const.get!, contains] using List.getValueD_filter_of_not_contains_right + +theorem get!_inter_of_not_contains_left [EquivBEq α] [LawfulHashable α] [Inhabited β] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) + {k : α} (h : m₁.contains k = false) : + Const.get! (m₁.inter m₂) k = default := by + revert h + simp_to_model [inter, Const.get!, contains] using List.getValueD_filter_of_not_contains_left + +end Const + section Alter theorem isEmpty_alter_eq_isEmpty_erase [LawfulBEq α] (h : m.1.WF) {k : α} diff --git a/src/Std/Data/Internal/List/Associative.lean b/src/Std/Data/Internal/List/Associative.lean index b89fd14f4b13..21e19194359a 100644 --- a/src/Std/Data/Internal/List/Associative.lean +++ b/src/Std/Data/Internal/List/Associative.lean @@ -5305,6 +5305,18 @@ theorem getValueCast_filter_containsKey_eq_containsKey [BEq α] [LawfulBEq α] . exact h₁.2 . exact dl₁ +theorem getKey_filter_containsKey_eq_containsKey [BEq α] [EquivBEq α] + {l₁ l₂ : List ((a : α) × β a)} {k : α} + (dl₁ : DistinctKeys l₁) {h₁ h₂} : + getKey k (List.filter (fun p => containsKey p.fst l₂) l₁) h₁ = getKey k l₁ h₂ := by + suffices some (getKey k (List.filter (fun p => containsKey p.fst l₂) l₁) h₁) = some (getKey k l₁ h₂) by + injections + simp only [← getKey?_eq_some_getKey] + apply getKey?_filter_containsKey_eq_containsKey dl₁ + · rw [contains_filter_iff] at h₁ + · exact h₁.2 + · exact dl₁ + theorem not_contains_filter_of_not_contains [BEq α] [EquivBEq α] {l₁ l₂ : List ((a : α) × β a)} {hl₁ : DistinctKeys l₁} {k : α} : containsKey k l₁ = false → containsKey k (List.filter (fun p => containsKey p.fst l₂) l₁) = false := by intro h @@ -5319,6 +5331,64 @@ theorem not_contains_filter_of_not_containsKey_pred [BEq α] [EquivBEq α] {l₁ . simp [h] . exact hl₁ +theorem getValue?_filter_containsKey_eq_if_containsKey {β : Type v} [BEq α] [EquivBEq α] + {l₁ l₂ : List ((_ : α) × β)} {k : α} + (dl₁ : DistinctKeys l₁) : + getValue? k (List.filter (fun p => containsKey p.fst l₂) l₁) = + if containsKey k l₂ = true then getValue? k l₁ else none := by + simp only [getValue?_eq_getEntry?] + split + case isTrue h => + congr 1 + rw [getEntry?_filter] + . simp only [Option.filter] + split + case h_1 _ val heq => + simp [heq] + have := PartialEquivBEq.symm <| @beq_of_getEntry?_eq_some α (fun _ => β) _ l₁ k val heq + rw [← containsKey_congr this] + exact h + case h_2 _ heq => + simp [heq] + . exact dl₁ + case isFalse h => + simp only [Option.map_eq_none_iff, getEntry?_eq_none] + simp only [Bool.not_eq_true] at h + apply not_contains_filter_of_not_containsKey_pred + . exact dl₁ + . exact h + +theorem getValue?_filter_containsKey_eq_containsKey {β : Type v} [BEq α] [EquivBEq α] + {l₁ l₂ : List ((_ : α) × β)} {k : α} + (dl₁ : DistinctKeys l₁) : + containsKey k l₂ → getValue? k (List.filter (fun p => containsKey p.fst l₂) l₁) = + getValue? k l₁ := by + intro h + rw [getValue?_filter_containsKey_eq_if_containsKey, h] + · simp only [↓reduceIte] + · exact dl₁ + +theorem getValue?_filter_of_not_contains {β : Type v} [BEq α] [EquivBEq α] + {l₁ l₂ : List ((_ : α) × β)} {k : α} + (dl₁ : DistinctKeys l₁) : + containsKey k l₁ = false → getValue? k (List.filter (fun p => containsKey p.fst l₂) l₁) = none := by + intro h + rw [getValue?_filter_containsKey_eq_if_containsKey] + · split + · rw [@getValue?_eq_none α β _ l₁ k] + exact h + · rfl + · exact dl₁ + +theorem getValue?_filter_of_not_contains_right {β : Type v} [BEq α] [EquivBEq α] + {l₁ l₂ : List ((_ : α) × β)} {k : α} + (dl₁ : DistinctKeys l₁) : + containsKey k l₂ = false → getValue? k (List.filter (fun p => containsKey p.fst l₂) l₁) = none := by + intro h + rw [getValue?_filter_containsKey_eq_if_containsKey, h] + · simp only [Bool.false_eq_true, ↓reduceIte] + · exact dl₁ + theorem getValueCastD_filter_containsKey_eq_if_containsKey [BEq α] [LawfulBEq α] {l₁ l₂ : List ((a : α) × β a)} {k : α} {fallback : β k} (dl₁ : DistinctKeys l₁) : @@ -5365,6 +5435,110 @@ theorem getValueCastD_filter_of_not_contains_right [BEq α] [LawfulBEq α] . exact dl₁ . exact h +theorem getKeyD_filter_containsKey_eq_if_containsKey [BEq α] [EquivBEq α] + {l₁ l₂ : List ((a : α) × β a)} {k : α} {fallback : α} + (dl₁ : DistinctKeys l₁) : + getKeyD k (List.filter (fun p => containsKey p.fst l₂) l₁) fallback = + if containsKey k l₂ = true then getKeyD k l₁ fallback else fallback := by + split + case isTrue h => + rw [getKeyD_eq_getKey?, getKeyD_eq_getKey?] + congr 1 + apply getKey?_filter_containsKey_eq_containsKey dl₁ h + case isFalse h => + apply getKeyD_eq_fallback + apply not_contains_filter_of_not_containsKey_pred + · exact dl₁ + · simp only [h] + +theorem getKeyD_filter_containsKey_eq_containsKey [BEq α] [EquivBEq α] + {l₁ l₂ : List ((a : α) × β a)} {k : α} {fallback : α} + (dl₁ : DistinctKeys l₁) : + containsKey k l₂ → getKeyD k (List.filter (fun p => containsKey p.fst l₂) l₁) fallback = + getKeyD k l₁ fallback := by + intro h + rw [getKeyD_eq_getKey?, getKeyD_eq_getKey?] + congr 1 + apply getKey?_filter_containsKey_eq_containsKey dl₁ h + +theorem getKeyD_filter_of_not_contains_left [BEq α] [EquivBEq α] + {l₁ l₂ : List ((a : α) × β a)} {k : α} {fallback : α} + (dl₁ : DistinctKeys l₁) : + containsKey k l₁ = false → getKeyD k (List.filter (fun p => containsKey p.fst l₂) l₁) fallback = fallback := by + intro h + apply getKeyD_eq_fallback + apply not_contains_filter_of_not_contains + · exact dl₁ + · exact h + +theorem getKeyD_filter_of_not_contains_right [BEq α] [EquivBEq α] + {l₁ l₂ : List ((a : α) × β a)} {k : α} {fallback : α} + (dl₁ : DistinctKeys l₁) : + containsKey k l₂ = false → getKeyD k (List.filter (fun p => containsKey p.fst l₂) l₁) fallback = fallback := by + intro h + apply getKeyD_eq_fallback + apply not_contains_filter_of_not_containsKey_pred + · exact dl₁ + · exact h + +theorem getValue_filter_containsKey_eq_containsKey {β : Type v} [BEq α] [EquivBEq α] + {l₁ l₂ : List ((_ : α) × β)} {k : α} + (dl₁ : DistinctKeys l₁) {h₁ h₂} : + getValue k (List.filter (fun p => containsKey p.fst l₂) l₁) h₁ = getValue k l₁ h₂ := by + suffices some (getValue k (List.filter (fun p => containsKey p.fst l₂) l₁) h₁) = some (getValue k l₁ h₂) by + injections + simp only [← getValue?_eq_some_getValue] + apply getValue?_filter_containsKey_eq_containsKey dl₁ + · rw [contains_filter_iff] at h₁ + · exact h₁.2 + · exact dl₁ + +theorem getValueD_filter_containsKey_eq_if_containsKey {β : Type v} [BEq α] [EquivBEq α] + {l₁ l₂ : List ((_ : α) × β)} {k : α} {fallback : β} + (dl₁ : DistinctKeys l₁) : + getValueD k (List.filter (fun p => containsKey p.fst l₂) l₁) fallback = + if containsKey k l₂ = true then getValueD k l₁ fallback else fallback := by + split + case isTrue h => + rw [getValueD_eq_getValue?, getValueD_eq_getValue?] + congr 1 + apply getValue?_filter_containsKey_eq_containsKey dl₁ h + case isFalse h => + apply getValueD_eq_fallback + apply not_contains_filter_of_not_containsKey_pred + · exact dl₁ + · simp only [h] + +theorem getValueD_filter_containsKey_eq_containsKey {β : Type v} [BEq α] [EquivBEq α] + {l₁ l₂ : List ((_ : α) × β)} {k : α} {fallback : β} + (dl₁ : DistinctKeys l₁) : + containsKey k l₂ → getValueD k (List.filter (fun p => containsKey p.fst l₂) l₁) fallback = + getValueD k l₁ fallback := by + intro h + rw [getValueD_eq_getValue?, getValueD_eq_getValue?] + congr 1 + apply getValue?_filter_containsKey_eq_containsKey dl₁ h + +theorem getValueD_filter_of_not_contains_left {β : Type v} [BEq α] [EquivBEq α] + {l₁ l₂ : List ((_ : α) × β)} {k : α} {fallback : β} + (dl₁ : DistinctKeys l₁) : + containsKey k l₁ = false → getValueD k (List.filter (fun p => containsKey p.fst l₂) l₁) fallback = fallback := by + intro h + apply getValueD_eq_fallback + apply not_contains_filter_of_not_contains + · exact dl₁ + · exact h + +theorem getValueD_filter_of_not_contains_right {β : Type v} [BEq α] [EquivBEq α] + {l₁ l₂ : List ((_ : α) × β)} {k : α} {fallback : β} + (dl₁ : DistinctKeys l₁) : + containsKey k l₂ = false → getValueD k (List.filter (fun p => containsKey p.fst l₂) l₁) fallback = fallback := by + intro h + apply getValueD_eq_fallback + apply not_contains_filter_of_not_containsKey_pred + · exact dl₁ + · exact h + theorem containsKey_of_containsKey_filterMap' [BEq α] [EquivBEq α] {f : ((a : α) × β a) → Option ((a : α) × γ a)} (hf : ∀ p, (f p).all (·.1 == p.1)) @@ -5872,6 +6046,70 @@ theorem length_filter_key_eq_length_iff [BEq α] [EquivBEq α] {f : α → Bool} simp only [getKey, getKey?_eq_getEntry?, this] at h exact h +theorem List.length_filter_containsKey_of_forall [BEq α] [EquivBEq α] + {l₁ l₂ : List ((a : α) × β a)} (hl₁ : DistinctKeys l₁) + (w : ∀ (a : α), containsKey a l₁ → containsKey a l₂) : + (l₁.filter fun p => containsKey p.fst l₂).length = l₁.length := by + rw [← List.countP_eq_length_filter] + induction l₁ + case nil => simp + case cons h t ih => + simp + rw [List.countP_cons_of_pos] + . rw [List.distinctKeys_cons_iff] at hl₁ + specialize ih hl₁.1 + simp only [Nat.add_right_cancel_iff] + apply ih + intro a hc + apply w + rw [containsKey_cons] + simp [hc] + . apply w + rw [containsKey_cons, BEq.rfl, Bool.true_or] + +theorem List.length_filter_containsKey_right [BEq α] [EquivBEq α] + {l₁ l₂ : List ((a : α) × β a)} (hl₁ : DistinctKeys l₁) + (w : ∀ (a : α), containsKey a l₂ → containsKey a l₁) : + (l₁.filter fun p => containsKey p.fst l₂).length = l₂.length := by + sorry + +theorem nil_of_containsKey_eq_false [BEq α] [EquivBEq α] {l : List ((a : α) × β a)} : + (∀ k, containsKey k l = false) ↔ l = [] := by + constructor + case mp => + intro hyp + rw [List.eq_nil_iff_forall_not_mem] + intro a + specialize hyp a.fst + rw [List.containsKey_eq_false_iff_forall_mem_keys] at hyp + specialize hyp a.fst + have : (a.fst == a.fst) → a.fst ∉keys l := by + intro h₁ h₂ + specialize hyp h₂ + rw [hyp] at h₁ + contradiction + specialize this (by simp) + intro h + have := Internal.List.fst_mem_keys_of_mem h + contradiction + case mpr => + intro hyp k + simp [hyp] + +theorem isEmpty_filter_containsKey_left [BEq α] [EquivBEq α] + {l₁ l₂ : List ((a : α) × β a)} : + l₁.isEmpty → (List.filter (fun p => containsKey p.fst l₂) l₁).isEmpty := by + intro hyp + simp at hyp + simp [hyp] + +theorem isEmpty_filter_containsKey_right [BEq α] [EquivBEq α] + {l₁ l₂ : List ((a : α) × β a)} : + l₂.isEmpty → (List.filter (fun p => containsKey p.fst l₂) l₁).isEmpty := by + intro hyp + simp at hyp + simp [hyp] + theorem perm_filter_self_iff {f : α → Bool} {l : List α} : (l.filter f).Perm l ↔ ∀ a ∈ l, f a = true := by rw (occs := [1]) [← (List.filter_append_perm f _).congr_right, From c7d1e136568b531e0bdfc0f88f5dc6058691d7ee Mon Sep 17 00:00:00 2001 From: Wojciech Rozowski Date: Thu, 6 Nov 2025 19:27:39 +0100 Subject: [PATCH 17/88] Remove sorried proof --- src/Std/Data/Internal/List/Associative.lean | 6 ------ 1 file changed, 6 deletions(-) diff --git a/src/Std/Data/Internal/List/Associative.lean b/src/Std/Data/Internal/List/Associative.lean index 21e19194359a..4cd53f228077 100644 --- a/src/Std/Data/Internal/List/Associative.lean +++ b/src/Std/Data/Internal/List/Associative.lean @@ -6067,12 +6067,6 @@ theorem List.length_filter_containsKey_of_forall [BEq α] [EquivBEq α] . apply w rw [containsKey_cons, BEq.rfl, Bool.true_or] -theorem List.length_filter_containsKey_right [BEq α] [EquivBEq α] - {l₁ l₂ : List ((a : α) × β a)} (hl₁ : DistinctKeys l₁) - (w : ∀ (a : α), containsKey a l₂ → containsKey a l₁) : - (l₁.filter fun p => containsKey p.fst l₂).length = l₂.length := by - sorry - theorem nil_of_containsKey_eq_false [BEq α] [EquivBEq α] {l : List ((a : α) × β a)} : (∀ k, containsKey k l = false) ↔ l = [] := by constructor From 3bf6ad47b5dcb204fd5670a62c32f55b900815c5 Mon Sep 17 00:00:00 2001 From: Wojciech Rozowski Date: Thu, 6 Nov 2025 21:03:29 +0100 Subject: [PATCH 18/88] push unfinished work --- src/Std/Data/Internal/List/Associative.lean | 41 ++++++++++++++++++++- 1 file changed, 40 insertions(+), 1 deletion(-) diff --git a/src/Std/Data/Internal/List/Associative.lean b/src/Std/Data/Internal/List/Associative.lean index 4cd53f228077..2619fbccff3b 100644 --- a/src/Std/Data/Internal/List/Associative.lean +++ b/src/Std/Data/Internal/List/Associative.lean @@ -6191,7 +6191,6 @@ theorem isEmpty_filter_key_iff [BEq α] [EquivBEq α] {f : α → Bool} simp only [getKey, getKey?_eq_getEntry?, this] at h exact h - namespace Const theorem getKey_getValue_mem [BEq α] [EquivBEq α] {β : Type v} {l : List ((_ : α) × β)} {k : α} {h} : @@ -6463,6 +6462,46 @@ theorem length_filter_key_eq_length_iff {β : Type v} [BEq α] [EquivBEq α] simp [← List.filterMap_eq_filter, forall_mem_iff_forall_contains_getKey_getValue (p := fun a b => f a = true) distinct] +theorem length_filter_containsKey_le [BEq α] [EquivBEq α] + {l₁ l₂ acc : List ((a : α) × β a)} + (hyp : ∀ e, e ∈ l₁ → containsKey e.fst acc = false) + (dl₁ : DistinctKeys (l₂ ++ acc)) + (dl₂ : DistinctKeys (l₁ ++ acc)) : + (l₁.filter fun p => containsKey p.fst (acc ++ l₂)).length ≤ l₂.length + acc.length := by + induction l₂ generalizing acc with + | nil => + simp only [List.append_nil, List.length_nil] + induction l₁ + case nil => simp + case cons h t ih => + simp [List.filter_cons] + simp [hyp h (by simp)] + simp at ih + apply ih + intro e mem + apply hyp + simp [mem] + | cons h t ih => + by_cases mem : containsKey h.fst l₁ + . specialize @ih (acc ++ [h]) ?refine1 ?refine2 ?refine3 + case refine1 => sorry + case refine2 => sorry + sorry + sorry + . suffices sublemma : (List.filter (fun p => containsKey p.fst (acc ++ h :: t)) l₁) = (List.filter (fun p => containsKey p.fst (acc ++ t)) l₁) by + rw [sublemma] + apply Nat.le_trans + . apply ih + . exact hyp + . sorry + . exact dl₂ + simp + sorry + + + + + theorem isEmpty_filterMap_eq_true [BEq α] [EquivBEq α] {β : Type v} {γ : Type w} {f : (_ : α) → β → Option γ} {l : List ((_ : α) × β)} (distinct : DistinctKeys l) : (l.filterMap fun p => (f p.1 p.2).map (fun x => (⟨p.1, x⟩ : (_ : α) × γ))).isEmpty = true ↔ From 20bd320ecfd063ebec97b8874eaa36f095366591 Mon Sep 17 00:00:00 2001 From: Wojciech Rozowski Date: Fri, 7 Nov 2025 11:05:10 +0000 Subject: [PATCH 19/88] Prove length lemma --- src/Std/Data/Internal/List/Associative.lean | 118 ++++++++++++++------ 1 file changed, 81 insertions(+), 37 deletions(-) diff --git a/src/Std/Data/Internal/List/Associative.lean b/src/Std/Data/Internal/List/Associative.lean index 2619fbccff3b..4f86c754e3d9 100644 --- a/src/Std/Data/Internal/List/Associative.lean +++ b/src/Std/Data/Internal/List/Associative.lean @@ -6463,44 +6463,88 @@ theorem length_filter_key_eq_length_iff {β : Type v} [BEq α] [EquivBEq α] forall_mem_iff_forall_contains_getKey_getValue (p := fun a b => f a = true) distinct] theorem length_filter_containsKey_le [BEq α] [EquivBEq α] - {l₁ l₂ acc : List ((a : α) × β a)} - (hyp : ∀ e, e ∈ l₁ → containsKey e.fst acc = false) - (dl₁ : DistinctKeys (l₂ ++ acc)) - (dl₂ : DistinctKeys (l₁ ++ acc)) : - (l₁.filter fun p => containsKey p.fst (acc ++ l₂)).length ≤ l₂.length + acc.length := by - induction l₂ generalizing acc with - | nil => - simp only [List.append_nil, List.length_nil] - induction l₁ - case nil => simp - case cons h t ih => - simp [List.filter_cons] - simp [hyp h (by simp)] - simp at ih - apply ih - intro e mem - apply hyp - simp [mem] - | cons h t ih => - by_cases mem : containsKey h.fst l₁ - . specialize @ih (acc ++ [h]) ?refine1 ?refine2 ?refine3 - case refine1 => sorry - case refine2 => sorry - sorry - sorry - . suffices sublemma : (List.filter (fun p => containsKey p.fst (acc ++ h :: t)) l₁) = (List.filter (fun p => containsKey p.fst (acc ++ t)) l₁) by - rw [sublemma] - apply Nat.le_trans - . apply ih - . exact hyp - . sorry - . exact dl₂ + {l₁ l₂ : List ((a : α) × β a)} + (dl₁ : DistinctKeys l₂) + (dl₂ : DistinctKeys l₁) : + (l₁.filter fun p => containsKey p.fst l₂).length ≤ l₂.length := by + induction l₁ generalizing l₂ + case nil => simp + case cons h t ih => + by_cases heq : containsKey h.fst l₂ + case pos => + simp [← List.countP_eq_length_filter] + simp only [heq, List.countP_cons_of_pos] + have len_eq : l₂.length = (eraseKey h.fst l₂).length + 1 := by + induction l₂ + case nil => + simp at heq + case cons h' t' ih' => simp - sorry - - - - + simp [eraseKey] + by_cases heq2 : h'.fst == h.fst + case pos => + simp [heq2] + case neg => + simp at heq2 + simp [heq2] + apply ih' + . rw [List.distinctKeys_cons_iff] at dl₁ + exact dl₁.1 + . rw [containsKey_cons] at heq + simp [heq2] at heq + exact heq + rw [len_eq] + simp only [Nat.add_le_add_iff_right, ge_iff_le] + specialize @ih (eraseKey h.fst l₂) ?refine1 ?refine2 + case refine1 => + apply DistinctKeys.eraseKey dl₁ + case refine2 => + rw [List.distinctKeys_cons_iff] at dl₂ + exact dl₂.1 + rw [← List.countP_eq_length_filter] at ih + apply Nat.le_trans ?refine3 ih + case refine3 => + clear ih len_eq + induction t + case nil => simp + case cons h' t' ih' => + by_cases heq2 : h.fst == h'.fst + case pos => + rw [List.distinctKeys_cons_iff] at dl₂ + replace dl₂ := dl₂.2 + rw [containsKey_cons] at dl₂ + simp at dl₂ + replace dl₂ := dl₂.1 + rw [PartialEquivBEq.symm] at dl₂ + . contradiction + . exact heq2 + case neg => + simp at heq2 + simp only [List.countP_cons] + split + case isTrue contains => + rw [containsKey_eraseKey] + simp [contains, heq2] + apply ih' + . rw [List.distinctKeys_cons_iff, List.distinctKeys_cons_iff, containsKey_cons] at dl₂ + simp at dl₂ + apply DistinctKeys.cons + all_goals simp [dl₂] + . exact dl₁ + case isFalse contains => + rw [containsKey_eraseKey] + simp [heq2, contains] + apply ih' + . rw [List.distinctKeys_cons_iff, List.distinctKeys_cons_iff, containsKey_cons] at dl₂ + simp at dl₂ + apply DistinctKeys.cons + all_goals simp [dl₂] + . exact dl₁ + case neg => + simp only [List.filter_cons, heq, Bool.false_eq_true, ↓reduceIte] + apply ih dl₁ + rw [List.distinctKeys_cons_iff] at dl₂ + exact dl₂.1 theorem isEmpty_filterMap_eq_true [BEq α] [EquivBEq α] {β : Type v} {γ : Type w} {f : (_ : α) → β → Option γ} {l : List ((_ : α) × β)} (distinct : DistinctKeys l) : From 8bf7034e45ce6bf3aa01f2afef1900b128c12aca Mon Sep 17 00:00:00 2001 From: Wojciech Rozowski Date: Fri, 7 Nov 2025 11:10:49 +0000 Subject: [PATCH 20/88] push changes --- src/Std/Data/DHashMap/Internal/RawLemmas.lean | 3 +-- src/Std/Data/Internal/List/Associative.lean | 27 ++++++++++--------- 2 files changed, 15 insertions(+), 15 deletions(-) diff --git a/src/Std/Data/DHashMap/Internal/RawLemmas.lean b/src/Std/Data/DHashMap/Internal/RawLemmas.lean index b64cf9312088..5c126d22a2fd 100644 --- a/src/Std/Data/DHashMap/Internal/RawLemmas.lean +++ b/src/Std/Data/DHashMap/Internal/RawLemmas.lean @@ -2978,8 +2978,7 @@ theorem size_inter_le_size_left [EquivBEq α] [LawfulHashable α] theorem size_inter_le_size_right [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) : (m₁.inter m₂).1.size ≤ m₂.1.size := by - simp_to_model [inter, size, contains] -- using List.length_filter_containsKey_le - sorry + simp_to_model [inter, size, contains] using List.length_filter_containsKey_le theorem size_inter_eq_size_left [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) diff --git a/src/Std/Data/Internal/List/Associative.lean b/src/Std/Data/Internal/List/Associative.lean index 4f86c754e3d9..100b16ac2f10 100644 --- a/src/Std/Data/Internal/List/Associative.lean +++ b/src/Std/Data/Internal/List/Associative.lean @@ -6479,19 +6479,18 @@ theorem length_filter_containsKey_le [BEq α] [EquivBEq α] case nil => simp at heq case cons h' t' ih' => - simp - simp [eraseKey] + simp only [List.length_cons, eraseKey, Nat.add_right_cancel_iff] by_cases heq2 : h'.fst == h.fst case pos => simp [heq2] case neg => - simp at heq2 - simp [heq2] + simp only [Bool.not_eq_true] at heq2 + simp only [heq2, cond_false, List.length_cons] apply ih' . rw [List.distinctKeys_cons_iff] at dl₁ exact dl₁.1 . rw [containsKey_cons] at heq - simp [heq2] at heq + simp only [heq2, Bool.false_or] at heq exact heq rw [len_eq] simp only [Nat.add_le_add_iff_right, ge_iff_le] @@ -6513,32 +6512,34 @@ theorem length_filter_containsKey_le [BEq α] [EquivBEq α] rw [List.distinctKeys_cons_iff] at dl₂ replace dl₂ := dl₂.2 rw [containsKey_cons] at dl₂ - simp at dl₂ + simp only [Bool.or_eq_false_iff] at dl₂ replace dl₂ := dl₂.1 rw [PartialEquivBEq.symm] at dl₂ . contradiction . exact heq2 case neg => - simp at heq2 + simp only [Bool.not_eq_true] at heq2 simp only [List.countP_cons] split case isTrue contains => rw [containsKey_eraseKey] - simp [contains, heq2] + simp only [heq2, Bool.not_false, contains, Bool.and_self, ↓reduceIte, + Nat.add_le_add_iff_right] apply ih' . rw [List.distinctKeys_cons_iff, List.distinctKeys_cons_iff, containsKey_cons] at dl₂ - simp at dl₂ + simp only [Bool.or_eq_false_iff] at dl₂ apply DistinctKeys.cons - all_goals simp [dl₂] + all_goals simp only [dl₂] . exact dl₁ case isFalse contains => rw [containsKey_eraseKey] - simp [heq2, contains] + simp only [Nat.add_zero, heq2, Bool.not_false, contains, Bool.and_false, + Bool.false_eq_true, ↓reduceIte] apply ih' . rw [List.distinctKeys_cons_iff, List.distinctKeys_cons_iff, containsKey_cons] at dl₂ - simp at dl₂ + simp only [Bool.or_eq_false_iff] at dl₂ apply DistinctKeys.cons - all_goals simp [dl₂] + all_goals simp only [dl₂] . exact dl₁ case neg => simp only [List.filter_cons, heq, Bool.false_eq_true, ↓reduceIte] From db949a8858668fc11791cfe7c604ce81b0038db9 Mon Sep 17 00:00:00 2001 From: Wojciech Rozowski Date: Fri, 7 Nov 2025 16:12:48 +0100 Subject: [PATCH 21/88] Finish sublemma --- src/Std/Data/DHashMap/Internal/RawLemmas.lean | 3 +- src/Std/Data/Internal/List/Associative.lean | 284 ++++++++++++------ 2 files changed, 200 insertions(+), 87 deletions(-) diff --git a/src/Std/Data/DHashMap/Internal/RawLemmas.lean b/src/Std/Data/DHashMap/Internal/RawLemmas.lean index 5c126d22a2fd..9c76eb65e023 100644 --- a/src/Std/Data/DHashMap/Internal/RawLemmas.lean +++ b/src/Std/Data/DHashMap/Internal/RawLemmas.lean @@ -2992,8 +2992,7 @@ theorem size_inter_eq_size_right [EquivBEq α] [LawfulHashable α] (h : ∀ (a : α), m₂.contains a → m₁.contains a) : (m₁.inter m₂).1.size = m₂.1.size := by revert h - simp_to_model [inter, size, contains] - sorry + simp_to_model [inter, size, contains] using List.length_filter_containsKey_of_forall_right /- isEmpty -/ @[simp] diff --git a/src/Std/Data/Internal/List/Associative.lean b/src/Std/Data/Internal/List/Associative.lean index 100b16ac2f10..bc6dab231740 100644 --- a/src/Std/Data/Internal/List/Associative.lean +++ b/src/Std/Data/Internal/List/Associative.lean @@ -6067,6 +6067,203 @@ theorem List.length_filter_containsKey_of_forall [BEq α] [EquivBEq α] . apply w rw [containsKey_cons, BEq.rfl, Bool.true_or] +theorem List.length_filter_containsKey_of_forall_right [BEq α] [EquivBEq α] + {l₁ l₂ : List ((a : α) × β a)} (hl₁ : DistinctKeys l₁) (hl₂ : DistinctKeys l₂) + (w : ∀ (a : α), containsKey a l₂ → containsKey a l₁) : + (l₁.filter fun p => containsKey p.fst l₂).length = l₂.length := by + rw [← List.countP_eq_length_filter] + induction l₁ generalizing l₂ + case nil => + simp only [List.countP_nil] + by_cases heq : l₂ = [] + case pos => + simp only [heq, List.length_nil] + case neg => + have ⟨⟨ek, ev⟩, emem⟩ := @List.exists_mem_of_ne_nil _ l₂ (by simp [heq]) + specialize w ek + rw [List.containsKey_eq_true_iff_exists_mem] at w + specialize w ?_ + . refine ⟨⟨ek,ev⟩, emem, ?_⟩ + simp only [BEq.rfl] + . simp at w + case cons h t ih => + by_cases heq : containsKey h.fst l₂ + case pos => + have len_eq : l₂.length = (eraseKey h.fst l₂).length + 1 := by + induction l₂ + case nil => simp at heq + case cons h' t' ih' => + simp only [List.length_cons, eraseKey, Nat.add_right_cancel_iff] + by_cases heq2 : h'.fst == h.fst + case pos => + simp [heq2] + case neg => + simp only [Bool.not_eq_true] at heq2 + simp only [heq2, cond_false, List.length_cons] + apply ih' + . rw [List.distinctKeys_cons_iff] at hl₂ + exact hl₂.1 + . intro a mem + specialize w a + apply w + rw [containsKey_cons] + simp only [mem, Bool.or_true] + . rw [containsKey_cons] at heq + simp only [heq2, Bool.false_or] at heq + exact heq + rw [len_eq] + rw [List.countP_cons] + simp only [heq, ↓reduceIte, Nat.add_right_cancel_iff] + suffices List.countP (fun p => containsKey p.fst l₂) t = List.countP (fun p => containsKey p.fst (eraseKey h.fst l₂)) t by + rw [this] + apply ih + . rw [List.distinctKeys_cons_iff] at hl₁ + exact hl₁.1 + . apply DistinctKeys.eraseKey hl₂ + . intro a + rw [containsKey_eraseKey] + simp + intro neq contains + specialize w a contains + rw [containsKey_cons] at w + . simp only [neq, Bool.false_or] at w + exact w + . exact hl₂ + clear ih w + induction t + case nil => simp + case cons h' t' ih' => + simp only [List.countP_cons] + split + case isTrue heq2 => + rw [containsKey_eraseKey] + . simp only [heq2, Bool.and_true, Bool.not_eq_eq_eq_not, Bool.not_true] + suffices (h.fst == h'.fst) = false by + simp [this] + apply ih' + . rw [List.distinctKeys_cons_iff, List.distinctKeys_cons_iff, containsKey_cons] at hl₁ + simp only [Bool.or_eq_false_iff] at hl₁ + apply DistinctKeys.cons + all_goals simp only [hl₁] + . rw [List.distinctKeys_cons_iff, containsKey_cons] at hl₁ + apply BEq.symm_false + simp only [Bool.or_eq_false_iff] at hl₁ + exact hl₁.2.1 + . exact hl₂ + case isFalse heq2 => + simp only [Bool.not_eq_true] at heq2 + rw [containsKey_eraseKey] + simp only [Nat.add_zero, heq2, Bool.and_false, Bool.false_eq_true, ↓reduceIte] + apply ih' + . rw [List.distinctKeys_cons_iff, List.distinctKeys_cons_iff, containsKey_cons] at hl₁ + simp only [Bool.or_eq_false_iff] at hl₁ + apply DistinctKeys.cons + all_goals simp only [hl₁] + . exact hl₂ + case neg => + simp only [Bool.not_eq_true] at heq + simp only [heq, Bool.false_eq_true, not_false_eq_true, List.countP_cons_of_neg] + apply ih + . rw [List.distinctKeys_cons_iff] at hl₁ + exact hl₁.1 + . exact hl₂ + . intro a mem + specialize w a mem + rw [containsKey_cons] at w + simp at w + rcases w with inl | inr + case w.inl => + rw [containsKey_congr inl] at heq + rw [mem] at heq + contradiction + case w.inr => + exact inr + +theorem length_filter_containsKey_le [BEq α] [EquivBEq α] + {l₁ l₂ : List ((a : α) × β a)} + (dl₁ : DistinctKeys l₂) + (dl₂ : DistinctKeys l₁) : + (l₁.filter fun p => containsKey p.fst l₂).length ≤ l₂.length := by + induction l₁ generalizing l₂ + case nil => simp + case cons h t ih => + by_cases heq : containsKey h.fst l₂ + case pos => + simp [← List.countP_eq_length_filter] + simp only [heq, List.countP_cons_of_pos] + have len_eq : l₂.length = (eraseKey h.fst l₂).length + 1 := by + induction l₂ + case nil => + simp at heq + case cons h' t' ih' => + simp only [List.length_cons, eraseKey, Nat.add_right_cancel_iff] + by_cases heq2 : h'.fst == h.fst + case pos => + simp [heq2] + case neg => + simp only [Bool.not_eq_true] at heq2 + simp only [heq2, cond_false, List.length_cons] + apply ih' + . rw [List.distinctKeys_cons_iff] at dl₁ + exact dl₁.1 + . rw [containsKey_cons] at heq + simp only [heq2, Bool.false_or] at heq + exact heq + rw [len_eq] + simp only [Nat.add_le_add_iff_right, ge_iff_le] + specialize @ih (eraseKey h.fst l₂) ?refine1 ?refine2 + case refine1 => + apply DistinctKeys.eraseKey dl₁ + case refine2 => + rw [List.distinctKeys_cons_iff] at dl₂ + exact dl₂.1 + rw [← List.countP_eq_length_filter] at ih + apply Nat.le_trans ?refine3 ih + case refine3 => + clear ih len_eq + induction t + case nil => simp + case cons h' t' ih' => + by_cases heq2 : h.fst == h'.fst + case pos => + rw [List.distinctKeys_cons_iff] at dl₂ + replace dl₂ := dl₂.2 + rw [containsKey_cons] at dl₂ + simp only [Bool.or_eq_false_iff] at dl₂ + replace dl₂ := dl₂.1 + rw [PartialEquivBEq.symm] at dl₂ + . contradiction + . exact heq2 + case neg => + simp only [Bool.not_eq_true] at heq2 + simp only [List.countP_cons] + split + case isTrue contains => + rw [containsKey_eraseKey] + simp only [heq2, Bool.not_false, contains, Bool.and_self, ↓reduceIte, + Nat.add_le_add_iff_right] + apply ih' + . rw [List.distinctKeys_cons_iff, List.distinctKeys_cons_iff, containsKey_cons] at dl₂ + simp only [Bool.or_eq_false_iff] at dl₂ + apply DistinctKeys.cons + all_goals simp only [dl₂] + . exact dl₁ + case isFalse contains => + rw [containsKey_eraseKey] + simp only [Nat.add_zero, heq2, Bool.not_false, contains, Bool.and_false, + Bool.false_eq_true, ↓reduceIte] + apply ih' + . rw [List.distinctKeys_cons_iff, List.distinctKeys_cons_iff, containsKey_cons] at dl₂ + simp only [Bool.or_eq_false_iff] at dl₂ + apply DistinctKeys.cons + all_goals simp only [dl₂] + . exact dl₁ + case neg => + simp only [List.filter_cons, heq, Bool.false_eq_true, ↓reduceIte] + apply ih dl₁ + rw [List.distinctKeys_cons_iff] at dl₂ + exact dl₂.1 + theorem nil_of_containsKey_eq_false [BEq α] [EquivBEq α] {l : List ((a : α) × β a)} : (∀ k, containsKey k l = false) ↔ l = [] := by constructor @@ -6090,6 +6287,8 @@ theorem nil_of_containsKey_eq_false [BEq α] [EquivBEq α] {l : List ((a : α) intro hyp k simp [hyp] + + theorem isEmpty_filter_containsKey_left [BEq α] [EquivBEq α] {l₁ l₂ : List ((a : α) × β a)} : l₁.isEmpty → (List.filter (fun p => containsKey p.fst l₂) l₁).isEmpty := by @@ -6462,91 +6661,6 @@ theorem length_filter_key_eq_length_iff {β : Type v} [BEq α] [EquivBEq α] simp [← List.filterMap_eq_filter, forall_mem_iff_forall_contains_getKey_getValue (p := fun a b => f a = true) distinct] -theorem length_filter_containsKey_le [BEq α] [EquivBEq α] - {l₁ l₂ : List ((a : α) × β a)} - (dl₁ : DistinctKeys l₂) - (dl₂ : DistinctKeys l₁) : - (l₁.filter fun p => containsKey p.fst l₂).length ≤ l₂.length := by - induction l₁ generalizing l₂ - case nil => simp - case cons h t ih => - by_cases heq : containsKey h.fst l₂ - case pos => - simp [← List.countP_eq_length_filter] - simp only [heq, List.countP_cons_of_pos] - have len_eq : l₂.length = (eraseKey h.fst l₂).length + 1 := by - induction l₂ - case nil => - simp at heq - case cons h' t' ih' => - simp only [List.length_cons, eraseKey, Nat.add_right_cancel_iff] - by_cases heq2 : h'.fst == h.fst - case pos => - simp [heq2] - case neg => - simp only [Bool.not_eq_true] at heq2 - simp only [heq2, cond_false, List.length_cons] - apply ih' - . rw [List.distinctKeys_cons_iff] at dl₁ - exact dl₁.1 - . rw [containsKey_cons] at heq - simp only [heq2, Bool.false_or] at heq - exact heq - rw [len_eq] - simp only [Nat.add_le_add_iff_right, ge_iff_le] - specialize @ih (eraseKey h.fst l₂) ?refine1 ?refine2 - case refine1 => - apply DistinctKeys.eraseKey dl₁ - case refine2 => - rw [List.distinctKeys_cons_iff] at dl₂ - exact dl₂.1 - rw [← List.countP_eq_length_filter] at ih - apply Nat.le_trans ?refine3 ih - case refine3 => - clear ih len_eq - induction t - case nil => simp - case cons h' t' ih' => - by_cases heq2 : h.fst == h'.fst - case pos => - rw [List.distinctKeys_cons_iff] at dl₂ - replace dl₂ := dl₂.2 - rw [containsKey_cons] at dl₂ - simp only [Bool.or_eq_false_iff] at dl₂ - replace dl₂ := dl₂.1 - rw [PartialEquivBEq.symm] at dl₂ - . contradiction - . exact heq2 - case neg => - simp only [Bool.not_eq_true] at heq2 - simp only [List.countP_cons] - split - case isTrue contains => - rw [containsKey_eraseKey] - simp only [heq2, Bool.not_false, contains, Bool.and_self, ↓reduceIte, - Nat.add_le_add_iff_right] - apply ih' - . rw [List.distinctKeys_cons_iff, List.distinctKeys_cons_iff, containsKey_cons] at dl₂ - simp only [Bool.or_eq_false_iff] at dl₂ - apply DistinctKeys.cons - all_goals simp only [dl₂] - . exact dl₁ - case isFalse contains => - rw [containsKey_eraseKey] - simp only [Nat.add_zero, heq2, Bool.not_false, contains, Bool.and_false, - Bool.false_eq_true, ↓reduceIte] - apply ih' - . rw [List.distinctKeys_cons_iff, List.distinctKeys_cons_iff, containsKey_cons] at dl₂ - simp only [Bool.or_eq_false_iff] at dl₂ - apply DistinctKeys.cons - all_goals simp only [dl₂] - . exact dl₁ - case neg => - simp only [List.filter_cons, heq, Bool.false_eq_true, ↓reduceIte] - apply ih dl₁ - rw [List.distinctKeys_cons_iff] at dl₂ - exact dl₂.1 - theorem isEmpty_filterMap_eq_true [BEq α] [EquivBEq α] {β : Type v} {γ : Type w} {f : (_ : α) → β → Option γ} {l : List ((_ : α) × β)} (distinct : DistinctKeys l) : (l.filterMap fun p => (f p.1 p.2).map (fun x => (⟨p.1, x⟩ : (_ : α) × γ))).isEmpty = true ↔ From 8c035db6a156988649d5acfc810db3da5581eecf Mon Sep 17 00:00:00 2001 From: Wojciech Rozowski Date: Fri, 7 Nov 2025 16:28:10 +0100 Subject: [PATCH 22/88] add wf predicate --- src/Std/Data/DHashMap/Internal/RawLemmas.lean | 1 + src/Std/Data/DHashMap/Internal/WF.lean | 5 +---- 2 files changed, 2 insertions(+), 4 deletions(-) diff --git a/src/Std/Data/DHashMap/Internal/RawLemmas.lean b/src/Std/Data/DHashMap/Internal/RawLemmas.lean index 9c76eb65e023..8a9a978d6513 100644 --- a/src/Std/Data/DHashMap/Internal/RawLemmas.lean +++ b/src/Std/Data/DHashMap/Internal/RawLemmas.lean @@ -2994,6 +2994,7 @@ theorem size_inter_eq_size_right [EquivBEq α] [LawfulHashable α] revert h simp_to_model [inter, size, contains] using List.length_filter_containsKey_of_forall_right + /- isEmpty -/ @[simp] theorem isEmpty_inter_left [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) (h : m₁.1.isEmpty): diff --git a/src/Std/Data/DHashMap/Internal/WF.lean b/src/Std/Data/DHashMap/Internal/WF.lean index 68578f235d05..86b4c28084e7 100644 --- a/src/Std/Data/DHashMap/Internal/WF.lean +++ b/src/Std/Data/DHashMap/Internal/WF.lean @@ -1601,10 +1601,7 @@ theorem wf_inter₀ [BEq α] [Hashable α] [LawfulHashable α] (h'₂ : m₂.WF) : (Raw₀.inter ⟨m₁, h₁⟩ ⟨m₂, h₂⟩).1.WF := by rw [inter] - split - . sorry - . sorry - + apply Raw.WF.inter₀ h'₁ h'₂ /-! # `Const.insertListₘ` -/ From ad469502324ad7aac34aadedbb2f1396c2dd1015 Mon Sep 17 00:00:00 2001 From: Wojciech Rozowski Date: Fri, 7 Nov 2025 16:57:26 +0100 Subject: [PATCH 23/88] Add getEntry?/getEntry lemmas --- src/Std/Data/DHashMap/Internal/RawLemmas.lean | 31 ++++++++ src/Std/Data/Internal/List/Associative.lean | 74 +++++++++++++++++++ 2 files changed, 105 insertions(+) diff --git a/src/Std/Data/DHashMap/Internal/RawLemmas.lean b/src/Std/Data/DHashMap/Internal/RawLemmas.lean index 8a9a978d6513..1dcdb98c66bd 100644 --- a/src/Std/Data/DHashMap/Internal/RawLemmas.lean +++ b/src/Std/Data/DHashMap/Internal/RawLemmas.lean @@ -2969,6 +2969,37 @@ theorem getKey!_inter_of_not_contains_left [EquivBEq α] [LawfulHashable α] [In revert h simp_to_model [inter, getKey!, contains] using List.getKeyD_filter_of_not_contains_left +/- getEntry? -/ +theorem getEntry?_inter [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} : + (m₁.inter m₂).getEntry? k = + if m₂.contains k then m₁.getEntry? k else none := by + simp_to_model [inter, getEntry?, contains] using List.getEntry?_filter_containsKey_eq_if_containsKey + +theorem getEntry?_inter_of_contains_right [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) + {k : α} (h : m₂.contains k) : + (m₁.inter m₂).getEntry? k = m₁.getEntry? k := by + revert h + simp_to_model [inter, getEntry?, contains] using List.getEntry?_filter_containsKey_eq_containsKey + +theorem getEntry?_inter_of_not_contains_left [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) + {k : α} (h : m₁.contains k = false) : + (m₁.inter m₂).getEntry? k = none := by + revert h + simp_to_model [inter, getEntry?, contains] using List.getEntry?_filter_of_not_contains + +theorem getEntry?_inter_of_not_contains_right [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) + {k : α} (h : m₂.contains k = false) : + (m₁.inter m₂).getEntry? k = none := by + revert h + simp_to_model [inter, getEntry?, contains] using List.getEntry?_filter_of_not_contains_right + +/- getEntry -/ +theorem getEntry_inter [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) + {k : α} {h_contains : (m₁.inter m₂).contains k} : + (m₁.inter m₂).getEntry k h_contains = + m₁.getEntry k (((@contains_inter_iff α β _ _ m₁ m₂ _ _ h₁ h₂ k).1 h_contains).1) := by + simp_to_model [inter, getEntry, contains] using List.getEntry_filter_containsKey_eq_containsKey + /- size -/ theorem size_inter_le_size_left [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) : diff --git a/src/Std/Data/Internal/List/Associative.lean b/src/Std/Data/Internal/List/Associative.lean index bc6dab231740..0a9c2e4f92c8 100644 --- a/src/Std/Data/Internal/List/Associative.lean +++ b/src/Std/Data/Internal/List/Associative.lean @@ -5254,6 +5254,68 @@ theorem getValueCast?_filter_of_not_contains_right [BEq α] [LawfulBEq α] . simp only [Bool.false_eq_true, ↓reduceIte] . exact dl₁ +theorem getEntry?_filter_containsKey_eq_if_containsKey [BEq α] [EquivBEq α] + {l₁ l₂ : List ((a : α) × β a)} {k : α} + (dl₁ : DistinctKeys l₁) : + getEntry? k (List.filter (fun p => containsKey p.fst l₂) l₁) = + if containsKey k l₂ = true then getEntry? k l₁ else none := by + split + case isTrue h => + rw [getEntry?_filter] + · simp [Option.filter] + split + case h_1 _ val heq => + have key_eq := beq_of_getEntry?_eq_some heq + rw [@containsKey_congr _ _ _ _ l₂ _ _ (PartialEquivBEq.symm key_eq)] at h + simp [h, heq] + case h_2 _ heq => + simp [heq] + · exact dl₁ + case isFalse h => + suffices (getEntry? k (List.filter (fun p => containsKey p.fst l₂) l₁)) = none by + simp [this] + rw [getEntry?_filter] + · simp [Option.filter] + split + case h_1 _ val heq => + have key_eq := beq_of_getEntry?_eq_some heq + simp at h + rw [@containsKey_congr _ _ _ _ l₂ _ _ (PartialEquivBEq.symm key_eq)] at h + simp [h] + case h_2 => rfl + · exact dl₁ + +theorem getEntry?_filter_containsKey_eq_containsKey [BEq α] [EquivBEq α] + {l₁ l₂ : List ((a : α) × β a)} {k : α} + (dl₁ : DistinctKeys l₁) : + containsKey k l₂ → getEntry? k (List.filter (fun p => containsKey p.fst l₂) l₁) = + getEntry? k l₁ := by + intro h + rw [getEntry?_filter_containsKey_eq_if_containsKey, h] + · simp only [↓reduceIte] + · exact dl₁ + +theorem getEntry?_filter_of_not_contains [BEq α] [EquivBEq α] + {l₁ l₂ : List ((a : α) × β a)} {k : α} + (dl₁ : DistinctKeys l₁) : + containsKey k l₁ = false → getEntry? k (List.filter (fun p => containsKey p.fst l₂) l₁) = none := by + intro h + rw [getEntry?_filter_containsKey_eq_if_containsKey] + · split + · apply (@getEntry?_eq_none _ _ _ l₁ k).2 + exact h + · rfl + · exact dl₁ + +theorem getEntry?_filter_of_not_contains_right [BEq α] [EquivBEq α] + {l₁ l₂ : List ((a : α) × β a)} {k : α} + (dl₁ : DistinctKeys l₁) : + containsKey k l₂ = false → getEntry? k (List.filter (fun p => containsKey p.fst l₂) l₁) = none := by + intro h + rw [getEntry?_filter_containsKey_eq_if_containsKey, h] + · simp only [Bool.false_eq_true, ↓reduceIte] + · exact dl₁ + theorem getEntry?_filterMap [BEq α] [EquivBEq α] {f : (a : α) → β a → Option (γ a)} {l : List ((a : α) × β a)} {k : α} (hl : DistinctKeys l) : @@ -5305,6 +5367,18 @@ theorem getValueCast_filter_containsKey_eq_containsKey [BEq α] [LawfulBEq α] . exact h₁.2 . exact dl₁ +theorem getEntry_filter_containsKey_eq_containsKey [BEq α] [EquivBEq α] + {l₁ l₂ : List ((a : α) × β a)} {k : α} + (dl₁ : DistinctKeys l₁) {h₁ h₂} : + getEntry k (List.filter (fun p => containsKey p.fst l₂) l₁) h₁ = getEntry k l₁ h₂ := by + suffices some (getEntry k (List.filter (fun p => containsKey p.fst l₂) l₁) h₁) = some (getEntry k l₁ h₂) by + injections + simp only [← getEntry?_eq_some_getEntry] + apply getEntry?_filter_containsKey_eq_containsKey dl₁ + · rw [contains_filter_iff] at h₁ + · exact h₁.2 + · exact dl₁ + theorem getKey_filter_containsKey_eq_containsKey [BEq α] [EquivBEq α] {l₁ l₂ : List ((a : α) × β a)} {k : α} (dl₁ : DistinctKeys l₁) {h₁ h₂} : From 247fb584f893c87b93036cd52cfc0a0cb8265a89 Mon Sep 17 00:00:00 2001 From: Wojciech Rozowski Date: Mon, 10 Nov 2025 17:16:08 +0100 Subject: [PATCH 24/88] Add getEntry! and getEntryD --- .../DHashMap/Internal/AssocList/Basic.lean | 10 +++ .../DHashMap/Internal/AssocList/Lemmas.lean | 16 ++++ src/Std/Data/DHashMap/Internal/Defs.lean | 14 ++++ src/Std/Data/DHashMap/Internal/Model.lean | 14 ++++ src/Std/Data/DHashMap/Internal/RawLemmas.lean | 6 ++ src/Std/Data/DHashMap/Internal/WF.lean | 20 +++++ src/Std/Data/Internal/List/Associative.lean | 76 ++++++++++++++++++- 7 files changed, 155 insertions(+), 1 deletion(-) diff --git a/src/Std/Data/DHashMap/Internal/AssocList/Basic.lean b/src/Std/Data/DHashMap/Internal/AssocList/Basic.lean index b73ded470597..b1f52e721835 100644 --- a/src/Std/Data/DHashMap/Internal/AssocList/Basic.lean +++ b/src/Std/Data/DHashMap/Internal/AssocList/Basic.lean @@ -132,6 +132,16 @@ def getEntry [BEq α] (a : α) : (l : AssocList α β) → l.contains a → (a : | cons k v es, h => if hka : k == a then ⟨k, v⟩ else es.getEntry a (by rw [← h, contains, Bool.of_not_eq_true hka, Bool.false_or]) +/-- Internal implementation detail of the hash map -/ +def getEntryD [BEq α] (a : α) (fallback : (a : α) × β a) : AssocList α β → (a : α) × β a + | nil => fallback + | cons k v es => if k == a then ⟨k, v⟩ else es.getEntryD a fallback + +/-- Internal implementation detail of the hash map -/ +def getEntry! [BEq α] (a : α) [Inhabited ((a : α) × β a)] : AssocList α β → (a : α) × β a + | nil => default + | cons k v es => if k == a then ⟨k, v⟩ else es.getEntry! a + /-- Internal implementation detail of the hash map -/ def getKey [BEq α] (a : α) : (l : AssocList α β) → l.contains a → α | cons k _ es, h => if hka : k == a then k diff --git a/src/Std/Data/DHashMap/Internal/AssocList/Lemmas.lean b/src/Std/Data/DHashMap/Internal/AssocList/Lemmas.lean index a8abd66ec54f..eab2b81c8e88 100644 --- a/src/Std/Data/DHashMap/Internal/AssocList/Lemmas.lean +++ b/src/Std/Data/DHashMap/Internal/AssocList/Lemmas.lean @@ -98,6 +98,22 @@ theorem getEntry?_eq [BEq α] {l : AssocList α β} {a : α} : next k v t ih => simp only [getEntry?, ih, toList_cons, getEntry?_cons, Bool.ite_eq_cond_iff] +@[simp] +theorem getEntryD_eq [BEq α] {l : AssocList α β} {a : α} {fallback : (a : α) × β a} : + l.getEntryD a fallback = List.getEntryD a fallback l.toList := by + induction l + · simp only [getEntryD, toList_nil, getEntryD_nil] + next k v t ih => + simp only [getEntryD, ih, toList_cons, getEntryD_cons, Bool.ite_eq_cond_iff] + +@[simp] +theorem getEntry!_eq [BEq α] {l : AssocList α β} {a : α} [Inhabited ((a : α) × β a)] : + l.getEntry! a = List.getEntry! a l.toList := by + induction l + · simp only [getEntry!, toList_nil, getEntry!_nil] + next k v t ih => + simp only [getEntry!, ih, toList_cons, List.getEntry!_cons, Bool.ite_eq_cond_iff] + @[simp] theorem getCastD_eq [BEq α] [LawfulBEq α] {l : AssocList α β} {a : α} {fallback : β a} : l.getCastD a fallback = getValueCastD a l.toList fallback := by diff --git a/src/Std/Data/DHashMap/Internal/Defs.lean b/src/Std/Data/DHashMap/Internal/Defs.lean index 15eb9089058d..b776bfc730b6 100644 --- a/src/Std/Data/DHashMap/Internal/Defs.lean +++ b/src/Std/Data/DHashMap/Internal/Defs.lean @@ -382,6 +382,20 @@ def getEntry? [BEq α] [Hashable α] (m : Raw₀ α β) (a : α) : let idx := mkIdx buckets.size h (hash a) buckets[idx.1].getEntry? a +/-- Internal implementation detail of the hash map -/ +def getEntryD [BEq α] [Hashable α] (m : Raw₀ α β) (a : α) (fallback : (a : α) × β a) : + (a : α) × β a := + let ⟨⟨_, buckets⟩, h⟩ := m + let idx := mkIdx buckets.size h (hash a) + buckets[idx.1].getEntryD a fallback + +/-- Internal implementation detail of the hash map -/ +def getEntry! [BEq α] [Hashable α] (m : Raw₀ α β) (a : α) [Inhabited ((a : α) × β a)] : + (a : α) × β a := + let ⟨⟨_, buckets⟩, h⟩ := m + let idx := mkIdx buckets.size h (hash a) + buckets[idx.1].getEntry! a + /-- Internal implementation detail of the hash map -/ def getD [BEq α] [LawfulBEq α] [Hashable α] (m : Raw₀ α β) (a : α) (fallback : β a) : β a := diff --git a/src/Std/Data/DHashMap/Internal/Model.lean b/src/Std/Data/DHashMap/Internal/Model.lean index a65660f96646..001a84f4d303 100644 --- a/src/Std/Data/DHashMap/Internal/Model.lean +++ b/src/Std/Data/DHashMap/Internal/Model.lean @@ -302,6 +302,14 @@ def getEntryₘ [BEq α] [Hashable α] (m : Raw₀ α β) (a : α) (h : m.contai def getEntry?ₘ [BEq α] [Hashable α] (m : Raw₀ α β) (a : α) : Option ((a : α) × β a) := (bucket m.1.buckets m.2 a).getEntry? a +/-- Internal implementation detail of the hash map -/ +def getEntryDₘ [BEq α] [Hashable α] (m : Raw₀ α β) (a : α) (fallback : (a : α) × β a) : (a : α) × β a := + (bucket m.1.buckets m.2 a).getEntryD a fallback + +/-- Internal implementation detail of the hash map -/ +def getEntry!ₘ [BEq α] [Hashable α] [Inhabited ((a : α) × β a)] (m : Raw₀ α β) (a : α) : (a : α) × β a := + (bucket m.1.buckets m.2 a).getEntry! a + /-- Internal implementation detail of the hash map -/ def getDₘ [BEq α] [LawfulBEq α] [Hashable α] (m : Raw₀ α β) (a : α) (fallback : β a) : β a := (m.get?ₘ a).getD fallback @@ -466,6 +474,12 @@ theorem getEntry_eq_getEntryₘ [BEq α] [Hashable α] (m : Raw₀ α β) (a : theorem getEntry?_eq_getEntry?ₘ [BEq α] [Hashable α] (m : Raw₀ α β) (a : α) : getEntry? m a = getEntry?ₘ m a := (rfl) +theorem getEntryD_eq_getEntryDₘ [BEq α] [Hashable α] (m : Raw₀ α β) (a : α) (fallback : (a : α) × β a) : + getEntryD m a fallback = getEntryDₘ m a fallback := (rfl) + +theorem getEntry!_eq_getEntry!ₘ [BEq α] [Hashable α] [Inhabited ((a : α) × β a)] (m : Raw₀ α β) (a : α) : + getEntry! m a = getEntry!ₘ m a := (rfl) + theorem getD_eq_getDₘ [BEq α] [LawfulBEq α] [Hashable α] (m : Raw₀ α β) (a : α) (fallback : β a) : getD m a fallback = getDₘ m a fallback := by simp [getD, getDₘ, get?ₘ, List.getValueCastD_eq_getValueCast?, bucket] diff --git a/src/Std/Data/DHashMap/Internal/RawLemmas.lean b/src/Std/Data/DHashMap/Internal/RawLemmas.lean index b4ce6c6dbb1d..4c2c5d381a80 100644 --- a/src/Std/Data/DHashMap/Internal/RawLemmas.lean +++ b/src/Std/Data/DHashMap/Internal/RawLemmas.lean @@ -142,6 +142,8 @@ private meta def queryMap : Std.DHashMap Name (fun _ => Name × Array (MacroM (T ⟨`getKey!, (``getKey!_eq_getKey!, #[`(getKey!_of_perm _)])⟩, ⟨`getEntry, (``getEntry_eq_getEntry, #[`(getEntry_of_perm _)])⟩, ⟨`getEntry?, (``getEntry?_eq_getEntry?, #[`(getEntry?_of_perm _)])⟩, + ⟨`getEntryD, (``getEntryD_eq_getEntryD, #[`(getEntryD_of_perm _)])⟩, + -- ⟨`getEntry!, (``getEntry!_eq_getEntry!, #[`(getEntry!_of_perm _)])⟩, ⟨`toList, (``Raw.toList_eq_toListModel, #[])⟩, ⟨`keys, (``Raw.keys_eq_keys_toListModel, #[`(perm_keys_congr_left)])⟩, ⟨`Const.toList, (``Raw.Const.toList_eq_toListModel_map, #[`(perm_map_congr_left)])⟩, @@ -290,6 +292,10 @@ theorem containsThenInsertIfNew_snd {k : α} {v : β k} : theorem get?_emptyWithCapacity [LawfulBEq α] {a : α} {c} : (emptyWithCapacity c : Raw₀ α β).get? a = none := by simp [get?] +@[simp] +theorem getEntry?_emptyWithCapacity [LawfulBEq α] {a : α} {c} : (emptyWithCapacity c : Raw₀ α β).getEntry? a = none := by + simp [getEntry?] + theorem get?_of_isEmpty [LawfulBEq α] (h : m.1.WF) {a : α} : m.1.isEmpty = true → m.get? a = none := by simp_to_model [isEmpty, get?]; empty diff --git a/src/Std/Data/DHashMap/Internal/WF.lean b/src/Std/Data/DHashMap/Internal/WF.lean index c1a88f7b07bc..99ce514992e1 100644 --- a/src/Std/Data/DHashMap/Internal/WF.lean +++ b/src/Std/Data/DHashMap/Internal/WF.lean @@ -430,6 +430,26 @@ theorem getEntry?_eq_getEntry? [BEq α] [Hashable α] [PartialEquivBEq α] [Lawf m.getEntry? a = List.getEntry? a (toListModel m.1.buckets) := by rw [getEntry?_eq_getEntry?ₘ, getEntry?ₘ_eq_getEntry? hm] +theorem getEntryDₘ_eq_getEntryD [BEq α] [PartialEquivBEq α] [Hashable α] [LawfulHashable α] {m : Raw₀ α β} (hm : Raw.WFImp m.1) + {a : α} {fallback : (a : α) × β a} : + m.getEntryDₘ a fallback = List.getEntryD a fallback (toListModel m.1.buckets) := + apply_bucket hm AssocList.getEntryD_eq getEntryD_of_perm getEntryD_append_of_containsKey_eq_false + +theorem getEntryD_eq_getEntryD [BEq α] [Hashable α] [PartialEquivBEq α] [LawfulHashable α] {m : Raw₀ α β} (hm : Raw.WFImp m.1) + {a : α} {fallback : (a : α) × β a} : + m.getEntryD a fallback = List.getEntryD a fallback (toListModel m.1.buckets) := by + rw [getEntryD_eq_getEntryDₘ, getEntryDₘ_eq_getEntryD hm] + +theorem getEntry!ₘ_eq_getEntry! [BEq α] [PartialEquivBEq α] [Hashable α] [LawfulHashable α] {m : Raw₀ α β} (hm : Raw.WFImp m.1) + {a : α} [Inhabited ((a : α) × β a)] : + m.getEntry!ₘ a = List.getEntry! a (toListModel m.1.buckets) := + apply_bucket hm AssocList.getEntry!_eq getEntry!_of_perm getEntry!_append_of_containsKey_eq_false + +theorem getEntry!_eq_getEntry! [BEq α] [Hashable α] [PartialEquivBEq α] [LawfulHashable α] {m : Raw₀ α β} (hm : Raw.WFImp m.1) + {a : α} [Inhabited ((a : α) × β a)] : + m.getEntry! a = List.getEntry! a (toListModel m.1.buckets) := by + rw [getEntry!_eq_getEntry!ₘ, getEntry!ₘ_eq_getEntry! hm] + theorem get!ₘ_eq_getValueCast! [BEq α] [Hashable α] [LawfulBEq α] {m : Raw₀ α β} (hm : Raw.WFImp m.1) {a : α} [Inhabited (β a)] : m.get!ₘ a = getValueCast! a (toListModel m.1.buckets) := by diff --git a/src/Std/Data/Internal/List/Associative.lean b/src/Std/Data/Internal/List/Associative.lean index 28feb9d97e15..8715687abf9a 100644 --- a/src/Std/Data/Internal/List/Associative.lean +++ b/src/Std/Data/Internal/List/Associative.lean @@ -48,13 +48,28 @@ def getEntry? [BEq α] (a : α) : List ((a : α) × β a) → Option ((a : α) | [] => none | ⟨k, v⟩ :: l => bif k == a then some ⟨k, v⟩ else getEntry? a l +/-- Internal implementation detail of the hash map -/ +def getEntryD [BEq α] (a : α) (fallback : (a : α) × β a) : List ((a : α) × β a) → (a : α) × β a + | [] => fallback + | ⟨k, v⟩ :: l => bif k == a then ⟨k, v⟩ else getEntryD a fallback l + +/-- Internal implementation detail of the hash map -/ +def getEntry! [BEq α] (a : α) [Inhabited ((a : α) × β a)] : List ((a : α) × β a) → (a : α) × β a + | [] => panic! "key is not present in associative list" + | ⟨k, v⟩ :: l => bif k == a then ⟨k, v⟩ else getEntry! a l + @[simp] theorem getEntry?_nil [BEq α] {a : α} : getEntry? a ([] : List ((a : α) × β a)) = none := (rfl) +@[simp] theorem getEntryD_nil [BEq α] {a : α} (fallback : (a : α) × β a) : + getEntryD a fallback ([] : List ((a : α) × β a)) = fallback := (rfl) + +@[simp] theorem getEntry!_nil [BEq α] {a : α} [Inhabited ((a : α) × β a)] : + getEntry! a ([] : List ((a : α) × β a)) = default := (rfl) + theorem getEntry?_cons [BEq α] {l : List ((a : α) × β a)} {k a : α} {v : β k} : getEntry? a (⟨k, v⟩ :: l) = bif k == a then some ⟨k, v⟩ else getEntry? a l := (rfl) - theorem getEntry?_eq_find [BEq α] {k : α} {l : List ((a : α) × β a)} : getEntry? k l = l.find? (·.1 == k) := by induction l using assoc_induction with @@ -74,6 +89,38 @@ theorem getEntry?_cons_self [BEq α] [ReflBEq α] {l : List ((a : α) × β a)} getEntry? k (⟨k, v⟩ :: l) = some ⟨k, v⟩ := getEntry?_cons_of_true BEq.rfl +theorem getEntryD_cons [BEq α] {l : List ((a : α) × β a)} {k a : α} {v : β k} {fallback : (a : α) × β a} : + getEntryD a fallback (⟨k, v⟩ :: l) = bif k == a then ⟨k, v⟩ else getEntryD a fallback l := (rfl) + +theorem getEntryD_cons_of_true [BEq α] {l : List ((a : α) × β a)} {k a : α} {v : β k} {fallback : (a : α) × β a} + (h : k == a) : getEntryD a fallback (⟨k, v⟩ :: l) = ⟨k, v⟩ := by + simp [getEntryD, h] + +theorem getEntryD_cons_of_false [BEq α] {l : List ((a : α) × β a)} {k a : α} {v : β k} {fallback : (a : α) × β a} + (h : (k == a) = false) : getEntryD a fallback (⟨k, v⟩ :: l) = getEntryD a fallback l := by + simp [getEntryD, h] + +@[simp] +theorem getEntryD_cons_self [BEq α] [ReflBEq α] {l : List ((a : α) × β a)} {k : α} {v : β k} {fallback : (a : α) × β a} : + getEntryD k fallback (⟨k, v⟩ :: l) = ⟨k, v⟩ := + getEntryD_cons_of_true BEq.rfl + +theorem getEntry!_cons [BEq α] {l : List ((a : α) × β a)} {k a : α} {v : β k} [Inhabited ((a : α) × β a)] : + getEntry! a (⟨k, v⟩ :: l) = bif k == a then ⟨k, v⟩ else getEntry! a l := (rfl) + +theorem getEntry!_cons_of_true [BEq α] {l : List ((a : α) × β a)} {k a : α} {v : β k} [Inhabited ((a : α) × β a)] + (h : k == a) : getEntry! a (⟨k, v⟩ :: l) = ⟨k, v⟩ := by + simp [getEntry!, h] + +theorem getEntry!_cons_of_false [BEq α] {l : List ((a : α) × β a)} {k a : α} {v : β k} [Inhabited ((a : α) × β a)] + (h : (k == a) = false) : getEntry! a (⟨k, v⟩ :: l) = getEntry! a l := by + simp [getEntry!, h] + +@[simp] +theorem getEntry!_cons_self [BEq α] [ReflBEq α] {l : List ((a : α) × β a)} {k : α} {v : β k} [Inhabited ((a : α) × β a)] : + getEntry! k (⟨k, v⟩ :: l) = ⟨k, v⟩ := + getEntry!_cons_of_true BEq.rfl + theorem beq_of_getEntry?_eq_some [BEq α] {l : List ((a : α) × β a)} {a : α} {p : (a : α) × β a} (h : getEntry? a l = some p) : p.1 == a := by induction l using assoc_induction @@ -96,6 +143,18 @@ theorem getEntry?_congr [BEq α] [PartialEquivBEq α] {l : List ((a : α) × β rw [getEntry?_cons_of_false h', getEntry?_cons_of_false h₂, ih] · rw [getEntry?_cons_of_true h', getEntry?_cons_of_true (BEq.trans h' h)] +theorem getEntryD_eq_getEntry? [BEq α] {l : List ((a : α) × β a)} {a : α} {fallback : (a : α) × β a} : + getEntryD a fallback l = (getEntry? a l).getD fallback := by + induction l using assoc_induction with + | nil => simp [getEntryD, getEntry?] + | cons k v t ih => cases h : k == a <;> simp_all [getEntryD, getEntry?] + +theorem getEntry!_eq_getEntry? [BEq α] {l : List ((a : α) × β a)} {a : α} [Inhabited ((a : α) × β a)] : + getEntry! a l = (getEntry? a l).get! := by + induction l using assoc_induction with + | nil => rfl + | cons k v t ih => cases h : k == a <;> simp_all [getEntry!, getEntry?] + theorem keys_eq_map {l : List ((a : α) × β a)} : keys l = l.map (·.1) := by induction l with @@ -2172,6 +2231,13 @@ theorem getEntry?_of_perm [BEq α] [PartialEquivBEq α] {l l' : List ((a : α) exact ((Bool.eq_false_iff.1 hl.2.1).elim (BEq.trans h₁ (BEq.symm h₂))).elim next l₁ l₂ l₃ hl₁₂ _ ih₁ ih₂ => exact (ih₁ hl).trans (ih₂ (hl.perm (hl₁₂.symm))) +theorem getEntryD_of_perm [BEq α] [PartialEquivBEq α] {l l' : List ((a : α) × β a)} {a : α} {fallback : (a : α) × β a} + (hl : DistinctKeys l) (h : Perm l l') : getEntryD a fallback l = getEntryD a fallback l' := by + rw [getEntryD_eq_getEntry?, getEntryD_eq_getEntry?, getEntry?_of_perm hl h] + +theorem getEntry!_of_perm [BEq α] [PartialEquivBEq α] {l l' : List ((a : α) × β a)} {a : α} [Inhabited ((a : α) × β a)] + (hl : DistinctKeys l) (h : Perm l l') : getEntry! a l = getEntry! a l' := by + rw [getEntry!_eq_getEntry?, getEntry!_eq_getEntry?, getEntry?_of_perm hl h] theorem containsKey_of_perm [BEq α] [PartialEquivBEq α] {l l' : List ((a : α) × β a)} {k : α} (h : Perm l l') : containsKey k l = containsKey k l' := by @@ -2374,6 +2440,14 @@ theorem getEntry?_append_of_containsKey_eq_false [BEq α] {l l' : List ((a : α) (h : containsKey a l' = false) : getEntry? a (l ++ l') = getEntry? a l := by rw [getEntry?_append, getEntry?_eq_none.2 h, Option.or_none] +theorem getEntryD_append_of_containsKey_eq_false [BEq α] {l l' : List ((a : α) × β a)} {a : α} {fallback : (a : α) × β a} + (h : containsKey a l' = false) : getEntryD a fallback (l ++ l') = getEntryD a fallback l := by + simp [getEntryD_eq_getEntry?, getEntry?_append_of_containsKey_eq_false h] + +theorem getEntry!_append_of_containsKey_eq_false [BEq α] {l l' : List ((a : α) × β a)} {a : α} [Inhabited ((a : α) × β a)] + (h : containsKey a l' = false) : getEntry! a (l ++ l') = getEntry! a l := by + simp [getEntry!_eq_getEntry?, getEntry?_append_of_containsKey_eq_false h] + @[simp] theorem containsKey_append [BEq α] {l l' : List ((a : α) × β a)} {a : α} : containsKey a (l ++ l') = (containsKey a l || containsKey a l') := by From 5b6274db9f0df4e1b5e90051c39cd77ee1a23630 Mon Sep 17 00:00:00 2001 From: Wojciech Rozowski Date: Mon, 10 Nov 2025 17:57:49 +0100 Subject: [PATCH 25/88] fix naming --- src/Std/Data/DHashMap/Internal/RawLemmas.lean | 75 ++++++----- src/Std/Data/Internal/List/Associative.lean | 116 +++++++++--------- 2 files changed, 95 insertions(+), 96 deletions(-) diff --git a/src/Std/Data/DHashMap/Internal/RawLemmas.lean b/src/Std/Data/DHashMap/Internal/RawLemmas.lean index 1dcdb98c66bd..ee47a419f378 100644 --- a/src/Std/Data/DHashMap/Internal/RawLemmas.lean +++ b/src/Std/Data/DHashMap/Internal/RawLemmas.lean @@ -2778,8 +2778,7 @@ end Const section Inter -variable (m₁ m₂ : Raw₀ α β) -variable {m₁ m₂} +variable {m₁ m₂ : Raw₀ α β} /- contains -/ theorem contains_inter [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) @@ -2790,215 +2789,215 @@ theorem contains_inter [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) theorem contains_inter_iff [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} : (m₁.inter m₂).contains k ↔ m₁.contains k ∧ m₂.contains k := by - simp_to_model [inter, contains] using List.contains_filter_iff + simp_to_model [inter, contains] using List.containsKey_filter_containsKey_iff theorem not_contains_inter_of_not_contains_left [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} (h : m₁.contains k = false) : (m₁.inter m₂).contains k = false := by revert h - simp_to_model [inter, contains] using List.not_contains_filter_of_not_contains + simp_to_model [inter, contains] using List.not_contains_filter_containsKey_of_not_containsKey_left theorem not_contains_inter_of_not_contains_right [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} (h : m₂.contains k = false) : (m₁.inter m₂).contains k = false := by revert h - simp_to_model [inter, contains] using List.not_contains_filter_of_not_containsKey_pred + simp_to_model [inter, contains] using List.not_contains_filter_containsKey_of_not_containsKey_right /- get? -/ theorem get?_inter [LawfulBEq α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} : (m₁.inter m₂).get? k = if m₂.contains k then m₁.get? k else none := by - simp_to_model [inter, get?, contains] using List.getValueCast?_filter_containsKey_eq_if_containsKey + simp_to_model [inter, get?, contains] using List.getValueCast?_filter_containsKey theorem get?_inter_of_contains_right [LawfulBEq α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} (h : m₂.contains k) : (m₁.inter m₂).get? k = m₁.get? k := by revert h - simp_to_model [inter, get?, contains] using List.getValueCast?_filter_containsKey_eq_containsKey + simp_to_model [inter, get?, contains] using List.getValueCast?_filter_containsKey_of_containsKey_right theorem get?_inter_of_not_contains_left [LawfulBEq α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} (h : m₁.contains k = false) : (m₁.inter m₂).get? k = none := by revert h - simp_to_model [inter, get?, contains] using List.getValueCast?_filter_of_not_contains + simp_to_model [inter, get?, contains] using List.getValueCast?_filter_containsKey_of_not_containsKey_left theorem get?_inter_of_not_contains_right [LawfulBEq α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} (h : m₂.contains k = false) : (m₁.inter m₂).get? k = none := by revert h - simp_to_model [inter, get?, contains] using List.getValueCast?_filter_of_not_contains_right + simp_to_model [inter, get?, contains] using List.getValueCast?_filter_containsKey_of_not_containsKey_right /- get -/ theorem get_inter [LawfulBEq α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} {h_contains : (m₁.inter m₂).contains k} : (m₁.inter m₂).get k h_contains = m₁.get k (((@contains_inter_iff α β _ _ m₁ m₂ _ _ h₁ h₂ k).1 h_contains).1) := by - simp_to_model [inter, get, contains] using List.getValueCast_filter_containsKey_eq_containsKey + simp_to_model [inter, get, contains] using List.getValueCast_filter_containsKey /- getD -/ theorem getD_inter [LawfulBEq α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} {fallback : β k} : (m₁.inter m₂).getD k fallback = if m₂.contains k then m₁.getD k fallback else fallback := by - simp_to_model [inter, getD, contains] using List.getValueCastD_filter_containsKey_eq_if_containsKey + simp_to_model [inter, getD, contains] using List.getValueCastD_filter_containsKey theorem getD_inter_of_contains_right [LawfulBEq α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} {fallback : β k} (h : m₂.contains k) : (m₁.inter m₂).getD k fallback = m₁.getD k fallback := by revert h - simp_to_model [inter, getD, contains] using List.getValueCastD_filter_containsKey_eq_containsKey + simp_to_model [inter, getD, contains] using List.getValueCastD_filter_containsKey_of_containsKey_right theorem getD_inter_of_not_contains_right [LawfulBEq α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} {fallback : β k} (h : m₂.contains k = false) : (m₁.inter m₂).getD k fallback = fallback := by revert h - simp_to_model [inter, getD, contains] using List.getValueCastD_filter_of_not_contains_right + simp_to_model [inter, getD, contains] using List.getValueCastD_filter_containsKey_of_not_containsKey_right theorem getD_inter_of_not_contains_left [LawfulBEq α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} {fallback : β k} (h : m₁.contains k = false) : (m₁.inter m₂).getD k fallback = fallback := by revert h - simp_to_model [inter, getD, contains] using List.getValueCastD_filter_of_not_contains_left + simp_to_model [inter, getD, contains] using List.getValueCastD_filter_containsKey_of_not_containsKey_left /- get! -/ theorem get!_inter [LawfulBEq α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} [Inhabited (β k)] : (m₁.inter m₂).get! k = if m₂.contains k then m₁.get! k else default := by - simp_to_model [inter, get!, contains] using List.getValueCastD_filter_containsKey_eq_if_containsKey + simp_to_model [inter, get!, contains] using List.getValueCastD_filter_containsKey theorem get!_inter_of_contains_right [LawfulBEq α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} [Inhabited (β k)] (h : m₂.contains k) : (m₁.inter m₂).get! k = m₁.get! k := by revert h - simp_to_model [inter, get!, contains] using List.getValueCastD_filter_containsKey_eq_containsKey + simp_to_model [inter, get!, contains] using List.getValueCastD_filter_containsKey_of_containsKey_right theorem get!_inter_of_not_contains_right [LawfulBEq α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} [Inhabited (β k)] (h : m₂.contains k = false) : (m₁.inter m₂).get! k = default := by revert h - simp_to_model [inter, get!, contains] using List.getValueCastD_filter_of_not_contains_right + simp_to_model [inter, get!, contains] using List.getValueCastD_filter_containsKey_of_not_containsKey_right theorem get!_inter_of_not_contains_left [LawfulBEq α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} [Inhabited (β k)] (h : m₁.contains k = false) : (m₁.inter m₂).get! k = default := by revert h - simp_to_model [inter, get!, contains] using List.getValueCastD_filter_of_not_contains_left + simp_to_model [inter, get!, contains] using List.getValueCastD_filter_containsKey_of_not_containsKey_left /- getKey? -/ theorem getKey?_inter [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} : (m₁.inter m₂).getKey? k = if m₂.contains k then m₁.getKey? k else none := by - simp_to_model [inter, contains, getKey?] using List.getKey?_filter_containsKey_eq_if_containsKey + simp_to_model [inter, contains, getKey?] using List.getKey?_filter_containsKey theorem getKey?_inter_of_contains_right [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} (h : m₂.contains k) : (m₁.inter m₂).getKey? k = m₁.getKey? k := by revert h - simp_to_model [contains, getKey?, inter] using List.getKey?_filter_containsKey_eq_containsKey + simp_to_model [contains, getKey?, inter] using List.getKey?_filter_containsKey_of_containsKey_right theorem getKey?_inter_of_not_contains_right [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} (h : m₂.contains k = false) : (m₁.inter m₂).getKey? k = none := by revert h - simp_to_model [contains, getKey?, inter] using List.getKey?_filter_of_not_contains_right + simp_to_model [contains, getKey?, inter] using List.getKey?_filter_containsKey_of_not_containsKey_right theorem getKey?_inter_of_not_contains_left [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} (h : m₁.contains k = false) : (m₁.inter m₂).getKey? k = none := by revert h - simp_to_model [contains, getKey?, inter] using List.getKey?_filter_of_not_contains + simp_to_model [contains, getKey?, inter] using List.getKey?_filter_containsKey_of_not_contains_left /- getKey -/ theorem getKey_inter [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} {h_contains : (m₁.inter m₂).contains k} : (m₁.inter m₂).getKey k h_contains = m₁.getKey k (by simp [contains_inter_iff h₁ h₂] at h_contains; exact h_contains.1) := by - simp_to_model [inter, contains, getKey] using List.getKey_filter_containsKey_eq_containsKey + simp_to_model [inter, contains, getKey] using List.getKey_filter_containsKey /- getKeyD -/ theorem getKeyD_inter [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k fallback : α} : (m₁.inter m₂).getKeyD k fallback = if m₂.contains k then m₁.getKeyD k fallback else fallback := by - simp_to_model [inter, getKeyD, contains] using List.getKeyD_filter_containsKey_eq_if_containsKey + simp_to_model [inter, getKeyD, contains] using List.getKeyD_filter_containsKey theorem getKeyD_inter_of_contains_right [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k fallback : α} (h : m₂.contains k) : (m₁.inter m₂).getKeyD k fallback = m₁.getKeyD k fallback := by revert h - simp_to_model [inter, getKeyD, contains] using List.getKeyD_filter_containsKey_eq_containsKey + simp_to_model [inter, getKeyD, contains] using List.getKeyD_filter_containsKey_of_containsKey_right theorem getKeyD_inter_of_not_contains_right [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k fallback : α} (h : m₂.contains k = false) : (m₁.inter m₂).getKeyD k fallback = fallback := by revert h - simp_to_model [inter, getKeyD, contains] using List.getKeyD_filter_of_not_contains_right + simp_to_model [inter, getKeyD, contains] using List.getKeyD_filter_containsKey_of_not_containsKey_right theorem getKeyD_inter_of_not_contains_left [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k fallback : α} (h : m₁.contains k = false) : (m₁.inter m₂).getKeyD k fallback = fallback := by revert h - simp_to_model [inter, getKeyD, contains] using List.getKeyD_filter_of_not_contains_left + simp_to_model [inter, getKeyD, contains] using List.getKeyD_filter_containsKey_of_not_containsKey_left /- getKey! -/ theorem getKey!_inter [EquivBEq α] [LawfulHashable α] [Inhabited α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} [Inhabited α] : (m₁.inter m₂).getKey! k = if m₂.contains k then m₁.getKey! k else default := by - simp_to_model [inter, getKey!, contains] using List.getKeyD_filter_containsKey_eq_if_containsKey + simp_to_model [inter, getKey!, contains] using List.getKeyD_filter_containsKey theorem getKey!_inter_of_contains_right [EquivBEq α] [LawfulHashable α] [Inhabited α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} (h : m₂.contains k) : (m₁.inter m₂).getKey! k = m₁.getKey! k := by revert h - simp_to_model [inter, getKey!, contains] using List.getKeyD_filter_containsKey_eq_containsKey + simp_to_model [inter, getKey!, contains] using List.getKeyD_filter_containsKey_of_containsKey_right theorem getKey!_inter_of_not_contains_right [EquivBEq α] [LawfulHashable α] [Inhabited α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} (h : m₂.contains k = false) : (m₁.inter m₂).getKey! k = default := by revert h - simp_to_model [inter, getKey!, contains] using List.getKeyD_filter_of_not_contains_right + simp_to_model [inter, getKey!, contains] using List.getKeyD_filter_containsKey_of_not_containsKey_right theorem getKey!_inter_of_not_contains_left [EquivBEq α] [LawfulHashable α] [Inhabited α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} (h : m₁.contains k = false) : (m₁.inter m₂).getKey! k = default := by revert h - simp_to_model [inter, getKey!, contains] using List.getKeyD_filter_of_not_contains_left + simp_to_model [inter, getKey!, contains] using List.getKeyD_filter_containsKey_of_not_containsKey_left /- getEntry? -/ theorem getEntry?_inter [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} : (m₁.inter m₂).getEntry? k = if m₂.contains k then m₁.getEntry? k else none := by - simp_to_model [inter, getEntry?, contains] using List.getEntry?_filter_containsKey_eq_if_containsKey + simp_to_model [inter, getEntry?, contains] using List.getEntry?_filter_containsKey theorem getEntry?_inter_of_contains_right [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} (h : m₂.contains k) : (m₁.inter m₂).getEntry? k = m₁.getEntry? k := by revert h - simp_to_model [inter, getEntry?, contains] using List.getEntry?_filter_containsKey_eq_containsKey + simp_to_model [inter, getEntry?, contains] using List.getEntry?_filter_containsKey_of_containsKey_right theorem getEntry?_inter_of_not_contains_left [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} (h : m₁.contains k = false) : (m₁.inter m₂).getEntry? k = none := by revert h - simp_to_model [inter, getEntry?, contains] using List.getEntry?_filter_of_not_contains + simp_to_model [inter, getEntry?, contains] using List.getEntry?_filter_containsKey_of_not_containsKey_left theorem getEntry?_inter_of_not_contains_right [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} (h : m₂.contains k = false) : (m₁.inter m₂).getEntry? k = none := by revert h - simp_to_model [inter, getEntry?, contains] using List.getEntry?_filter_of_not_contains_right + simp_to_model [inter, getEntry?, contains] using List.getEntry?_filter_containsKey_of_not_containsKey_right /- getEntry -/ theorem getEntry_inter [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} {h_contains : (m₁.inter m₂).contains k} : (m₁.inter m₂).getEntry k h_contains = m₁.getEntry k (((@contains_inter_iff α β _ _ m₁ m₂ _ _ h₁ h₂ k).1 h_contains).1) := by - simp_to_model [inter, getEntry, contains] using List.getEntry_filter_containsKey_eq_containsKey + simp_to_model [inter, getEntry, contains] using List.getEntry_filter_containsKey /- size -/ theorem size_inter_le_size_left [EquivBEq α] [LawfulHashable α] @@ -3016,14 +3015,14 @@ theorem size_inter_eq_size_left [EquivBEq α] [LawfulHashable α] (h : ∀ (a : α), m₁.contains a → m₂.contains a) : (m₁.inter m₂).1.size = m₁.1.size := by revert h - simp_to_model [inter, size, contains] using List.length_filter_containsKey_of_forall + simp_to_model [inter, size, contains] using List.length_filter_containsKey_eq_length_left theorem size_inter_eq_size_right [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) (h : ∀ (a : α), m₂.contains a → m₁.contains a) : (m₁.inter m₂).1.size = m₂.1.size := by revert h - simp_to_model [inter, size, contains] using List.length_filter_containsKey_of_forall_right + simp_to_model [inter, size, contains] using List.length_filter_containsKey_of_length_right /- isEmpty -/ diff --git a/src/Std/Data/Internal/List/Associative.lean b/src/Std/Data/Internal/List/Associative.lean index 0a9c2e4f92c8..32b43fa89abb 100644 --- a/src/Std/Data/Internal/List/Associative.lean +++ b/src/Std/Data/Internal/List/Associative.lean @@ -5128,7 +5128,7 @@ theorem getEntry?_filter [BEq α] [EquivBEq α] intro p simp only [Option.all_guard, BEq.rfl, Bool.or_true] -theorem getValueCast?_filter_containsKey_eq_if_containsKey [BEq α] [LawfulBEq α] +theorem List.getValueCast?_filter_containsKey [BEq α] [LawfulBEq α] {l₁ l₂ : List ((a : α) × β a)} {k : α} (dl₁ : DistinctKeys l₁) : getValueCast? k (List.filter (fun p => containsKey p.fst l₂) l₁) = @@ -5160,7 +5160,7 @@ theorem getValueCast?_filter_containsKey_eq_if_containsKey [BEq α] [LawfulBEq case h_2 => rfl exact dl₁ -theorem getKey?_filter_containsKey_eq_if_containsKey [BEq α] [EquivBEq α] +theorem List.getKey?_filter_containsKey [BEq α] [EquivBEq α] {l₁ l₂ : List ((a : α) × β a)} {k : α} (dl₁ : DistinctKeys l₁) : getKey? k (List.filter (fun p => containsKey p.fst l₂) l₁) = @@ -5194,67 +5194,67 @@ theorem getKey?_filter_containsKey_eq_if_containsKey [BEq α] [EquivBEq α] case h_2 => rfl · exact dl₁ -theorem getKey?_filter_containsKey_eq_containsKey [BEq α] [EquivBEq α] +theorem List.getKey?_filter_containsKey_of_containsKey_right [BEq α] [EquivBEq α] {l₁ l₂ : List ((a : α) × β a)} {k : α} (dl₁ : DistinctKeys l₁) : containsKey k l₂ → getKey? k (List.filter (fun p => containsKey p.fst l₂) l₁) = getKey? k l₁ := by intro h - rw [getKey?_filter_containsKey_eq_if_containsKey, h] + rw [List.getKey?_filter_containsKey, h] · simp only [↓reduceIte] · exact dl₁ -theorem getKey?_filter_of_not_contains [BEq α] [EquivBEq α] +theorem List.getKey?_filter_containsKey_of_not_contains_left [BEq α] [EquivBEq α] {l₁ l₂ : List ((a : α) × β a)} {k : α} (dl₁ : DistinctKeys l₁) : containsKey k l₁ = false → getKey? k (List.filter (fun p => containsKey p.fst l₂) l₁) = none := by intro h - rw [getKey?_filter_containsKey_eq_if_containsKey] + rw [List.getKey?_filter_containsKey] · split · rw [getKey?_eq_none h] · rfl · exact dl₁ -theorem getKey?_filter_of_not_contains_right [BEq α] [EquivBEq α] +theorem List.getKey?_filter_containsKey_of_not_containsKey_right [BEq α] [EquivBEq α] {l₁ l₂ : List ((a : α) × β a)} {k : α} (dl₁ : DistinctKeys l₁) : containsKey k l₂ = false → getKey? k (List.filter (fun p => containsKey p.fst l₂) l₁) = none := by intro h - rw [getKey?_filter_containsKey_eq_if_containsKey, h] + rw [List.getKey?_filter_containsKey, h] · simp only [Bool.false_eq_true, ↓reduceIte] · exact dl₁ -theorem getValueCast?_filter_containsKey_eq_containsKey [BEq α] [LawfulBEq α] +theorem List.getValueCast?_filter_containsKey_of_containsKey_right [BEq α] [LawfulBEq α] {l₁ l₂ : List ((a : α) × β a)} {k : α} (dl₁ : DistinctKeys l₁) : containsKey k l₂ → getValueCast? k (List.filter (fun p => containsKey p.fst l₂) l₁) = getValueCast? k l₁ := by intro h - rw [getValueCast?_filter_containsKey_eq_if_containsKey, h] + rw [List.getValueCast?_filter_containsKey, h] . simp only [↓reduceIte] . exact dl₁ -theorem getValueCast?_filter_of_not_contains [BEq α] [LawfulBEq α] +theorem List.getValueCast?_filter_containsKey_of_not_containsKey_left [BEq α] [LawfulBEq α] {l₁ l₂ : List ((a : α) × β a)} {k : α} (dl₁ : DistinctKeys l₁) : containsKey k l₁ = false → getValueCast? k (List.filter (fun p => containsKey p.fst l₂) l₁) = none := by intro h - rw [getValueCast?_filter_containsKey_eq_if_containsKey] + rw [List.getValueCast?_filter_containsKey] . split . rw [getValueCast?_eq_none h] . rfl . exact dl₁ -theorem getValueCast?_filter_of_not_contains_right [BEq α] [LawfulBEq α] +theorem List.getValueCast?_filter_containsKey_of_not_containsKey_right [BEq α] [LawfulBEq α] {l₁ l₂ : List ((a : α) × β a)} {k : α} (dl₁ : DistinctKeys l₁) : containsKey k l₂ = false → getValueCast? k (List.filter (fun p => containsKey p.fst l₂) l₁) = none := by intro h - rw [getValueCast?_filter_containsKey_eq_if_containsKey, h] + rw [List.getValueCast?_filter_containsKey, h] . simp only [Bool.false_eq_true, ↓reduceIte] . exact dl₁ -theorem getEntry?_filter_containsKey_eq_if_containsKey [BEq α] [EquivBEq α] +theorem List.getEntry?_filter_containsKey [BEq α] [EquivBEq α] {l₁ l₂ : List ((a : α) × β a)} {k : α} (dl₁ : DistinctKeys l₁) : getEntry? k (List.filter (fun p => containsKey p.fst l₂) l₁) = @@ -5285,34 +5285,34 @@ theorem getEntry?_filter_containsKey_eq_if_containsKey [BEq α] [EquivBEq α] case h_2 => rfl · exact dl₁ -theorem getEntry?_filter_containsKey_eq_containsKey [BEq α] [EquivBEq α] +theorem List.getEntry?_filter_containsKey_of_containsKey_right [BEq α] [EquivBEq α] {l₁ l₂ : List ((a : α) × β a)} {k : α} (dl₁ : DistinctKeys l₁) : containsKey k l₂ → getEntry? k (List.filter (fun p => containsKey p.fst l₂) l₁) = getEntry? k l₁ := by intro h - rw [getEntry?_filter_containsKey_eq_if_containsKey, h] + rw [List.getEntry?_filter_containsKey, h] · simp only [↓reduceIte] · exact dl₁ -theorem getEntry?_filter_of_not_contains [BEq α] [EquivBEq α] +theorem List.getEntry?_filter_containsKey_of_not_containsKey_left [BEq α] [EquivBEq α] {l₁ l₂ : List ((a : α) × β a)} {k : α} (dl₁ : DistinctKeys l₁) : containsKey k l₁ = false → getEntry? k (List.filter (fun p => containsKey p.fst l₂) l₁) = none := by intro h - rw [getEntry?_filter_containsKey_eq_if_containsKey] + rw [List.getEntry?_filter_containsKey] · split · apply (@getEntry?_eq_none _ _ _ l₁ k).2 exact h · rfl · exact dl₁ -theorem getEntry?_filter_of_not_contains_right [BEq α] [EquivBEq α] +theorem List.getEntry?_filter_containsKey_of_not_containsKey_right [BEq α] [EquivBEq α] {l₁ l₂ : List ((a : α) × β a)} {k : α} (dl₁ : DistinctKeys l₁) : containsKey k l₂ = false → getEntry? k (List.filter (fun p => containsKey p.fst l₂) l₁) = none := by intro h - rw [getEntry?_filter_containsKey_eq_if_containsKey, h] + rw [List.getEntry?_filter_containsKey, h] · simp only [Bool.false_eq_true, ↓reduceIte] · exact dl₁ @@ -5349,56 +5349,56 @@ theorem containsKey_filter_containsKey [BEq α] [EquivBEq α] {l₁ l₂ : List apply List.beq_of_getEntry?_eq_some heq . exact hl₁ -theorem contains_filter_iff [BEq α] [EquivBEq α] {l₁ l₂ : List ((a : α) × β a)} {hl₁ : DistinctKeys l₁} {k : α} : +theorem List.containsKey_filter_containsKey_iff [BEq α] [EquivBEq α] {l₁ l₂ : List ((a : α) × β a)} {hl₁ : DistinctKeys l₁} {k : α} : (containsKey k (List.filter (fun p => containsKey p.fst l₂) l₁)) ↔ (containsKey k l₁ ∧ containsKey k l₂) := by rw [containsKey_filter_containsKey] . simp . exact hl₁ -theorem getValueCast_filter_containsKey_eq_containsKey [BEq α] [LawfulBEq α] +theorem List.getValueCast_filter_containsKey [BEq α] [LawfulBEq α] {l₁ l₂ : List ((a : α) × β a)} {k : α} (dl₁ : DistinctKeys l₁) {h₁ h₂} : getValueCast k (List.filter (fun p => containsKey p.fst l₂) l₁) h₁ = getValueCast k l₁ h₂ := by suffices some (getValueCast k (List.filter (fun p => containsKey p.fst l₂) l₁) h₁) = some (getValueCast k l₁ h₂) by injections simp only [← getValueCast?_eq_some_getValueCast] - apply getValueCast?_filter_containsKey_eq_containsKey dl₁ - . rw [contains_filter_iff] at h₁ + apply List.getValueCast?_filter_containsKey_of_containsKey_right dl₁ + . rw [List.containsKey_filter_containsKey_iff] at h₁ . exact h₁.2 . exact dl₁ -theorem getEntry_filter_containsKey_eq_containsKey [BEq α] [EquivBEq α] +theorem List.getEntry_filter_containsKey [BEq α] [EquivBEq α] {l₁ l₂ : List ((a : α) × β a)} {k : α} (dl₁ : DistinctKeys l₁) {h₁ h₂} : getEntry k (List.filter (fun p => containsKey p.fst l₂) l₁) h₁ = getEntry k l₁ h₂ := by suffices some (getEntry k (List.filter (fun p => containsKey p.fst l₂) l₁) h₁) = some (getEntry k l₁ h₂) by injections simp only [← getEntry?_eq_some_getEntry] - apply getEntry?_filter_containsKey_eq_containsKey dl₁ - · rw [contains_filter_iff] at h₁ + apply List.getEntry?_filter_containsKey_of_containsKey_right dl₁ + · rw [List.containsKey_filter_containsKey_iff] at h₁ · exact h₁.2 · exact dl₁ -theorem getKey_filter_containsKey_eq_containsKey [BEq α] [EquivBEq α] +theorem List.getKey_filter_containsKey [BEq α] [EquivBEq α] {l₁ l₂ : List ((a : α) × β a)} {k : α} (dl₁ : DistinctKeys l₁) {h₁ h₂} : getKey k (List.filter (fun p => containsKey p.fst l₂) l₁) h₁ = getKey k l₁ h₂ := by suffices some (getKey k (List.filter (fun p => containsKey p.fst l₂) l₁) h₁) = some (getKey k l₁ h₂) by injections simp only [← getKey?_eq_some_getKey] - apply getKey?_filter_containsKey_eq_containsKey dl₁ - · rw [contains_filter_iff] at h₁ + apply List.getKey?_filter_containsKey_of_containsKey_right dl₁ + · rw [List.containsKey_filter_containsKey_iff] at h₁ · exact h₁.2 · exact dl₁ -theorem not_contains_filter_of_not_contains [BEq α] [EquivBEq α] {l₁ l₂ : List ((a : α) × β a)} {hl₁ : DistinctKeys l₁} {k : α} : +theorem List.not_contains_filter_containsKey_of_not_containsKey_left [BEq α] [EquivBEq α] {l₁ l₂ : List ((a : α) × β a)} {hl₁ : DistinctKeys l₁} {k : α} : containsKey k l₁ = false → containsKey k (List.filter (fun p => containsKey p.fst l₂) l₁) = false := by intro h rw [containsKey_filter_containsKey] . simp [h] . exact hl₁ -theorem not_contains_filter_of_not_containsKey_pred [BEq α] [EquivBEq α] {l₁ l₂ : List ((a : α) × β a)} {hl₁ : DistinctKeys l₁} {k : α} : +theorem List.not_contains_filter_containsKey_of_not_containsKey_right [BEq α] [EquivBEq α] {l₁ l₂ : List ((a : α) × β a)} {hl₁ : DistinctKeys l₁} {k : α} : containsKey k l₂ = false → containsKey k (List.filter (fun p => containsKey p.fst l₂) l₁) = false := by intro h rw [containsKey_filter_containsKey] @@ -5428,7 +5428,7 @@ theorem getValue?_filter_containsKey_eq_if_containsKey {β : Type v} [BEq α] [E case isFalse h => simp only [Option.map_eq_none_iff, getEntry?_eq_none] simp only [Bool.not_eq_true] at h - apply not_contains_filter_of_not_containsKey_pred + apply List.not_contains_filter_containsKey_of_not_containsKey_right . exact dl₁ . exact h @@ -5463,7 +5463,7 @@ theorem getValue?_filter_of_not_contains_right {β : Type v} [BEq α] [EquivBEq · simp only [Bool.false_eq_true, ↓reduceIte] · exact dl₁ -theorem getValueCastD_filter_containsKey_eq_if_containsKey [BEq α] [LawfulBEq α] +theorem List.getValueCastD_filter_containsKey [BEq α] [LawfulBEq α] {l₁ l₂ : List ((a : α) × β a)} {k : α} {fallback : β k} (dl₁ : DistinctKeys l₁) : getValueCastD k (List.filter (fun p => containsKey p.fst l₂) l₁) fallback = @@ -5472,14 +5472,14 @@ theorem getValueCastD_filter_containsKey_eq_if_containsKey [BEq α] [LawfulBEq case isTrue h => rw [getValueCastD_eq_getValueCast?, getValueCastD_eq_getValueCast?] congr 1 - apply getValueCast?_filter_containsKey_eq_containsKey dl₁ h + apply List.getValueCast?_filter_containsKey_of_containsKey_right dl₁ h case isFalse h => apply getValueCastD_eq_fallback - apply not_contains_filter_of_not_containsKey_pred + apply List.not_contains_filter_containsKey_of_not_containsKey_right . exact dl₁ . simp only [h] -theorem getValueCastD_filter_containsKey_eq_containsKey [BEq α] [LawfulBEq α] +theorem List.getValueCastD_filter_containsKey_of_containsKey_right [BEq α] [LawfulBEq α] {l₁ l₂ : List ((a : α) × β a)} {k : α} {fallback: β k} (dl₁ : DistinctKeys l₁) : containsKey k l₂ → getValueCastD k (List.filter (fun p => containsKey p.fst l₂) l₁) fallback = @@ -5487,29 +5487,29 @@ theorem getValueCastD_filter_containsKey_eq_containsKey [BEq α] [LawfulBEq α] intro h rw [getValueCastD_eq_getValueCast?, getValueCastD_eq_getValueCast?] congr 1 - apply getValueCast?_filter_containsKey_eq_containsKey dl₁ h + apply List.getValueCast?_filter_containsKey_of_containsKey_right dl₁ h -theorem getValueCastD_filter_of_not_contains_left [BEq α] [LawfulBEq α] +theorem List.getValueCastD_filter_containsKey_of_not_containsKey_left [BEq α] [LawfulBEq α] {l₁ l₂ : List ((a : α) × β a)} {k : α} {fallback} (dl₁ : DistinctKeys l₁) : containsKey k l₁ = false → getValueCastD k (List.filter (fun p => containsKey p.fst l₂) l₁) fallback = fallback := by intro h apply getValueCastD_eq_fallback - apply not_contains_filter_of_not_contains + apply List.not_contains_filter_containsKey_of_not_containsKey_left . exact dl₁ . exact h -theorem getValueCastD_filter_of_not_contains_right [BEq α] [LawfulBEq α] +theorem List.getValueCastD_filter_containsKey_of_not_containsKey_right [BEq α] [LawfulBEq α] {l₁ l₂ : List ((a : α) × β a)} {k : α} {fallback} (dl₁ : DistinctKeys l₁) : containsKey k l₂ = false → getValueCastD k (List.filter (fun p => containsKey p.fst l₂) l₁) fallback = fallback := by intro h apply getValueCastD_eq_fallback - apply not_contains_filter_of_not_containsKey_pred + apply List.not_contains_filter_containsKey_of_not_containsKey_right . exact dl₁ . exact h -theorem getKeyD_filter_containsKey_eq_if_containsKey [BEq α] [EquivBEq α] +theorem List.getKeyD_filter_containsKey [BEq α] [EquivBEq α] {l₁ l₂ : List ((a : α) × β a)} {k : α} {fallback : α} (dl₁ : DistinctKeys l₁) : getKeyD k (List.filter (fun p => containsKey p.fst l₂) l₁) fallback = @@ -5518,14 +5518,14 @@ theorem getKeyD_filter_containsKey_eq_if_containsKey [BEq α] [EquivBEq α] case isTrue h => rw [getKeyD_eq_getKey?, getKeyD_eq_getKey?] congr 1 - apply getKey?_filter_containsKey_eq_containsKey dl₁ h + apply List.getKey?_filter_containsKey_of_containsKey_right dl₁ h case isFalse h => apply getKeyD_eq_fallback - apply not_contains_filter_of_not_containsKey_pred + apply List.not_contains_filter_containsKey_of_not_containsKey_right · exact dl₁ · simp only [h] -theorem getKeyD_filter_containsKey_eq_containsKey [BEq α] [EquivBEq α] +theorem List.getKeyD_filter_containsKey_of_containsKey_right [BEq α] [EquivBEq α] {l₁ l₂ : List ((a : α) × β a)} {k : α} {fallback : α} (dl₁ : DistinctKeys l₁) : containsKey k l₂ → getKeyD k (List.filter (fun p => containsKey p.fst l₂) l₁) fallback = @@ -5533,25 +5533,25 @@ theorem getKeyD_filter_containsKey_eq_containsKey [BEq α] [EquivBEq α] intro h rw [getKeyD_eq_getKey?, getKeyD_eq_getKey?] congr 1 - apply getKey?_filter_containsKey_eq_containsKey dl₁ h + apply List.getKey?_filter_containsKey_of_containsKey_right dl₁ h -theorem getKeyD_filter_of_not_contains_left [BEq α] [EquivBEq α] +theorem List.getKeyD_filter_containsKey_of_not_containsKey_left [BEq α] [EquivBEq α] {l₁ l₂ : List ((a : α) × β a)} {k : α} {fallback : α} (dl₁ : DistinctKeys l₁) : containsKey k l₁ = false → getKeyD k (List.filter (fun p => containsKey p.fst l₂) l₁) fallback = fallback := by intro h apply getKeyD_eq_fallback - apply not_contains_filter_of_not_contains + apply List.not_contains_filter_containsKey_of_not_containsKey_left · exact dl₁ · exact h -theorem getKeyD_filter_of_not_contains_right [BEq α] [EquivBEq α] +theorem List.getKeyD_filter_containsKey_of_not_containsKey_right [BEq α] [EquivBEq α] {l₁ l₂ : List ((a : α) × β a)} {k : α} {fallback : α} (dl₁ : DistinctKeys l₁) : containsKey k l₂ = false → getKeyD k (List.filter (fun p => containsKey p.fst l₂) l₁) fallback = fallback := by intro h apply getKeyD_eq_fallback - apply not_contains_filter_of_not_containsKey_pred + apply List.not_contains_filter_containsKey_of_not_containsKey_right · exact dl₁ · exact h @@ -5563,7 +5563,7 @@ theorem getValue_filter_containsKey_eq_containsKey {β : Type v} [BEq α] [Equiv injections simp only [← getValue?_eq_some_getValue] apply getValue?_filter_containsKey_eq_containsKey dl₁ - · rw [contains_filter_iff] at h₁ + · rw [List.containsKey_filter_containsKey_iff] at h₁ · exact h₁.2 · exact dl₁ @@ -5579,7 +5579,7 @@ theorem getValueD_filter_containsKey_eq_if_containsKey {β : Type v} [BEq α] [E apply getValue?_filter_containsKey_eq_containsKey dl₁ h case isFalse h => apply getValueD_eq_fallback - apply not_contains_filter_of_not_containsKey_pred + apply List.not_contains_filter_containsKey_of_not_containsKey_right · exact dl₁ · simp only [h] @@ -5599,7 +5599,7 @@ theorem getValueD_filter_of_not_contains_left {β : Type v} [BEq α] [EquivBEq containsKey k l₁ = false → getValueD k (List.filter (fun p => containsKey p.fst l₂) l₁) fallback = fallback := by intro h apply getValueD_eq_fallback - apply not_contains_filter_of_not_contains + apply List.not_contains_filter_containsKey_of_not_containsKey_left · exact dl₁ · exact h @@ -5609,7 +5609,7 @@ theorem getValueD_filter_of_not_contains_right {β : Type v} [BEq α] [EquivBEq containsKey k l₂ = false → getValueD k (List.filter (fun p => containsKey p.fst l₂) l₁) fallback = fallback := by intro h apply getValueD_eq_fallback - apply not_contains_filter_of_not_containsKey_pred + apply List.not_contains_filter_containsKey_of_not_containsKey_right · exact dl₁ · exact h @@ -6120,7 +6120,7 @@ theorem length_filter_key_eq_length_iff [BEq α] [EquivBEq α] {f : α → Bool} simp only [getKey, getKey?_eq_getEntry?, this] at h exact h -theorem List.length_filter_containsKey_of_forall [BEq α] [EquivBEq α] +theorem List.length_filter_containsKey_eq_length_left [BEq α] [EquivBEq α] {l₁ l₂ : List ((a : α) × β a)} (hl₁ : DistinctKeys l₁) (w : ∀ (a : α), containsKey a l₁ → containsKey a l₂) : (l₁.filter fun p => containsKey p.fst l₂).length = l₁.length := by @@ -6141,7 +6141,7 @@ theorem List.length_filter_containsKey_of_forall [BEq α] [EquivBEq α] . apply w rw [containsKey_cons, BEq.rfl, Bool.true_or] -theorem List.length_filter_containsKey_of_forall_right [BEq α] [EquivBEq α] +theorem List.length_filter_containsKey_of_length_right [BEq α] [EquivBEq α] {l₁ l₂ : List ((a : α) × β a)} (hl₁ : DistinctKeys l₁) (hl₂ : DistinctKeys l₂) (w : ∀ (a : α), containsKey a l₂ → containsKey a l₁) : (l₁.filter fun p => containsKey p.fst l₂).length = l₂.length := by From 8a8072941b6d3811bfff3d815597e05df94a92ad Mon Sep 17 00:00:00 2001 From: Wojciech Rozowski Date: Tue, 11 Nov 2025 21:42:24 +0100 Subject: [PATCH 26/88] push what i have --- src/Std/Data/DHashMap/Basic.lean | 2 +- src/Std/Data/DHashMap/Internal/Defs.lean | 4 +-- src/Std/Data/DHashMap/Internal/WF.lean | 38 +++++++++++++++++++----- src/Std/Data/DHashMap/Raw.lean | 29 ++++++------------ src/Std/Data/DHashMap/RawDef.lean | 11 +++++++ 5 files changed, 53 insertions(+), 31 deletions(-) diff --git a/src/Std/Data/DHashMap/Basic.lean b/src/Std/Data/DHashMap/Basic.lean index 88dc4ac067ee..288edeba7ddb 100644 --- a/src/Std/Data/DHashMap/Basic.lean +++ b/src/Std/Data/DHashMap/Basic.lean @@ -327,7 +327,7 @@ This function always merges the smaller map into the larger map, so the expected -/ @[inline] def inter [BEq α] [Hashable α] (m₁ m₂ : DHashMap α β) : DHashMap α β where inner := Raw₀.inter ⟨m₁.1, m₁.2.size_buckets_pos⟩ ⟨m₂.1, m₂.2.size_buckets_pos⟩ - wf := Std.DHashMap.Raw.WF.inter₀ m₁.2 m₂.2 + wf := sorry instance [BEq α] [Hashable α] : Union (DHashMap α β) := ⟨union⟩ diff --git a/src/Std/Data/DHashMap/Internal/Defs.lean b/src/Std/Data/DHashMap/Internal/Defs.lean index a7a0d14181e4..6c1554ae3b55 100644 --- a/src/Std/Data/DHashMap/Internal/Defs.lean +++ b/src/Std/Data/DHashMap/Internal/Defs.lean @@ -464,8 +464,8 @@ def interSmallerFn [BEq α] [Hashable α] (m sofar : Raw₀ α β) (k : α) : Ra | none => sofar /-- Internal implementation detail of the hash map -/ -def interSmaller [BEq α] [Hashable α] (m₁ m₂ : Raw₀ α β) : Raw₀ α β := - (m₂.foldl (fun sofar k _ => interSmallerFn m₁ sofar k) emptyWithCapacity).1 +def interSmaller [BEq α] [Hashable α] (m₁ : Raw₀ α β) (m₂ : Raw α β) : Raw₀ α β := + (m₂.fold (fun sofar k _ => interSmallerFn m₁ sofar k) emptyWithCapacity) /-- Internal implementation detail of the hash map -/ @[inline] def union [BEq α] [Hashable α] (m₁ m₂ : Raw₀ α β) : Raw₀ α β := diff --git a/src/Std/Data/DHashMap/Internal/WF.lean b/src/Std/Data/DHashMap/Internal/WF.lean index 86b4c28084e7..704dd90762e0 100644 --- a/src/Std/Data/DHashMap/Internal/WF.lean +++ b/src/Std/Data/DHashMap/Internal/WF.lean @@ -162,6 +162,23 @@ theorem foldM_eq_foldlM_toListModel {δ : Type w} {m : Type w → Type w'} [Mona funext init' rw [ih] +theorem fold_induction {δ : Type w} + {f : δ → (a : α) → β a → δ} {init : δ} {b : Raw α β} {P : δ → Prop} + (base : P init) (step : ∀ acc a b , P acc → P (f acc a b)) : + P (b.fold f init) := by + simp [Raw.fold, Raw.foldM, ← Array.foldlM_toList] + induction b.buckets.toList generalizing init with + | nil => simp [base] + | cons hd tl ih => + apply ih + induction hd generalizing init with + | nil => simp [AssocList.foldlM, pure, base] + | cons hda hdb tl ih => + simp only [AssocList.foldlM, pure_bind] + apply ih + apply step + exact base + theorem fold_eq_foldl_toListModel {l : Raw α β} {f : γ → (a : α) → β a → γ} {init : γ} : l.fold f init = (toListModel l.buckets).foldl (fun a b => f a b.1 b.2) init := by simp [Raw.fold, foldM_eq_foldlM_toListModel] @@ -1311,7 +1328,7 @@ theorem foldl_eq_foldlₘ {δ} (f : δ → (a : α) → β a → δ) (init : δ) /-- Internal implementation detail of the hash map -/ def interSmallerₘ [BEq α] [Hashable α] (m₁ : Raw₀ α β) (l : List ((a : α) × β a)) : Raw₀ α β := - foldlₘ (fun sofar k _ => interSmallerFnₘ m₁ sofar k) emptyWithCapacity l + l.foldl (fun sofar k => interSmallerFnₘ m₁ sofar k.1) emptyWithCapacity theorem foldl_perm_cong [BEq α] {init₁ init₂ : List ((a : α) × β a)} {l : List ((a : α) × β a)} {f : List ((a : α) × β a) → ((a : α) × β a) → List ((a : α) × β a)} (h₁ : Perm init₁ init₂) @@ -1412,9 +1429,7 @@ theorem toListModel_interSmallerₘ [BEq α] [EquivBEq α] [Hashable α] [Lawful theorem interSmaller_eq_interSmallerₘ [BEq α] [Hashable α] (m₁ m₂ : Raw₀ α β) : interSmaller m₁ m₂ = interSmallerₘ m₁ (toListModel m₂.1.buckets) := by - rw [interSmaller, foldl_eq_foldlₘ] - rw [interSmallerₘ] - rw [Raw.toList_eq_toListModel] + rw [interSmaller, interSmallerₘ, Raw.fold_eq_foldl_toListModel] simp only [interSmallerFn_eq_interSmallerFnₘ] theorem getEntry_foldl [BEq α] [EquivBEq α] (a : α) {acc l₁ l₂ : List ((a : α) × β a)} (hd : DistinctKeys l₁) (hyp : ∀ (a : α) (kv''), List.getEntry? a l₁ = some kv'' → (List.getEntry? a acc = none) ∨ (List.getEntry? a acc = some kv'')): List.getEntry? a @@ -1573,7 +1588,7 @@ theorem toListModel_inter [BEq α] [EquivBEq α] [Hashable α] [LawfulHashable exact hm₂ . apply Perm.trans . rw [interSmaller_eq_interSmallerₘ] - . rw [interSmallerₘ, foldlₘ] + . rw [interSmallerₘ] apply Perm.trans . apply toListModel_interSmallerₘ . exact hm₁ @@ -1597,11 +1612,18 @@ theorem toListModel_inter [BEq α] [EquivBEq α] [Hashable α] [LawfulHashable simp theorem wf_inter₀ [BEq α] [Hashable α] [LawfulHashable α] - {m₁ m₂ : Raw α β} {h₁ : 0 < m₁.buckets.size} {h₂ : 0 < m₂.buckets.size} (h'₁ : m₁.WF) - (h'₂ : m₂.WF) : + {m₁ m₂ : Raw α β} {h₁ : 0 < m₁.buckets.size} {h₂ : 0 < m₂.buckets.size} (wh₁ : m₁.WF) : (Raw₀.inter ⟨m₁, h₁⟩ ⟨m₂, h₂⟩).1.WF := by rw [inter] - apply Raw.WF.inter₀ h'₁ h'₂ + split + . apply Raw.WF.filter₀ wh₁ + . rw [interSmaller] + apply @Raw.fold_induction _ _ _ (fun sofar k x => interSmallerFn ⟨m₁, h₁⟩ sofar k) _ _ (·.val.WF) Raw.WF.emptyWithCapacity₀ + intro acc a b wf + rw [interSmallerFn] + split + . apply Raw.WF.insert₀ wf + . apply wf /-! # `Const.insertListₘ` -/ diff --git a/src/Std/Data/DHashMap/Raw.lean b/src/Std/Data/DHashMap/Raw.lean index e3668f8da573..3d8a01bc91c8 100644 --- a/src/Std/Data/DHashMap/Raw.lean +++ b/src/Std/Data/DHashMap/Raw.lean @@ -349,17 +349,6 @@ This function ensures that the value is used linearly. else ∅ -/-- -Monadically computes a value by folding the given function over the mappings in the hash -map in some order. --/ -@[inline] def foldM (f : δ → (a : α) → β a → m δ) (init : δ) (b : Raw α β) : m δ := - b.buckets.foldlM (fun acc l => l.foldlM f acc) init - -/-- Folds the given function over the mappings in the hash map in some order. -/ -@[inline] def fold (f : δ → (a : α) → β a → δ) (init : δ) (b : Raw α β) : δ := - Id.run (b.foldM (pure <| f · · ·) init) - namespace Internal /-- @@ -733,15 +722,15 @@ theorem WF.inter₀ [BEq α] [Hashable α] {m₁ m₂ : Raw α β} (h₁ : m₁. split . apply WF.filter₀ h₁ . rw [Raw₀.interSmaller] - have := (@Raw₀.foldl α β _ (fun sofar k _ => Raw₀.interSmallerFn ⟨m₁, h₁.size_buckets_pos⟩ sofar k) Raw₀.emptyWithCapacity ⟨m₂, h₂.size_buckets_pos⟩).2 (fun x => x.val.WF) - apply this - . intro ⟨d, hd⟩ k v hw - rw [Raw₀.interSmallerFn] - split - . apply insert₀ - simp [hw] - . exact hw - . apply emptyWithCapacity₀ + sorry + + + + + + + + theorem WF.union [BEq α] [Hashable α] {m₁ m₂ : Raw α β} (h₁ : m₁.WF) (h₂ : m₂.WF) : (m₁.union m₂ : Raw α β).WF := by simp [Std.DHashMap.Raw.union, h₁.size_buckets_pos, h₂.size_buckets_pos] diff --git a/src/Std/Data/DHashMap/RawDef.lean b/src/Std/Data/DHashMap/RawDef.lean index 74ba4c56d695..f17b03f1c739 100644 --- a/src/Std/Data/DHashMap/RawDef.lean +++ b/src/Std/Data/DHashMap/RawDef.lean @@ -56,6 +56,17 @@ namespace Raw variable {α : Type u} {β : α → Type v} {δ : Type w} {m : Type w → Type w'} +/-- +Monadically computes a value by folding the given function over the mappings in the hash +map in some order. +-/ +@[inline] def foldM [Monad m] (f : δ → (a : α) → β a → m δ) (init : δ) (b : Raw α β) : m δ := + b.buckets.foldlM (fun acc l => l.foldlM f acc) init + +/-- Folds the given function over the mappings in the hash map in some order. -/ +@[inline] def fold (f : δ → (a : α) → β a → δ) (init : δ) (b : Raw α β) : δ := + Id.run (b.foldM (pure <| f · · ·) init) + /-- Carries out a monadic action on each mapping in the hash map in some order. -/ @[inline] def forM [Monad m] (f : (a : α) → β a → m PUnit) (b : Raw α β) : m PUnit := b.buckets.forM (AssocList.forM f) From 804fba0e088571a1eefd1097f8047ca68ddea661 Mon Sep 17 00:00:00 2001 From: Wojciech Rozowski Date: Tue, 11 Nov 2025 21:48:09 +0100 Subject: [PATCH 27/88] Cleanup old fold --- src/Std/Data/DHashMap/Internal/Defs.lean | 10 ---------- src/Std/Data/DHashMap/Internal/Model.lean | 4 ---- src/Std/Data/DHashMap/Internal/WF.lean | 19 ------------------- src/Std/Data/DHashMap/Raw.lean | 14 +------------- 4 files changed, 1 insertion(+), 46 deletions(-) diff --git a/src/Std/Data/DHashMap/Internal/Defs.lean b/src/Std/Data/DHashMap/Internal/Defs.lean index 6c1554ae3b55..24cd60b453ae 100644 --- a/src/Std/Data/DHashMap/Internal/Defs.lean +++ b/src/Std/Data/DHashMap/Internal/Defs.lean @@ -447,16 +447,6 @@ def insertMany {ρ : Type w} [ForIn Id ρ ((a : α) × β a)] [BEq α] [Hashable r := ⟨r.1.insertIfNew a b, fun _ h hm => h (r.2 _ h hm)⟩ return r -/-- Internal implementation detail of the hash map -/ -@[inline] def foldl {δ : Type w} (f : δ → (a : α) → β a → δ) (init : δ) - (m : Raw₀ α β) : { d : δ // ∀ (P : δ → Prop), - (∀ {d' a b}, P d' → P (f d' a b)) → P init → P d } := Id.run do - let mut r : { d : δ // ∀ (P : δ → Prop), - (∀ {d' a b}, P d' → P (f d' a b)) → P init → P d } := ⟨init, fun _ _ => id⟩ - for ⟨a, b⟩ in m.1 do - r := ⟨f r.1 a b, fun _ h hm => h (r.2 _ h hm)⟩ - return r - /-- Internal implementation detail of the hash map -/ def interSmallerFn [BEq α] [Hashable α] (m sofar : Raw₀ α β) (k : α) : Raw₀ α β := match m.getEntry? k with diff --git a/src/Std/Data/DHashMap/Internal/Model.lean b/src/Std/Data/DHashMap/Internal/Model.lean index e3a4d1cb91d9..5ced03b87933 100644 --- a/src/Std/Data/DHashMap/Internal/Model.lean +++ b/src/Std/Data/DHashMap/Internal/Model.lean @@ -412,10 +412,6 @@ def unionₘ [BEq α] [Hashable α] (m₁ m₂ : Raw₀ α β) : Raw₀ α β := else insertListₘ m₁ (toListModel m₂.1.buckets) -/-- Internal implementation detail of the hash map -/ -def foldlₘ {δ} (f : δ → (a : α) → β a → δ) (init : δ) (m : List ((a : α) × β a)) : δ := - m.foldl (fun acc ⟨k,v⟩ => f acc k v) init - /-- Internal implementation detail of the hash map -/ def interSmallerFnₘ [BEq α] [Hashable α] (m sofar : Raw₀ α β) (k : α) : Raw₀ α β := match m.getEntry?ₘ k with diff --git a/src/Std/Data/DHashMap/Internal/WF.lean b/src/Std/Data/DHashMap/Internal/WF.lean index 704dd90762e0..dc7367b32905 100644 --- a/src/Std/Data/DHashMap/Internal/WF.lean +++ b/src/Std/Data/DHashMap/Internal/WF.lean @@ -1307,25 +1307,6 @@ theorem toListModel_union [BEq α] [Hashable α] [EquivBEq α] [LawfulHashable rw [union_eq_unionₘ] exact toListModel_unionₘ h₁ h₂ -/-! # `foldl`-/ -theorem foldl_eq_foldlₘ {δ} (f : δ → (a : α) → β a → δ) (init : δ) (m : Raw₀ α β) : - foldl f init m = foldlₘ f init (m.1.toList) := by - simp only [foldl, ForIn.forIn, bind_pure_comp, map_pure, bind_pure, foldlₘ] - rw [Raw.forIn_eq_forIn_toListModel, Raw.toList_eq_toListModel] - generalize (toListModel m.val.buckets) = l - simp only [forIn_pure_yield_eq_foldl, Id.run_pure] - suffices ∀ (t : { d : δ // ∀ (P : δ → Prop), - (∀ {d' : δ} {a : α} {b : β a}, P d' → P (f d' a b)) → P init → P d }), - (List.foldl (fun d' p => ⟨f d'.val p.1 p.2, fun P h₁ h₂ => h₁ (d'.2 _ h₁ h₂)⟩) t l).val = - foldlₘ f t.val l from this _ - intro t - induction l generalizing t with - | nil => simp [foldlₘ] - | cons hd tl ih => - simp only [List.foldl_cons, foldlₘ] - simp only [foldlₘ, Subtype.forall] at ih - apply ih - /-- Internal implementation detail of the hash map -/ def interSmallerₘ [BEq α] [Hashable α] (m₁ : Raw₀ α β) (l : List ((a : α) × β a)) : Raw₀ α β := l.foldl (fun sofar k => interSmallerFnₘ m₁ sofar k.1) emptyWithCapacity diff --git a/src/Std/Data/DHashMap/Raw.lean b/src/Std/Data/DHashMap/Raw.lean index 3d8a01bc91c8..e5f4c19878c0 100644 --- a/src/Std/Data/DHashMap/Raw.lean +++ b/src/Std/Data/DHashMap/Raw.lean @@ -718,19 +718,7 @@ theorem WF.union₀ [BEq α] [Hashable α] {m₁ m₂ : Raw α β} (h₁ : m₁. . exact (Raw₀.insertMany ⟨m₁, h₁.size_buckets_pos⟩ m₂).2 _ WF.insert₀ h₁ theorem WF.inter₀ [BEq α] [Hashable α] {m₁ m₂ : Raw α β} (h₁ : m₁.WF) (h₂ : m₂.WF) : (Raw₀.inter ⟨m₁, h₁.size_buckets_pos⟩ ⟨m₂, h₂.size_buckets_pos⟩).val.WF := by - simp only [Raw₀.inter] - split - . apply WF.filter₀ h₁ - . rw [Raw₀.interSmaller] - sorry - - - - - - - - + sorry theorem WF.union [BEq α] [Hashable α] {m₁ m₂ : Raw α β} (h₁ : m₁.WF) (h₂ : m₂.WF) : (m₁.union m₂ : Raw α β).WF := by simp [Std.DHashMap.Raw.union, h₁.size_buckets_pos, h₂.size_buckets_pos] From 5411489503c801b92a130dbd372caf784825934f Mon Sep 17 00:00:00 2001 From: Wojciech Rozowski Date: Wed, 12 Nov 2025 11:49:03 +0100 Subject: [PATCH 28/88] Add isEmpty lemma --- src/Std/Data/DHashMap/Internal/Defs.lean | 1 + src/Std/Data/DHashMap/Internal/RawLemmas.lean | 14 ++++--- src/Std/Data/Internal/List/Associative.lean | 37 +++++++++++++++++++ 3 files changed, 47 insertions(+), 5 deletions(-) diff --git a/src/Std/Data/DHashMap/Internal/Defs.lean b/src/Std/Data/DHashMap/Internal/Defs.lean index 24cd60b453ae..012f8ab4dab0 100644 --- a/src/Std/Data/DHashMap/Internal/Defs.lean +++ b/src/Std/Data/DHashMap/Internal/Defs.lean @@ -448,6 +448,7 @@ def insertMany {ρ : Type w} [ForIn Id ρ ((a : α) × β a)] [BEq α] [Hashable return r /-- Internal implementation detail of the hash map -/ +@[inline] def interSmallerFn [BEq α] [Hashable α] (m sofar : Raw₀ α β) (k : α) : Raw₀ α β := match m.getEntry? k with | some kv' => sofar.insert kv'.1 kv'.2 diff --git a/src/Std/Data/DHashMap/Internal/RawLemmas.lean b/src/Std/Data/DHashMap/Internal/RawLemmas.lean index ee47a419f378..30fea3a7ca59 100644 --- a/src/Std/Data/DHashMap/Internal/RawLemmas.lean +++ b/src/Std/Data/DHashMap/Internal/RawLemmas.lean @@ -2833,7 +2833,7 @@ theorem get?_inter_of_not_contains_right [LawfulBEq α] (h₁ : m₁.val.WF) (h theorem get_inter [LawfulBEq α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} {h_contains : (m₁.inter m₂).contains k} : (m₁.inter m₂).get k h_contains = - m₁.get k (((@contains_inter_iff α β _ _ m₁ m₂ _ _ h₁ h₂ k).1 h_contains).1) := by + m₁.get k ((contains_inter_iff h₁ h₂).1 h_contains).1 := by simp_to_model [inter, get, contains] using List.getValueCast_filter_containsKey /- getD -/ @@ -2996,7 +2996,7 @@ theorem getEntry?_inter_of_not_contains_right [EquivBEq α] [LawfulHashable α] theorem getEntry_inter [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} {h_contains : (m₁.inter m₂).contains k} : (m₁.inter m₂).getEntry k h_contains = - m₁.getEntry k (((@contains_inter_iff α β _ _ m₁ m₂ _ _ h₁ h₂ k).1 h_contains).1) := by + m₁.getEntry k ((contains_inter_iff h₁ h₂).1 h_contains).1 := by simp_to_model [inter, getEntry, contains] using List.getEntry_filter_containsKey /- size -/ @@ -3024,21 +3024,25 @@ theorem size_inter_eq_size_right [EquivBEq α] [LawfulHashable α] revert h simp_to_model [inter, size, contains] using List.length_filter_containsKey_of_length_right - /- isEmpty -/ @[simp] -theorem isEmpty_inter_left [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) (h : m₁.1.isEmpty): +theorem isEmpty_inter_left [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) (h : m₁.1.isEmpty) : (m₁.inter m₂).1.isEmpty = true := by revert h simp_to_model [isEmpty, inter, contains] using List.isEmpty_filter_containsKey_left /- isEmpty -/ @[simp] -theorem isEmpty_inter_right [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) (h : m₁.1.isEmpty): +theorem isEmpty_inter_right [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) (h : m₁.1.isEmpty) : (m₁.inter m₂).1.isEmpty = true := by revert h simp_to_model [isEmpty, inter, contains] using List.isEmpty_filter_containsKey_left +theorem isEmpty_inter_iff [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) : + (m₁.inter m₂).1.isEmpty ↔ ∀ k, m₁.contains k → m₂.contains k = false := by + simp_to_model [inter, contains, isEmpty] using List.isEmpty_filter_containsKey_iff + + end Inter namespace Const diff --git a/src/Std/Data/Internal/List/Associative.lean b/src/Std/Data/Internal/List/Associative.lean index 32b43fa89abb..a602ae3d9fff 100644 --- a/src/Std/Data/Internal/List/Associative.lean +++ b/src/Std/Data/Internal/List/Associative.lean @@ -6338,6 +6338,43 @@ theorem length_filter_containsKey_le [BEq α] [EquivBEq α] rw [List.distinctKeys_cons_iff] at dl₂ exact dl₂.1 +theorem isEmpty_filter_containsKey_iff [BEq α] [EquivBEq α] + {l₁ l₂ : List ((a : α) × β a)} (dl₁ : DistinctKeys l₁) : + (List.filter (fun p => containsKey p.fst l₂) l₁).isEmpty = true ↔ + ∀ (k : α), containsKey k l₁ = true → containsKey k l₂ = false := by + constructor + case mpr => + intro hyp + rw [List.isEmpty_iff] + simp only [List.filter_eq_nil_iff, Bool.not_eq_true] + intro ⟨k,v⟩ mem + apply hyp + apply List.containsKey_of_mem mem + case mp => + intro hyp k mem + rw [List.isEmpty_iff] at hyp + induction l₁ with + | nil => simp at mem + | cons h t ih => + rw [containsKey_cons] at mem + simp only [Bool.or_eq_true] at mem + rw [List.filter_cons] at hyp + split at hyp + case cons.isTrue _ => + simp at hyp + case cons.isFalse heq => + simp only [Bool.not_eq_true] at heq + cases mem + case inl heq2 => + rw [containsKey_congr (PartialEquivBEq.symm heq2)] + exact heq + case inr heq2 => + apply ih + . rw [List.distinctKeys_cons_iff] at dl₁ + exact dl₁.1 + . exact hyp + . exact heq2 + theorem nil_of_containsKey_eq_false [BEq α] [EquivBEq α] {l : List ((a : α) × β a)} : (∀ k, containsKey k l = false) ↔ l = [] := by constructor From 86f2b74727bdaba5470cfe1444cc6b521f69c132 Mon Sep 17 00:00:00 2001 From: Wojciech Rozowski Date: Wed, 12 Nov 2025 13:21:02 +0100 Subject: [PATCH 29/88] add size lemma --- src/Std/Data/DHashMap/Internal/RawLemmas.lean | 6 ++- src/Std/Data/Internal/List/Associative.lean | 46 +++++++++++++++++++ 2 files changed, 51 insertions(+), 1 deletion(-) diff --git a/src/Std/Data/DHashMap/Internal/RawLemmas.lean b/src/Std/Data/DHashMap/Internal/RawLemmas.lean index 30fea3a7ca59..7617418fd0e2 100644 --- a/src/Std/Data/DHashMap/Internal/RawLemmas.lean +++ b/src/Std/Data/DHashMap/Internal/RawLemmas.lean @@ -3024,6 +3024,11 @@ theorem size_inter_eq_size_right [EquivBEq α] [LawfulHashable α] revert h simp_to_model [inter, size, contains] using List.length_filter_containsKey_of_length_right +theorem size_union_inter_eq_size [EquivBEq α] [LawfulHashable α] + (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) : + m₁.1.size + m₂.1.size = (m₁.union m₂).1.size + (m₁.inter m₂).1.size := by + simp_to_model [union, inter, size] using List.size_insertList_filter_containsKey_eq_size + /- isEmpty -/ @[simp] theorem isEmpty_inter_left [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) (h : m₁.1.isEmpty) : @@ -3042,7 +3047,6 @@ theorem isEmpty_inter_iff [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (m₁.inter m₂).1.isEmpty ↔ ∀ k, m₁.contains k → m₂.contains k = false := by simp_to_model [inter, contains, isEmpty] using List.isEmpty_filter_containsKey_iff - end Inter namespace Const diff --git a/src/Std/Data/Internal/List/Associative.lean b/src/Std/Data/Internal/List/Associative.lean index a602ae3d9fff..d75753c52cdf 100644 --- a/src/Std/Data/Internal/List/Associative.lean +++ b/src/Std/Data/Internal/List/Associative.lean @@ -6253,6 +6253,52 @@ theorem List.length_filter_containsKey_of_length_right [BEq α] [EquivBEq α] case w.inr => exact inr +theorem List.size_insertList_filter_containsKey_eq_size [BEq α] [EquivBEq α] + {l₁ l₂ : List ((a : α) × β a)} (hl₁ : DistinctKeys l₁) (hl₂ : DistinctKeys l₂) : + l₁.length + l₂.length = + (insertList l₁ l₂).length + + (List.filter (fun p => containsKey p.fst l₂) l₁).length := by + rw [← Perm.length_eq (insertListIfNew_perm_insertList hl₂ hl₁)] + induction l₁ generalizing l₂ + case nil => simp [insertListIfNew] + case cons h t ih => + simp [insertListIfNew, List.filter_cons] + split + case isTrue hyp => + simp only [insertEntryIfNew, hyp, cond_true, List.length_cons] + specialize ih (by rw [List.distinctKeys_cons_iff] at hl₁; exact hl₁.1) hl₂ + omega + case isFalse hyp => + rw [Bool.not_eq_true] at hyp + simp [insertEntryIfNew, hyp] + specialize @ih (⟨h.fst, h.snd⟩ :: l₂) (by rw [List.distinctKeys_cons_iff] at hl₁; exact hl₁.1) + rw [List.distinctKeys_cons_iff] at ih + specialize ih ⟨hl₂, hyp⟩ + simp at ih + suffices (List.filter (fun p => h.fst == p.fst || containsKey p.fst l₂) t).length = (List.filter (fun p => containsKey p.fst l₂) t).length by omega + clear ih hl₂ + induction t + case nil => simp + case cons h' t' ih' => + simp [List.filter_cons, Bool.or_eq_true] + by_cases hyp2 : containsKey h'.fst l₂ + case pos => + simp [hyp2] + apply ih' + rw [List.distinctKeys_cons_iff] + rw [List.distinctKeys_cons_iff, List.distinctKeys_cons_iff, List.containsKey_cons] at hl₁ + simp only [Bool.or_eq_false_iff] at hl₁ + exact ⟨hl₁.1.1, hl₁.2.2⟩ + case neg => + simp [hyp2] + simp at hyp2 + rw [List.distinctKeys_cons_iff, containsKey_cons] at hl₁ + simp at hl₁ + simp [BEq.symm_false hl₁.2.1] + apply ih' + rw [List.distinctKeys_cons_iff] at hl₁ |- + exact ⟨hl₁.1.1, hl₁.2.2⟩ + theorem length_filter_containsKey_le [BEq α] [EquivBEq α] {l₁ l₂ : List ((a : α) × β a)} (dl₁ : DistinctKeys l₂) From bae20a2651da9d0ddc557fdffa10bbc61a218e74 Mon Sep 17 00:00:00 2001 From: Wojciech Rozowski Date: Wed, 12 Nov 2025 14:11:58 +0100 Subject: [PATCH 30/88] fix simp_to_model --- src/Std/Data/DHashMap/Internal/RawLemmas.lean | 6 +++--- 1 file changed, 3 insertions(+), 3 deletions(-) diff --git a/src/Std/Data/DHashMap/Internal/RawLemmas.lean b/src/Std/Data/DHashMap/Internal/RawLemmas.lean index 4c2c5d381a80..34004ab3b515 100644 --- a/src/Std/Data/DHashMap/Internal/RawLemmas.lean +++ b/src/Std/Data/DHashMap/Internal/RawLemmas.lean @@ -143,7 +143,7 @@ private meta def queryMap : Std.DHashMap Name (fun _ => Name × Array (MacroM (T ⟨`getEntry, (``getEntry_eq_getEntry, #[`(getEntry_of_perm _)])⟩, ⟨`getEntry?, (``getEntry?_eq_getEntry?, #[`(getEntry?_of_perm _)])⟩, ⟨`getEntryD, (``getEntryD_eq_getEntryD, #[`(getEntryD_of_perm _)])⟩, - -- ⟨`getEntry!, (``getEntry!_eq_getEntry!, #[`(getEntry!_of_perm _)])⟩, + ⟨`getEntry!, (``getEntry!_eq_getEntry!, #[`(getEntry!_of_perm _)])⟩, ⟨`toList, (``Raw.toList_eq_toListModel, #[])⟩, ⟨`keys, (``Raw.keys_eq_keys_toListModel, #[`(perm_keys_congr_left)])⟩, ⟨`Const.toList, (``Raw.Const.toList_eq_toListModel_map, #[`(perm_map_congr_left)])⟩, @@ -2447,10 +2447,10 @@ theorem contains_of_contains_union_of_contains_eq_false_left [EquivBEq α] theorem union_insert_right_equiv_insert_union [EquivBEq α] [LawfulHashable α] {p : (a : α) × β a} (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) : (m₁.union (m₂.insert p.fst p.snd)).1.Equiv ((m₁.union m₂).insert p.fst p.snd).1 := by - simp_to_model [union, insert] + simp_to_model [Equiv, union, insert] apply List.Perm.trans . apply insertList_perm_of_perm_second - simp_to_model [insert] + simp_to_model [Equiv, insert] . apply insertEntry_of_perm . wf_trivial . apply List.Perm.refl From be0de5744685113cee7b280cbb4d7f9bf179c979 Mon Sep 17 00:00:00 2001 From: Wojciech Rozowski Date: Wed, 12 Nov 2025 15:30:48 +0100 Subject: [PATCH 31/88] cleanup --- .../DHashMap/Internal/AssocList/Basic.lean | 1 + src/Std/Data/DHashMap/Internal/RawLemmas.lean | 20 ------------------- src/Std/Data/Internal/List/Associative.lean | 1 - 3 files changed, 1 insertion(+), 21 deletions(-) diff --git a/src/Std/Data/DHashMap/Internal/AssocList/Basic.lean b/src/Std/Data/DHashMap/Internal/AssocList/Basic.lean index b1f52e721835..8bc28c0debd6 100644 --- a/src/Std/Data/DHashMap/Internal/AssocList/Basic.lean +++ b/src/Std/Data/DHashMap/Internal/AssocList/Basic.lean @@ -98,6 +98,7 @@ variable {β : Type v} def get? [BEq α] (a : α) : AssocList α (fun _ => β) → Option β | nil => none | cons k v es => bif k == a then some v else get? a es + end /-- Internal implementation detail of the hash map -/ diff --git a/src/Std/Data/DHashMap/Internal/RawLemmas.lean b/src/Std/Data/DHashMap/Internal/RawLemmas.lean index 34004ab3b515..353ffcbf5e91 100644 --- a/src/Std/Data/DHashMap/Internal/RawLemmas.lean +++ b/src/Std/Data/DHashMap/Internal/RawLemmas.lean @@ -292,10 +292,6 @@ theorem containsThenInsertIfNew_snd {k : α} {v : β k} : theorem get?_emptyWithCapacity [LawfulBEq α] {a : α} {c} : (emptyWithCapacity c : Raw₀ α β).get? a = none := by simp [get?] -@[simp] -theorem getEntry?_emptyWithCapacity [LawfulBEq α] {a : α} {c} : (emptyWithCapacity c : Raw₀ α β).getEntry? a = none := by - simp [getEntry?] - theorem get?_of_isEmpty [LawfulBEq α] (h : m.1.WF) {a : α} : m.1.isEmpty = true → m.get? a = none := by simp_to_model [isEmpty, get?]; empty @@ -384,22 +380,6 @@ theorem get_insert [LawfulBEq α] (h : m.1.WF) {k a : α} {v : β k} {h₁} : m.get a (contains_of_contains_insert _ h h₁ (Bool.eq_false_iff.2 h₂)) := by simp_to_model [insert, get] using List.getValueCast_insertEntry -theorem getEntry_insert [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {k a : α} {v : β k} {h₁} : - (m.insert k v).getEntry a h₁ = - if h₂ : k == a then - ⟨k, v⟩ - else - m.getEntry a (contains_of_contains_insert _ h h₁ (Bool.eq_false_iff.2 h₂)) := by - simp_to_model [insert, getEntry] using List.getEntry_insertEntry - -theorem beq_of_getEntry_eq [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {p : ((a : α) × β a)} {k : α} {h₁} : - (w : m.getEntry k h₁ = p) → p.fst == k := by - simp_to_model [getEntry, contains] using List.beq_of_getEntry_eq - -theorem getEntry_congr [EquivBEq α] [LawfulHashable α] (h : m.1.WF) {a b : α} (hab : a == b) {h₁ h₂} : - getEntry m a h₁ = getEntry m b h₂ := by - simp_to_model [getEntry, contains] using List.getEntry_congr - theorem get_insert_self [LawfulBEq α] (h : m.1.WF) {k : α} {v : β k} : (m.insert k v).get k (contains_insert_self _ h) = v := by simp_to_model [insert, get] using List.getValueCast_insertEntry_self diff --git a/src/Std/Data/Internal/List/Associative.lean b/src/Std/Data/Internal/List/Associative.lean index 8715687abf9a..d88f6fe51f50 100644 --- a/src/Std/Data/Internal/List/Associative.lean +++ b/src/Std/Data/Internal/List/Associative.lean @@ -684,7 +684,6 @@ def getValueCast [BEq α] [LawfulBEq α] (a : α) (l : List ((a : α) × β a)) β a := (getValueCast? a l).get <| containsKey_eq_isSome_getValueCast?.symm.trans h - theorem getValueCast?_eq_some_getValueCast [BEq α] [LawfulBEq α] {l : List ((a : α) × β a)} {a : α} (h : containsKey a l) : getValueCast? a l = some (getValueCast a l h) := by simp [getValueCast] From 0228303d0ca4c33b68b7ec7b7757e6b007217cea Mon Sep 17 00:00:00 2001 From: Wojciech Rozowski Date: Wed, 12 Nov 2025 15:32:17 +0100 Subject: [PATCH 32/88] Further cleanup --- src/Std/Data/DHashMap/Internal/Model.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/src/Std/Data/DHashMap/Internal/Model.lean b/src/Std/Data/DHashMap/Internal/Model.lean index 001a84f4d303..5237e19a4d0d 100644 --- a/src/Std/Data/DHashMap/Internal/Model.lean +++ b/src/Std/Data/DHashMap/Internal/Model.lean @@ -472,7 +472,7 @@ theorem getEntry_eq_getEntryₘ [BEq α] [Hashable α] (m : Raw₀ α β) (a : getEntry m a h = getEntryₘ m a (by exact h) := (rfl) theorem getEntry?_eq_getEntry?ₘ [BEq α] [Hashable α] (m : Raw₀ α β) (a : α) : - getEntry? m a = getEntry?ₘ m a := (rfl) + getEntry? m a = getEntry?ₘ m a := (rfl) theorem getEntryD_eq_getEntryDₘ [BEq α] [Hashable α] (m : Raw₀ α β) (a : α) (fallback : (a : α) × β a) : getEntryD m a fallback = getEntryDₘ m a fallback := (rfl) From d6253e8ba8f5b7befe2d03ea67017d2ed8b1c70c Mon Sep 17 00:00:00 2001 From: Wojciech Rozowski Date: Wed, 12 Nov 2025 15:56:12 +0100 Subject: [PATCH 33/88] add getEntry to publicly exposed DHashMap --- src/Std/Data/DHashMap/Basic.lean | 12 +++++++++++ src/Std/Data/DHashMap/Raw.lean | 37 ++++++++++++++++++++++++++++++++ 2 files changed, 49 insertions(+) diff --git a/src/Std/Data/DHashMap/Basic.lean b/src/Std/Data/DHashMap/Basic.lean index 1316a4a28e06..230d5a1e5c3f 100644 --- a/src/Std/Data/DHashMap/Basic.lean +++ b/src/Std/Data/DHashMap/Basic.lean @@ -190,6 +190,18 @@ end @[inline, inherit_doc Raw.getKeyD] def getKeyD (m : DHashMap α β) (a : α) (fallback : α) : α := Raw₀.getKeyD ⟨m.1, m.2.size_buckets_pos⟩ a fallback +@[inline, inherit_doc Raw.getKey?] def getEntry? (m : DHashMap α β) (a : α) : Option ((a : α) × β a) := + Raw₀.getEntry? ⟨m.1, m.2.size_buckets_pos⟩ a + +@[inline, inherit_doc Raw.getKey] def getEntry (m : DHashMap α β) (a : α) (h : a ∈ m) : (a : α) × β a := + Raw₀.getEntry ⟨m.1, m.2.size_buckets_pos⟩ a h + +@[inline, inherit_doc Raw.getKey!] def getEntry! [Inhabited ((a : α) × β a)] (m : DHashMap α β) (a : α) : (a : α) × β a := + Raw₀.getEntry! ⟨m.1, m.2.size_buckets_pos⟩ a + +@[inline, inherit_doc Raw.getKeyD] def getEntryD (m : DHashMap α β) (a : α) (fallback : (a : α) × β a) : (a : α) × β a := + Raw₀.getEntryD ⟨m.1, m.2.size_buckets_pos⟩ a fallback + @[inline, inherit_doc Raw.size] def size (m : DHashMap α β) : Nat := m.1.size diff --git a/src/Std/Data/DHashMap/Raw.lean b/src/Std/Data/DHashMap/Raw.lean index b0018b935cdf..6b92257a3de1 100644 --- a/src/Std/Data/DHashMap/Raw.lean +++ b/src/Std/Data/DHashMap/Raw.lean @@ -300,6 +300,43 @@ If no panic occurs the result is guaranteed to be pointer equal to the key in th Raw₀.getKey! ⟨m, h⟩ a else default -- will never happen for well-formed inputs +/-- +Checks if a mapping for the given key exists and returns the key-value pair if it does, otherwise `none`. +The result in the `some` case is guaranteed to have the key pointer equal to the key in the map. +-/ +@[inline] def getEntry? [BEq α] [Hashable α] (m : Raw α β) (a : α) : Option ((a : α) × β a) := + if h : 0 < m.buckets.size then + Raw₀.getEntry? ⟨m, h⟩ a + else none -- will never happen for well-formed inputs + +/-- +Retrieves the key-value pair, whose key matches `a`. Ensures that such a mapping exists by +requiring a proof of `a ∈ m`. The key of the result is guaranteed to be pointer equal to the key in the map. +-/ +@[inline] def getEntry [BEq α] [Hashable α] (m : Raw α β) (a : α) (h : a ∈ m) : (a : α) × β a := + Raw₀.getEntry ⟨m, by change dite .. = true at h; split at h <;> simp_all⟩ a + (by change dite .. = true at h; split at h <;> simp_all) + +/-- +Checks if a mapping for the given key exists and returns the key-value pair if it does, otherwise `fallback`. +If a mapping exists the result is guaranteed to have the key pointer equal to the key in the map. +-/ +@[inline] def getEntryD [BEq α] [Hashable α] (m : Raw α β) (a : α) (fallback : (a : α) × β a) : (a : α) × β a := + if h : 0 < m.buckets.size then + Raw₀.getEntryD ⟨m, h⟩ a fallback + else fallback -- will never happen for well-formed inputs + +/-- +Checks if a mapping for the given key exists and returns the key-value pair if it does, otherwise panics. +If no panic occurs the result is guaranteed to have the key pointer equal to the key in the map. +-/ +@[inline] def getEntry! [BEq α] [Hashable α] [Inhabited ((a : α) × β a)] (m : Raw α β) (a : α) : (a : α) × β a := + if h : 0 < m.buckets.size then + Raw₀.getEntry! ⟨m, h⟩ a + else default -- will never happen for well-formed inputs + + + /-- Returns `true` if the hash map contains no mappings. From fb2e980d820ebe5478e285b7c6b71c2a89847449 Mon Sep 17 00:00:00 2001 From: Wojciech Rozowski Date: Wed, 12 Nov 2025 16:33:46 +0100 Subject: [PATCH 34/88] Fix WF issues --- src/Std/Data/DHashMap/Basic.lean | 2 +- src/Std/Data/DHashMap/Internal/WF.lean | 16 +++++++++++++++- src/Std/Data/DHashMap/Raw.lean | 6 +++--- 3 files changed, 19 insertions(+), 5 deletions(-) diff --git a/src/Std/Data/DHashMap/Basic.lean b/src/Std/Data/DHashMap/Basic.lean index e9fa71b6445f..cc378bc77218 100644 --- a/src/Std/Data/DHashMap/Basic.lean +++ b/src/Std/Data/DHashMap/Basic.lean @@ -339,7 +339,7 @@ This function always merges the smaller map into the larger map, so the expected -/ @[inline] def inter [BEq α] [Hashable α] (m₁ m₂ : DHashMap α β) : DHashMap α β where inner := Raw₀.inter ⟨m₁.1, m₁.2.size_buckets_pos⟩ ⟨m₂.1, m₂.2.size_buckets_pos⟩ - wf := sorry + wf := Std.DHashMap.Raw.WF.inter₀ m₁.2 m₂.2 instance [BEq α] [Hashable α] : Union (DHashMap α β) := ⟨union⟩ diff --git a/src/Std/Data/DHashMap/Internal/WF.lean b/src/Std/Data/DHashMap/Internal/WF.lean index 6db0e1dd65ea..c5973425aeba 100644 --- a/src/Std/Data/DHashMap/Internal/WF.lean +++ b/src/Std/Data/DHashMap/Internal/WF.lean @@ -1115,7 +1115,6 @@ theorem insertMany_eq_insertListₘ_toListModel [BEq α] [Hashable α] (m m₂ : simp only [List.foldl_cons, insertListₘ] apply ih - theorem insertManyIfNew_eq_insertListIfNewₘ_toListModel [BEq α] [Hashable α] (m m₂ : Raw₀ α β) : insertManyIfNew m m₂.1 = insertListIfNewₘ m (toListModel m₂.1.buckets) := by simp only [insertManyIfNew, bind_pure_comp, map_pure, bind_pure] @@ -1150,6 +1149,20 @@ theorem toListModel_unionₘ [BEq α] [Hashable α] [EquivBEq α] [LawfulHashabl · exact toListModel_insertListIfNewₘ ‹_› · exact toListModel_insertListₘ ‹_› +theorem wfimp_inter [BEq α] [EquivBEq α] [Hashable α] [LawfulHashable α] + {m₁ m₂ : Raw α β} {h₁ : 0 < m₁.buckets.size} {h₂ : 0 < m₂.buckets.size} (wh₁ : Raw.WFImp m₁) : + Raw.WFImp (Raw₀.inter ⟨m₁, h₁⟩ ⟨m₂, h₂⟩).val := by + rw [inter] + split + . apply wfImp_filter wh₁ + . rw [interSmaller] + apply @Raw.fold_induction _ β _ (fun sofar k x => interSmallerFn ⟨m₁, h₁⟩ sofar k) emptyWithCapacity m₂ (Raw.WFImp ·.val) wfImp_emptyWithCapacity + intro acc a b wf + rw [interSmallerFn] + split + . apply wfImp_insert wf + . apply wf + end Raw₀ namespace Raw @@ -1171,6 +1184,7 @@ theorem WF.out [BEq α] [Hashable α] [i₁ : EquivBEq α] [i₂ : LawfulHashabl | constModify₀ _ h => exact Raw₀.Const.wfImp_modify (by apply h) | alter₀ _ h => exact Raw₀.wfImp_alter (by apply h) | constAlter₀ _ h => exact Raw₀.Const.wfImp_alter (by apply h) + | inter₀ _ _ h _ => exact Raw₀.wfimp_inter (by apply h) end Raw diff --git a/src/Std/Data/DHashMap/Raw.lean b/src/Std/Data/DHashMap/Raw.lean index fc9c82f4122d..fad90ac7d990 100644 --- a/src/Std/Data/DHashMap/Raw.lean +++ b/src/Std/Data/DHashMap/Raw.lean @@ -661,6 +661,8 @@ inductive WF : {α : Type u} → {β : α → Type v} → [BEq α] → [Hashable /-- Internal implementation detail of the hash map -/ | constAlter₀ {α} {β : Type v} [BEq α] [Hashable α] {m : Raw α (fun _ => β)} {h a} {f : Option β → Option β} : WF m → WF (Raw₀.Const.alter ⟨m, h⟩ a f).1 + /-- Internal implementation detail of the hash map -/ + | inter₀ {α β} [BEq α] [Hashable α] {m₁ m₂ : Raw α β} {h₁ h₂} : WF m₁ → WF m₂ → WF (Raw₀.inter ⟨m₁, h₁⟩ ⟨m₂, h₂⟩).1 -- TODO: this needs to be deprecated, but there is a bootstrapping issue. -- @[deprecated WF.emptyWithCapacity₀ (since := "2025-03-12")] @@ -683,6 +685,7 @@ theorem WF.size_buckets_pos [BEq α] [Hashable α] (m : Raw α β) : WF m → 0 | constModify₀ _ => (Raw₀.Const.modify _ _ _).2 | alter₀ _ => (Raw₀.alter _ _ _).2 | constAlter₀ _ => (Raw₀.Const.alter _ _ _).2 + | inter₀ _ _ => (Raw₀.inter _ _).2 @[simp] theorem WF.emptyWithCapacity [BEq α] [Hashable α] {c : Nat} : (Raw.emptyWithCapacity c : Raw α β).WF := .emptyWithCapacity₀ @@ -754,9 +757,6 @@ theorem WF.union₀ [BEq α] [Hashable α] {m₁ m₂ : Raw α β} (h₁ : m₁. . exact (Raw₀.insertManyIfNew ⟨m₂, h₂.size_buckets_pos⟩ m₁).2 _ WF.insertIfNew₀ h₂ . exact (Raw₀.insertMany ⟨m₁, h₁.size_buckets_pos⟩ m₂).2 _ WF.insert₀ h₁ -theorem WF.inter₀ [BEq α] [Hashable α] {m₁ m₂ : Raw α β} (h₁ : m₁.WF) (h₂ : m₂.WF) : (Raw₀.inter ⟨m₁, h₁.size_buckets_pos⟩ ⟨m₂, h₂.size_buckets_pos⟩).val.WF := by - sorry - theorem WF.union [BEq α] [Hashable α] {m₁ m₂ : Raw α β} (h₁ : m₁.WF) (h₂ : m₂.WF) : (m₁.union m₂ : Raw α β).WF := by simp [Std.DHashMap.Raw.union, h₁.size_buckets_pos, h₂.size_buckets_pos] exact WF.union₀ h₁ h₂ From 05e97527ddd61ba26ea83025100aa7fd18cdf149 Mon Sep 17 00:00:00 2001 From: Wojciech Rozowski Date: Wed, 12 Nov 2025 16:51:54 +0100 Subject: [PATCH 35/88] Fix comments --- src/Std/Data/DHashMap/Basic.lean | 8 ++++---- src/Std/Data/DHashMap/Raw.lean | 12 +++++------- 2 files changed, 9 insertions(+), 11 deletions(-) diff --git a/src/Std/Data/DHashMap/Basic.lean b/src/Std/Data/DHashMap/Basic.lean index 230d5a1e5c3f..e341420bcac4 100644 --- a/src/Std/Data/DHashMap/Basic.lean +++ b/src/Std/Data/DHashMap/Basic.lean @@ -190,16 +190,16 @@ end @[inline, inherit_doc Raw.getKeyD] def getKeyD (m : DHashMap α β) (a : α) (fallback : α) : α := Raw₀.getKeyD ⟨m.1, m.2.size_buckets_pos⟩ a fallback -@[inline, inherit_doc Raw.getKey?] def getEntry? (m : DHashMap α β) (a : α) : Option ((a : α) × β a) := +@[inline, inherit_doc Raw.getEntry?] def getEntry? (m : DHashMap α β) (a : α) : Option ((a : α) × β a) := Raw₀.getEntry? ⟨m.1, m.2.size_buckets_pos⟩ a -@[inline, inherit_doc Raw.getKey] def getEntry (m : DHashMap α β) (a : α) (h : a ∈ m) : (a : α) × β a := +@[inline, inherit_doc Raw.getEntry] def getEntry (m : DHashMap α β) (a : α) (h : a ∈ m) : (a : α) × β a := Raw₀.getEntry ⟨m.1, m.2.size_buckets_pos⟩ a h -@[inline, inherit_doc Raw.getKey!] def getEntry! [Inhabited ((a : α) × β a)] (m : DHashMap α β) (a : α) : (a : α) × β a := +@[inline, inherit_doc Raw.getEntry!] def getEntry! [Inhabited ((a : α) × β a)] (m : DHashMap α β) (a : α) : (a : α) × β a := Raw₀.getEntry! ⟨m.1, m.2.size_buckets_pos⟩ a -@[inline, inherit_doc Raw.getKeyD] def getEntryD (m : DHashMap α β) (a : α) (fallback : (a : α) × β a) : (a : α) × β a := +@[inline, inherit_doc Raw.getEntryD] def getEntryD (m : DHashMap α β) (a : α) (fallback : (a : α) × β a) : (a : α) × β a := Raw₀.getEntryD ⟨m.1, m.2.size_buckets_pos⟩ a fallback @[inline, inherit_doc Raw.size] def size (m : DHashMap α β) : Nat := diff --git a/src/Std/Data/DHashMap/Raw.lean b/src/Std/Data/DHashMap/Raw.lean index 6b92257a3de1..bec00f0634fb 100644 --- a/src/Std/Data/DHashMap/Raw.lean +++ b/src/Std/Data/DHashMap/Raw.lean @@ -302,16 +302,16 @@ If no panic occurs the result is guaranteed to be pointer equal to the key in th /-- Checks if a mapping for the given key exists and returns the key-value pair if it does, otherwise `none`. -The result in the `some` case is guaranteed to have the key pointer equal to the key in the map. +The key in the returned pair will be `BEq` to the input `a`. -/ @[inline] def getEntry? [BEq α] [Hashable α] (m : Raw α β) (a : α) : Option ((a : α) × β a) := if h : 0 < m.buckets.size then Raw₀.getEntry? ⟨m, h⟩ a - else none -- will never happen for well-formed inputs + else none -- will never happen for well-formed inputs /-- Retrieves the key-value pair, whose key matches `a`. Ensures that such a mapping exists by -requiring a proof of `a ∈ m`. The key of the result is guaranteed to be pointer equal to the key in the map. +requiring a proof of `a ∈ m`. The key in the returned pair will be `BEq` to the input `a`. -/ @[inline] def getEntry [BEq α] [Hashable α] (m : Raw α β) (a : α) (h : a ∈ m) : (a : α) × β a := Raw₀.getEntry ⟨m, by change dite .. = true at h; split at h <;> simp_all⟩ a @@ -319,7 +319,7 @@ requiring a proof of `a ∈ m`. The key of the result is guaranteed to be pointe /-- Checks if a mapping for the given key exists and returns the key-value pair if it does, otherwise `fallback`. -If a mapping exists the result is guaranteed to have the key pointer equal to the key in the map. +The key in the returned pair will be `BEq` to the input `a`. -/ @[inline] def getEntryD [BEq α] [Hashable α] (m : Raw α β) (a : α) (fallback : (a : α) × β a) : (a : α) × β a := if h : 0 < m.buckets.size then @@ -328,15 +328,13 @@ If a mapping exists the result is guaranteed to have the key pointer equal to th /-- Checks if a mapping for the given key exists and returns the key-value pair if it does, otherwise panics. -If no panic occurs the result is guaranteed to have the key pointer equal to the key in the map. +The key in the returned pair will be `BEq` to the input `a`. -/ @[inline] def getEntry! [BEq α] [Hashable α] [Inhabited ((a : α) × β a)] (m : Raw α β) (a : α) : (a : α) × β a := if h : 0 < m.buckets.size then Raw₀.getEntry! ⟨m, h⟩ a else default -- will never happen for well-formed inputs - - /-- Returns `true` if the hash map contains no mappings. From cfd637230a03efb1bb67dabf17e43dbafc885f31 Mon Sep 17 00:00:00 2001 From: Wojciech Rozowski Date: Wed, 12 Nov 2025 17:05:00 +0100 Subject: [PATCH 36/88] cleanup by applying Markus' suggestion --- src/Std/Data/DHashMap/Internal/WF.lean | 424 +++++++++---------------- 1 file changed, 145 insertions(+), 279 deletions(-) diff --git a/src/Std/Data/DHashMap/Internal/WF.lean b/src/Std/Data/DHashMap/Internal/WF.lean index c5973425aeba..d79ca8fdb296 100644 --- a/src/Std/Data/DHashMap/Internal/WF.lean +++ b/src/Std/Data/DHashMap/Internal/WF.lean @@ -1154,14 +1154,14 @@ theorem wfimp_inter [BEq α] [EquivBEq α] [Hashable α] [LawfulHashable α] Raw.WFImp (Raw₀.inter ⟨m₁, h₁⟩ ⟨m₂, h₂⟩).val := by rw [inter] split - . apply wfImp_filter wh₁ - . rw [interSmaller] + · apply wfImp_filter wh₁ + · rw [interSmaller] apply @Raw.fold_induction _ β _ (fun sofar k x => interSmallerFn ⟨m₁, h₁⟩ sofar k) emptyWithCapacity m₂ (Raw.WFImp ·.val) wfImp_emptyWithCapacity intro acc a b wf rw [interSmallerFn] split - . apply wfImp_insert wf - . apply wf + · apply wfImp_insert wf + · apply wf end Raw₀ @@ -1334,16 +1334,21 @@ theorem wf_union₀ [BEq α] [Hashable α] [EquivBEq α] [LawfulHashable α] · exact wf_insertManyIfNew₀ ‹_› · exact wf_insertMany₀ ‹_› -theorem toListModel_union [BEq α] [Hashable α] [EquivBEq α] [LawfulHashable α] {m₁ m₂ : Raw₀ α β} - (h₁ : Raw.WFImp m₁.1) (h₂ : Raw.WFImp m₂.1) : - Perm (toListModel (m₁.union m₂).1.buckets) - (List.insertList (toListModel m₁.1.buckets) (toListModel m₂.1.buckets)) := by - rw [union_eq_unionₘ] - exact toListModel_unionₘ h₁ h₂ +theorem wfImp_interSmallerFnₘ [BEq α] [EquivBEq α] [Hashable α] [LawfulHashable α] (m₁ : Raw₀ α β) (m₂ : Raw₀ α β) + (hm₂ : Raw.WFImp m₂.1) (k : α) : Raw.WFImp (m₁.interSmallerFnₘ m₂ k).1 := by + rw [interSmallerFnₘ] + split + · exact wfImp_insertₘ hm₂ + · exact hm₂ /-- Internal implementation detail of the hash map -/ -def interSmallerₘ [BEq α] [Hashable α] (m₁ : Raw₀ α β) (l : List ((a : α) × β a)) : Raw₀ α β := - l.foldl (fun sofar k => interSmallerFnₘ m₁ sofar k.1) emptyWithCapacity +def interSmallerₘ [BEq α] [Hashable α] (m₁ : Raw₀ α β) (m₂ : Raw α β) : Raw₀ α β := + m₂.fold (fun sofar k _ => interSmallerFnₘ m₁ sofar k) emptyWithCapacity + +theorem interSmaller_eq_interSmallerₘ [BEq α] [Hashable α] (m₁ : Raw₀ α β) (m₂ : Raw α β) : + m₁.interSmaller m₂ = m₁.interSmallerₘ m₂ := by + rw [interSmaller, interSmallerₘ] + simp only [interSmallerFn_eq_interSmallerFnₘ] theorem foldl_perm_cong [BEq α] {init₁ init₂ : List ((a : α) × β a)} {l : List ((a : α) × β a)} {f : List ((a : α) × β a) → ((a : α) × β a) → List ((a : α) × β a)} (h₁ : Perm init₁ init₂) @@ -1356,289 +1361,150 @@ theorem foldl_perm_cong [BEq α] {init₁ init₂ : List ((a : α) × β a)} {l case cons h t ih => simp only [foldl_cons] apply ih - . exact (h₂ h init₁ init₂ h₃ h₁).1 - . exact (h₂ h init₁ init₂ h₃ h₁).2 + · exact (h₂ h init₁ init₂ h₃ h₁).1 + · exact (h₂ h init₁ init₂ h₃ h₁).2 -theorem toListModel_interSmallerFnₘ [BEq α] [EquivBEq α] [Hashable α] [LawfulHashable α] (m sofar : Raw₀ α β) (hm₁ : Raw.WFImp m.1) (hm₂ : Raw.WFImp sofar.1) (k : α) : - Perm (toListModel ((interSmallerFnₘ m sofar k).1.buckets)) (match List.getEntry? k (toListModel m.1.buckets) with - | some kv' => List.insertEntry kv'.1 kv'.2 (toListModel sofar.1.buckets) - | none => (toListModel sofar.1.buckets)) := by - rw [interSmallerFnₘ] - rw [←getEntry?_eq_getEntry?ₘ, getEntry?_eq_getEntry?] - . split - case h_1 _ kv' heq => - simp only [heq] - apply toListModel_insertₘ - . exact hm₂ - case h_2 _ heq => - simp [heq] - . exact hm₁ - -theorem toListModel_interSmallerₘ [BEq α] [EquivBEq α] [Hashable α] [LawfulHashable α] (m₁ m₂ : Raw₀ α β) (hm₁ : Raw.WFImp m₁.1) (hm₂ : Raw.WFImp m₂.1) (l : List ((a : α) × β a)) : - toListModel - (List.foldl - (fun sofar x => - match x with - | ⟨k, _⟩ => m₁.interSmallerFnₘ sofar k) - m₂ l).val.buckets ~ - List.foldl - (fun sofar x => - match x with - | ⟨k, _⟩ => - match List.getEntry? k (toListModel m₁.val.buckets) with - | some kv' => insertEntry kv'.fst kv'.snd sofar - | none => sofar) - (toListModel m₂.val.buckets) l := by - induction l generalizing m₂ - case nil => - simp - case cons kv t ih => - simp - apply Perm.trans - . apply ih - . unfold interSmallerFnₘ - split - case _ kv' heq => - rw [← insert_eq_insertₘ] - apply wfImp_insert hm₂ - . exact hm₂ - . rw [interSmallerFnₘ] - apply foldl_perm_cong - . split - case h_1 _ kv' heq => - rw [← getEntry?_eq_getEntry?ₘ, getEntry?_eq_getEntry?] at heq - simp only [heq] - apply toListModel_insertₘ - . exact hm₂ - . exact hm₁ - case h_2 _ heq => - rw [← getEntry?_eq_getEntry?ₘ, getEntry?_eq_getEntry?] at heq - simp only [heq] - apply Perm.refl - . exact hm₁ - . intro h l₁ l₂ hd hp - simp - apply And.intro - . split - case h_1 _ kv' heq => - apply insertEntry_of_perm - . exact hd - . exact hp - case h_2 _ heq => - exact hp - . split - case h_1 _ kv' heq => - apply DistinctKeys.insertEntry - . exact hd - case h_2 _ heq => - . exact hd - . split - case h_1 _ kv' heq => - apply DistinctKeys.perm - . apply toListModel_insertₘ - . exact hm₂ - . apply DistinctKeys.insertEntry - . exact hm₂.distinct - case h_2 _ => - exact hm₂.distinct - -theorem interSmaller_eq_interSmallerₘ [BEq α] [Hashable α] (m₁ m₂ : Raw₀ α β) : - interSmaller m₁ m₂ = interSmallerₘ m₁ (toListModel m₂.1.buckets) := by - rw [interSmaller, interSmallerₘ, Raw.fold_eq_foldl_toListModel] - simp only [interSmallerFn_eq_interSmallerFnₘ] - -theorem getEntry_foldl [BEq α] [EquivBEq α] (a : α) {acc l₁ l₂ : List ((a : α) × β a)} (hd : DistinctKeys l₁) (hyp : ∀ (a : α) (kv''), List.getEntry? a l₁ = some kv'' → (List.getEntry? a acc = none) ∨ (List.getEntry? a acc = some kv'')): List.getEntry? a - (List.foldl - (fun sofar x => - match List.getEntry? x.fst l₁ with - | some kv' => insertEntry kv'.fst kv'.snd sofar - | none => sofar) - acc l₂) = List.getEntry? a (acc ++ List.filter (fun p => containsKey p.fst l₂) l₁) := by - induction l₂ generalizing acc - case nil => - simp - suffices (List.getEntry? a (List.filter (fun p => false) l₁)) = .none by simp [this] - have : (List.filter (fun p => false) l₁) = [] := by simp only [filter_eq_nil_iff, - Bool.false_eq_true, not_false_eq_true, implies_true] - simp only [getEntry?_eq_none, this, containsKey_nil] - case cons h t ih => - simp only [foldl_cons, getEntry?_append] - simp at ih - specialize @ih ((match List.getEntry? h.fst l₁ with - | some kv' => insertEntry kv'.fst kv'.snd acc - | none => acc)) - specialize ih ?goal - . intro a kv hkv - split - case h_1 _ kv' heq => - by_cases a == kv'.fst - case pos isTrue => - apply Or.inr - simp [getEntry?_insertEntry, PartialEquivBEq.symm isTrue] - rw [@getEntry?_congr α β _ _ l₁ _ _ (PartialEquivBEq.trans (PartialEquivBEq.symm <| beq_of_getEntry?_eq_some heq) ( PartialEquivBEq.symm <| isTrue))] at heq - rw [heq] at hkv - injections - case neg isFalse => - specialize hyp a kv hkv - apply Or.elim hyp - . intro hyp2 - apply Or.inl - simp only [containsKey_insertEntry, Bool.or_eq_false_iff] - simp only [Bool.not_eq_true] at isFalse - simp [BEq.symm_false isFalse] - rw [containsKey_eq_isSome_getEntry?, hyp2] - simp only [Option.isSome_none] - . intro isSome - apply Or.inr - simp only [Bool.not_eq_true] at isFalse - simp only [getEntry?_insertEntry, BEq.symm_false isFalse, Bool.false_eq_true, - ↓reduceIte, isSome] - case h_2 _ heq => - apply Or.elim <| hyp a kv hkv - . rw [containsKey_eq_isSome_getEntry?] - intro hyp - simp only [hyp, Option.isSome_none, reduceCtorEq, or_false] - . intro hyp - exact Or.inr hyp - rw [ih] - split - case h_1 _ kv' heq => - by_cases ha : a == kv'.fst - case pos => - simp only [getEntry?_insertEntry, PartialEquivBEq.symm ha, ↓reduceIte, Option.some_or] - have heq2 := beq_of_getEntry?_eq_some heq - specialize hyp a kv' (by rw [getEntry?_congr ha, getEntry?_congr heq2]; exact heq) - apply Or.elim hyp - . intro isNone - simp [isNone] - rw [getEntry?_filter, getEntry?_congr ha, getEntry?_congr heq2, heq] - simp only [Option.filter, containsKey, PartialEquivBEq.symm heq2, Bool.true_or, - ↓reduceIte] - . exact hd - . intro hyp - simp [hyp] - case neg => - simp at ha - simp [getEntry?_insertEntry, BEq.symm_false ha] - rw [getEntry?_filter, getEntry?_filter] - generalize heq3 : List.getEntry? a l₁ = x - cases x - case none => - simp - case some kv'' => - simp only [Option.filter, containsKey, Bool.or_eq_true] - split - case isTrue heq4 => - simp only [Option.or_some, heq4, or_true, ↓reduceIte] - case isFalse heq4 => - simp only [Option.or_none, - BEq.symm_false - (BEq.neq_of_beq_of_neq (beq_of_getEntry?_eq_some heq3) - (BEq.neq_of_neq_of_beq ha (beq_of_getEntry?_eq_some heq))), - Bool.false_eq_true, heq4, or_self, ↓reduceIte] - . exact hd - . exact hd - case h_2 _ heq => - congr 1 - rw [getEntry?_filter, getEntry?_filter] - generalize heq2 : List.getEntry? a l₁ = x - cases x - case e_a.none => - simp [Option.filter] - case e_a.some kv' => - simp only [Option.filter, containsKey, Bool.or_eq_true] - congr - simp only [eq_iff_iff, iff_or_self] - intro hyp - have heq3 := beq_of_getEntry?_eq_some heq2 - have heq4 : h.fst == a := by - apply PartialEquivBEq.trans - . exact hyp - . exact heq3 - rw [@getEntry?_congr α β _ _ l₁ h.fst a heq4] at heq - rw [heq] at heq2 - simp_all only [reduceCtorEq] - . exact hd - . exact hd - -theorem foldl_insertEntry_distinctKeys [BEq α] [PartialEquivBEq α] (l₁ l₂ : List ((a : α) × β a)) (acc : List ((a : α) × β a)) - (h_acc : DistinctKeys acc) : - DistinctKeys - (List.foldl - (fun sofar x => - match x with - | ⟨k, _⟩ => - match List.getEntry? k l₁ with - | some kv' => insertEntry kv'.fst kv'.snd sofar - | none => sofar) - acc l₂) := by - induction l₂ generalizing acc with - | nil => - simp [List.foldl] - exact h_acc - | cons h t ih => - simp only [List.foldl_cons] - cases heq : List.getEntry? h.fst l₁ with - | none => - apply ih - exact h_acc - | some kv' => - apply ih - split - case h_1 _ kv' heq => - apply DistinctKeys.insertEntry - exact h_acc - case h_2 => exact h_acc +/-- Internal implementation detail of the hash map -/ +def List.interSmallerFn [BEq α] (l sofar : List ((a : α) × β a)) (k : α) : List ((a : α) × β a) := + match List.getEntry? k l with + | some kv' => List.insertEntry kv'.1 kv'.2 sofar + | none => sofar + +theorem List.distinctKeys_interSmallerFn [BEq α] [PartialEquivBEq α] + (l sofar : List ((a : α) × β a)) (k : α) (hs : DistinctKeys sofar) : + DistinctKeys (interSmallerFn l sofar k) := by + rw [interSmallerFn] + split + · exact hs.insertEntry + · exact hs + +theorem List.getEntry?_interSmallerFn [BEq α] [PartialEquivBEq α] (l sofar : List ((a : α) × β a)) (k k' : α) : + List.getEntry? k' (interSmallerFn l sofar k) = + ((List.getEntry? k' l).filter (fun kv => k == kv.1)).or (sofar.getEntry? k') := by + rw [interSmallerFn] + split + · rename_i kv hkv + rw [getEntry?_insertEntry] + split <;> rename_i hk + · have hk' : k == k' := (BEq.trans (BEq.symm (List.beq_of_getEntry?_eq_some hkv)) hk) + simp [← List.getEntry?_congr hk', hkv, Option.filter_some, BEq.trans hk' (BEq.symm hk)] + · rw [Option.or_eq_right_of_none] + apply Option.filter_eq_none_iff.2 + intro p hp + have hp' := List.beq_of_getEntry?_eq_some hp + have hkv' := List.beq_of_getEntry?_eq_some hkv + exact fun h => hk (BEq.trans hkv' (BEq.trans h hp')) + · rename_i hk + rw [Option.or_eq_right_of_none] + apply Option.filter_eq_none_iff.2 + intro p hp + have := List.beq_of_getEntry?_eq_some hp + intro hkp + have := BEq.trans hkp this + simp [List.getEntry?_congr this, hp] at hk + +theorem toListModel_interSmallerFnₘ [BEq α] [EquivBEq α] [Hashable α] [LawfulHashable α] (m sofar : Raw₀ α β) + (l : List ((a : α) × β a)) + (hm : Raw.WFImp m.1) (hs : Raw.WFImp sofar.1) (k : α) (hml : toListModel sofar.1.buckets ~ l) : + Perm (toListModel ((interSmallerFnₘ m sofar k).1.buckets)) + (List.interSmallerFn (toListModel m.1.buckets) l k) := by + rw [interSmallerFnₘ, getEntry?ₘ_eq_getEntry? hm, List.interSmallerFn] + split + · simpa [*] using (toListModel_insertₘ hs).trans (List.insertEntry_of_perm hs.distinct hml) + · simp [*] + +/-- Internal implementation detail of the hash map -/ +def List.interSmaller [BEq α] (l₁ l₂ : List ((a : α) × β a)) : List ((a : α) × β a) := + l₂.foldl (fun sofar kv => List.interSmallerFn l₁ sofar kv.1) [] + +@[simp] +theorem Option.filter_false {o : Option α} : o.filter (fun _ => false) = none := by + cases o <;> simp + +theorem Option.filter_or {o : Option α} {p q : α → Bool} : o.filter (fun a => p a || q a) = + (o.filter p).or (o.filter q) := by + cases o with + | none => simp + | some a => + simp [Option.filter_some] + cases p a <;> cases q a <;> simp + +theorem List.getEntry?_interSmaller [BEq α] [PartialEquivBEq α] (l₁ l₂ : List ((a : α) × β a)) (k : α) : + List.getEntry? k (interSmaller l₁ l₂) = (List.getEntry? k l₁).filter (fun kv => containsKey kv.1 l₂) := by + rw [interSmaller] + suffices ∀ l₃, + List.getEntry? k (foldl (fun sofar kv => interSmallerFn l₁ sofar kv.fst) l₃ l₂) = + (Option.filter (fun kv => containsKey kv.fst l₂) (List.getEntry? k l₁)).or (List.getEntry? k l₃) by + simpa using this [] + intro l₃ + induction l₂ using assoc_induction generalizing l₃ with + | nil => simp + | cons k v tl ih => + rw [List.foldl_cons, ih] + simp only [List.containsKey_cons, Bool.or_comm (k == _), Option.filter_or, Option.or_assoc, + List.getEntry?_interSmallerFn] + +theorem List.distinctKeys_interSmaller [BEq α] [PartialEquivBEq α] {l₁ l₂ : List ((a : α) × β a)} : + DistinctKeys (interSmaller l₁ l₂) := by + rw [interSmaller] + suffices ∀ l, DistinctKeys l → DistinctKeys (l₂.foldl (fun sofar kv => List.interSmallerFn l₁ sofar kv.1) l) by + simpa using this [] (by simp) + intro l hl + induction l₂ generalizing l with + | nil => simpa + | cons ht tl ih => + rw [List.foldl_cons] + apply ih + exact distinctKeys_interSmallerFn _ _ _ hl + +theorem List.interSmaller_perm_filter [BEq α] [EquivBEq α] (l₁ l₂ : List ((a : α) × β a)) (h₁ : DistinctKeys l₁) : + List.interSmaller l₁ l₂ ~ l₁.filter (fun kv => containsKey kv.1 l₂) := by + apply List.getEntry?_ext + · exact List.distinctKeys_interSmaller + · exact h₁.filter (f := fun k v => containsKey k l₂) + · intro k' + rw [List.getEntry?_filter h₁, List.getEntry?_interSmaller] + +theorem toListModel_interSmallerₘ [BEq α] [EquivBEq α] [Hashable α] [LawfulHashable α] + (m₁ : Raw₀ α β) (m₂ : Raw α β) (hm₁ : Raw.WFImp m₁.1) : + toListModel (m₁.interSmallerₘ m₂).1.buckets ~ + List.interSmaller (toListModel m₁.1.buckets) (toListModel m₂.buckets) := by + rw [interSmallerₘ, Raw.fold_eq_foldl_toListModel, List.interSmaller] + generalize toListModel m₂.buckets = l + suffices ∀ m l', Raw.WFImp m.1 → toListModel m.1.buckets ~ l' → toListModel (foldl (fun a b => m₁.interSmallerFnₘ a b.fst) m l).val.buckets ~ + foldl (fun sofar kv => List.interSmallerFn (toListModel m₁.val.buckets) sofar kv.fst) l' l by + simpa using this emptyWithCapacity [] wfImp_emptyWithCapacity (by simp) + intro m l' hm hml' + induction l generalizing m l' with + | nil => simpa + | cons ht tl ih => + rw [List.foldl_cons, List.foldl_cons] + exact ih _ _ (wfImp_interSmallerFnₘ _ _ hm _) (toListModel_interSmallerFnₘ _ _ _ hm₁ hm _ hml') theorem toListModel_inter [BEq α] [EquivBEq α] [Hashable α] [LawfulHashable α] (m₁ m₂ : Raw₀ α β) (hm₁ : Raw.WFImp m₁.1) (hm₂ : Raw.WFImp m₂.1) : Perm (toListModel (m₁.inter m₂).1.buckets) ((toListModel m₁.1.buckets).filter fun p => containsKey p.1 (toListModel m₂.1.buckets) ) := by simp [inter] split - . rw [filter_eq_filterₘ] - apply Perm.trans - . apply toListModel_filterₘ - . suffices (fun p : (a : α) × β a => m₂.contains p.fst) = (fun p : (a : α) × β a => containsKey p.fst (toListModel m₂.val.buckets)) by rw [this] - ext x - rw [contains_eq_containsKey] - exact hm₂ - . apply Perm.trans - . rw [interSmaller_eq_interSmallerₘ] - . rw [interSmallerₘ] - apply Perm.trans - . apply toListModel_interSmallerₘ - . exact hm₁ - . apply wfImp_emptyWithCapacity - . generalize heq1 : (toListModel m₁.val.buckets) = l₁ - generalize heq2 : (toListModel m₂.val.buckets) = l₂ - apply getEntry?_ext - . apply foldl_insertEntry_distinctKeys - rw [toListModel_buckets_emptyWithCapacity] - simp - . apply @DistinctKeys.filter _ _ _ l₁ (fun k v => containsKey k l₂) - rw [← heq1] - exact hm₁.distinct - . intro a - simp - apply getEntry_foldl - . rw [← heq1] - exact hm₁.distinct - . intro a _ _ - apply Or.inl - simp + · rw [filter_eq_filterₘ] + simp only [contains_eq_containsKey hm₂] + exact toListModel_filterₘ + · rw [interSmaller_eq_interSmallerₘ] + exact Perm.trans (toListModel_interSmallerₘ _ _ hm₁) + (List.interSmaller_perm_filter _ _ hm₁.distinct) theorem wf_inter₀ [BEq α] [Hashable α] [LawfulHashable α] {m₁ m₂ : Raw α β} {h₁ : 0 < m₁.buckets.size} {h₂ : 0 < m₂.buckets.size} (wh₁ : m₁.WF) : (Raw₀.inter ⟨m₁, h₁⟩ ⟨m₂, h₂⟩).1.WF := by rw [inter] split - . apply Raw.WF.filter₀ wh₁ - . rw [interSmaller] + · apply Raw.WF.filter₀ wh₁ + · rw [interSmaller] apply @Raw.fold_induction _ _ _ (fun sofar k x => interSmallerFn ⟨m₁, h₁⟩ sofar k) _ _ (·.val.WF) Raw.WF.emptyWithCapacity₀ intro acc a b wf rw [interSmallerFn] split - . apply Raw.WF.insert₀ wf - . apply wf + · apply Raw.WF.insert₀ wf + · apply wf /-! # `Const.insertListₘ` -/ From 3b74c6706ed4fed598cae4ceaee2efa702ce3593 Mon Sep 17 00:00:00 2001 From: Wojciech Rozowski Date: Wed, 12 Nov 2025 17:06:56 +0100 Subject: [PATCH 37/88] fix deleted code --- src/Std/Data/DHashMap/Internal/WF.lean | 11 +++++++++++ 1 file changed, 11 insertions(+) diff --git a/src/Std/Data/DHashMap/Internal/WF.lean b/src/Std/Data/DHashMap/Internal/WF.lean index d79ca8fdb296..aa7dbf9c1caa 100644 --- a/src/Std/Data/DHashMap/Internal/WF.lean +++ b/src/Std/Data/DHashMap/Internal/WF.lean @@ -1334,6 +1334,17 @@ theorem wf_union₀ [BEq α] [Hashable α] [EquivBEq α] [LawfulHashable α] · exact wf_insertManyIfNew₀ ‹_› · exact wf_insertMany₀ ‹_› + +theorem toListModel_union [BEq α] [Hashable α] [EquivBEq α] [LawfulHashable α] {m₁ m₂ : Raw₀ α β} + (h₁ : Raw.WFImp m₁.1) (h₂ : Raw.WFImp m₂.1) : + Perm (toListModel (m₁.union m₂).1.buckets) + (List.insertList (toListModel m₁.1.buckets) (toListModel m₂.1.buckets)) := by + rw [union_eq_unionₘ] + exact toListModel_unionₘ h₁ h₂ + + +/-! # `inter` -/ + theorem wfImp_interSmallerFnₘ [BEq α] [EquivBEq α] [Hashable α] [LawfulHashable α] (m₁ : Raw₀ α β) (m₂ : Raw₀ α β) (hm₂ : Raw.WFImp m₂.1) (k : α) : Raw.WFImp (m₁.interSmallerFnₘ m₂ k).1 := by rw [interSmallerFnₘ] From 869351272be38ee79929dc20626ad17912da91e2 Mon Sep 17 00:00:00 2001 From: Wojciech Rozowski Date: Wed, 12 Nov 2025 17:28:58 +0100 Subject: [PATCH 38/88] renaming --- src/Std/Data/DHashMap/Internal/RawLemmas.lean | 109 +++++++----------- src/Std/Data/Internal/List/Associative.lean | 50 ++++---- 2 files changed, 64 insertions(+), 95 deletions(-) diff --git a/src/Std/Data/DHashMap/Internal/RawLemmas.lean b/src/Std/Data/DHashMap/Internal/RawLemmas.lean index 165f746ae3de..d81d0a349d4e 100644 --- a/src/Std/Data/DHashMap/Internal/RawLemmas.lean +++ b/src/Std/Data/DHashMap/Internal/RawLemmas.lean @@ -2777,19 +2777,19 @@ theorem contains_inter_iff [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF (m₁.inter m₂).contains k ↔ m₁.contains k ∧ m₂.contains k := by simp_to_model [inter, contains] using List.containsKey_filter_containsKey_iff -theorem not_contains_inter_of_not_contains_left [EquivBEq α] [LawfulHashable α] +theorem not_contains_inter_of_contains_eq_false_left [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} (h : m₁.contains k = false) : (m₁.inter m₂).contains k = false := by revert h - simp_to_model [inter, contains] using List.not_contains_filter_containsKey_of_not_containsKey_left + simp_to_model [inter, contains] using List.not_contains_filter_containsKey_of_containsKey_eq_false_left -theorem not_contains_inter_of_not_contains_right [EquivBEq α] [LawfulHashable α] +theorem not_contains_inter_of_contains_eq_false_right [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} (h : m₂.contains k = false) : (m₁.inter m₂).contains k = false := by revert h - simp_to_model [inter, contains] using List.not_contains_filter_containsKey_of_not_containsKey_right + simp_to_model [inter, contains] using List.not_contains_filter_containsKey_of_containsKey_eq_false_right /- get? -/ theorem get?_inter [LawfulBEq α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} : @@ -2803,17 +2803,17 @@ theorem get?_inter_of_contains_right [LawfulBEq α] (h₁ : m₁.val.WF) (h₂ : revert h simp_to_model [inter, get?, contains] using List.getValueCast?_filter_containsKey_of_containsKey_right -theorem get?_inter_of_not_contains_left [LawfulBEq α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) +theorem get?_inter_of_contains_eq_false_left [LawfulBEq α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} (h : m₁.contains k = false) : (m₁.inter m₂).get? k = none := by revert h - simp_to_model [inter, get?, contains] using List.getValueCast?_filter_containsKey_of_not_containsKey_left + simp_to_model [inter, get?, contains] using List.getValueCast?_filter_containsKey_of_containsKey_eq_false_left -theorem get?_inter_of_not_contains_right [LawfulBEq α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) +theorem get?_inter_of_contains_eq_false_right [LawfulBEq α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} (h : m₂.contains k = false) : (m₁.inter m₂).get? k = none := by revert h - simp_to_model [inter, get?, contains] using List.getValueCast?_filter_containsKey_of_not_containsKey_right + simp_to_model [inter, get?, contains] using List.getValueCast?_filter_containsKey_of_containsKey_eq_false_right /- get -/ theorem get_inter [LawfulBEq α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) @@ -2835,17 +2835,17 @@ theorem getD_inter_of_contains_right [LawfulBEq α] (h₁ : m₁.val.WF) (h₂ : revert h simp_to_model [inter, getD, contains] using List.getValueCastD_filter_containsKey_of_containsKey_right -theorem getD_inter_of_not_contains_right [LawfulBEq α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) +theorem getD_inter_of_contains_eq_false_right [LawfulBEq α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} {fallback : β k} (h : m₂.contains k = false) : (m₁.inter m₂).getD k fallback = fallback := by revert h - simp_to_model [inter, getD, contains] using List.getValueCastD_filter_containsKey_of_not_containsKey_right + simp_to_model [inter, getD, contains] using List.getValueCastD_filter_containsKey_of_containsKey_eq_false_right -theorem getD_inter_of_not_contains_left [LawfulBEq α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) +theorem getD_inter_of_contains_eq_false_left [LawfulBEq α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} {fallback : β k} (h : m₁.contains k = false) : (m₁.inter m₂).getD k fallback = fallback := by revert h - simp_to_model [inter, getD, contains] using List.getValueCastD_filter_containsKey_of_not_containsKey_left + simp_to_model [inter, getD, contains] using List.getValueCastD_filter_containsKey_of_containsKey_eq_false_left /- get! -/ theorem get!_inter [LawfulBEq α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) @@ -2860,17 +2860,17 @@ theorem get!_inter_of_contains_right [LawfulBEq α] (h₁ : m₁.val.WF) (h₂ : revert h simp_to_model [inter, get!, contains] using List.getValueCastD_filter_containsKey_of_containsKey_right -theorem get!_inter_of_not_contains_right [LawfulBEq α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) +theorem get!_inter_of_contains_eq_false_right [LawfulBEq α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} [Inhabited (β k)] (h : m₂.contains k = false) : (m₁.inter m₂).get! k = default := by revert h - simp_to_model [inter, get!, contains] using List.getValueCastD_filter_containsKey_of_not_containsKey_right + simp_to_model [inter, get!, contains] using List.getValueCastD_filter_containsKey_of_containsKey_eq_false_right -theorem get!_inter_of_not_contains_left [LawfulBEq α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) +theorem get!_inter_of_contains_eq_false_left [LawfulBEq α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} [Inhabited (β k)] (h : m₁.contains k = false) : (m₁.inter m₂).get! k = default := by revert h - simp_to_model [inter, get!, contains] using List.getValueCastD_filter_containsKey_of_not_containsKey_left + simp_to_model [inter, get!, contains] using List.getValueCastD_filter_containsKey_of_containsKey_eq_false_left /- getKey? -/ theorem getKey?_inter [EquivBEq α] [LawfulHashable α] @@ -2885,17 +2885,17 @@ theorem getKey?_inter_of_contains_right [EquivBEq α] [LawfulHashable α] revert h simp_to_model [contains, getKey?, inter] using List.getKey?_filter_containsKey_of_containsKey_right -theorem getKey?_inter_of_not_contains_right [EquivBEq α] [LawfulHashable α] +theorem getKey?_inter_of_contains_eq_false_right [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} (h : m₂.contains k = false) : (m₁.inter m₂).getKey? k = none := by revert h - simp_to_model [contains, getKey?, inter] using List.getKey?_filter_containsKey_of_not_containsKey_right + simp_to_model [contains, getKey?, inter] using List.getKey?_filter_containsKey_of_containsKey_eq_false_right -theorem getKey?_inter_of_not_contains_left [EquivBEq α] [LawfulHashable α] +theorem getKey?_inter_of_contains_eq_false_left [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} (h : m₁.contains k = false) : (m₁.inter m₂).getKey? k = none := by revert h - simp_to_model [contains, getKey?, inter] using List.getKey?_filter_containsKey_of_not_contains_left + simp_to_model [contains, getKey?, inter] using List.getKey?_filter_containsKey_of_containsKey_eq_false_left /- getKey -/ theorem getKey_inter [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) @@ -2917,17 +2917,17 @@ theorem getKeyD_inter_of_contains_right [EquivBEq α] [LawfulHashable α] (h₁ revert h simp_to_model [inter, getKeyD, contains] using List.getKeyD_filter_containsKey_of_containsKey_right -theorem getKeyD_inter_of_not_contains_right [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) +theorem getKeyD_inter_of_contains_eq_false_right [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k fallback : α} (h : m₂.contains k = false) : (m₁.inter m₂).getKeyD k fallback = fallback := by revert h - simp_to_model [inter, getKeyD, contains] using List.getKeyD_filter_containsKey_of_not_containsKey_right + simp_to_model [inter, getKeyD, contains] using List.getKeyD_filter_containsKey_of_containsKey_eq_false_right -theorem getKeyD_inter_of_not_contains_left [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) +theorem getKeyD_inter_of_contains_eq_false_left [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k fallback : α} (h : m₁.contains k = false) : (m₁.inter m₂).getKeyD k fallback = fallback := by revert h - simp_to_model [inter, getKeyD, contains] using List.getKeyD_filter_containsKey_of_not_containsKey_left + simp_to_model [inter, getKeyD, contains] using List.getKeyD_filter_containsKey_of_containsKey_eq_false_left /- getKey! -/ theorem getKey!_inter [EquivBEq α] [LawfulHashable α] [Inhabited α] (h₁ : m₁.val.WF) @@ -2942,48 +2942,17 @@ theorem getKey!_inter_of_contains_right [EquivBEq α] [LawfulHashable α] [Inhab revert h simp_to_model [inter, getKey!, contains] using List.getKeyD_filter_containsKey_of_containsKey_right -theorem getKey!_inter_of_not_contains_right [EquivBEq α] [LawfulHashable α] [Inhabited α] (h₁ : m₁.val.WF) +theorem getKey!_inter_of_contains_eq_false_right [EquivBEq α] [LawfulHashable α] [Inhabited α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} (h : m₂.contains k = false) : (m₁.inter m₂).getKey! k = default := by revert h - simp_to_model [inter, getKey!, contains] using List.getKeyD_filter_containsKey_of_not_containsKey_right + simp_to_model [inter, getKey!, contains] using List.getKeyD_filter_containsKey_of_containsKey_eq_false_right -theorem getKey!_inter_of_not_contains_left [EquivBEq α] [LawfulHashable α] [Inhabited α] (h₁ : m₁.val.WF) +theorem getKey!_inter_of_contains_eq_false_left [EquivBEq α] [LawfulHashable α] [Inhabited α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} (h : m₁.contains k = false) : (m₁.inter m₂).getKey! k = default := by revert h - simp_to_model [inter, getKey!, contains] using List.getKeyD_filter_containsKey_of_not_containsKey_left - -/- getEntry? -/ -theorem getEntry?_inter [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} : - (m₁.inter m₂).getEntry? k = - if m₂.contains k then m₁.getEntry? k else none := by - simp_to_model [inter, getEntry?, contains] using List.getEntry?_filter_containsKey - -theorem getEntry?_inter_of_contains_right [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) - {k : α} (h : m₂.contains k) : - (m₁.inter m₂).getEntry? k = m₁.getEntry? k := by - revert h - simp_to_model [inter, getEntry?, contains] using List.getEntry?_filter_containsKey_of_containsKey_right - -theorem getEntry?_inter_of_not_contains_left [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) - {k : α} (h : m₁.contains k = false) : - (m₁.inter m₂).getEntry? k = none := by - revert h - simp_to_model [inter, getEntry?, contains] using List.getEntry?_filter_containsKey_of_not_containsKey_left - -theorem getEntry?_inter_of_not_contains_right [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) - {k : α} (h : m₂.contains k = false) : - (m₁.inter m₂).getEntry? k = none := by - revert h - simp_to_model [inter, getEntry?, contains] using List.getEntry?_filter_containsKey_of_not_containsKey_right - -/- getEntry -/ -theorem getEntry_inter [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) - {k : α} {h_contains : (m₁.inter m₂).contains k} : - (m₁.inter m₂).getEntry k h_contains = - m₁.getEntry k ((contains_inter_iff h₁ h₂).1 h_contains).1 := by - simp_to_model [inter, getEntry, contains] using List.getEntry_filter_containsKey + simp_to_model [inter, getKey!, contains] using List.getKeyD_filter_containsKey_of_containsKey_eq_false_left /- size -/ theorem size_inter_le_size_left [EquivBEq α] [LawfulHashable α] @@ -3051,17 +3020,17 @@ theorem get?_inter_of_contains_right [EquivBEq α] [LawfulHashable α] (h₁ : m revert h simp_to_model [inter, Const.get?, contains] using List.getValue?_filter_containsKey_eq_containsKey -theorem get?_inter_of_not_contains_left [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) +theorem get?_inter_of_contains_eq_false_left [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} (h : m₁.contains k = false) : Const.get? (m₁.inter m₂) k = none := by revert h simp_to_model [inter, Const.get?, contains] using List.getValue?_filter_of_not_contains -theorem get?_inter_of_not_contains_right [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) +theorem get?_inter_of_contains_eq_false_right [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} (h : m₂.contains k = false) : Const.get? (m₁.inter m₂) k = none := by revert h - simp_to_model [inter, Const.get?, contains] using List.getValue?_filter_of_not_contains_right + simp_to_model [inter, Const.get?, contains] using List.getValue?_filter_of_containsKey_eq_false_right theorem get_inter [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} {h_contains : (m₁.inter m₂).contains k} : @@ -3082,17 +3051,17 @@ theorem getD_inter_of_contains_right [EquivBEq α] [LawfulHashable α] (h₁ : m revert h simp_to_model [inter, Const.getD, contains] using List.getValueD_filter_containsKey_eq_containsKey -theorem getD_inter_of_not_contains_right [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) +theorem getD_inter_of_contains_eq_false_right [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} {fallback : β} (h : m₂.contains k = false) : Const.getD (m₁.inter m₂) k fallback = fallback := by revert h - simp_to_model [inter, Const.getD, contains] using List.getValueD_filter_of_not_contains_right + simp_to_model [inter, Const.getD, contains] using List.getValueD_filter_of_containsKey_eq_false_right -theorem getD_inter_of_not_contains_left [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) +theorem getD_inter_of_contains_eq_false_left [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} {fallback : β} (h : m₁.contains k = false) : Const.getD (m₁.inter m₂) k fallback = fallback := by revert h - simp_to_model [inter, Const.getD, contains] using List.getValueD_filter_of_not_contains_left + simp_to_model [inter, Const.getD, contains] using List.getValueD_filter_of_containsKey_eq_false_left /- get! -/ theorem get!_inter [EquivBEq α] [LawfulHashable α] [Inhabited β] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) @@ -3107,17 +3076,17 @@ theorem get!_inter_of_contains_right [EquivBEq α] [LawfulHashable α] [Inhabite revert h simp_to_model [inter, Const.get!, contains] using List.getValueD_filter_containsKey_eq_containsKey -theorem get!_inter_of_not_contains_right [EquivBEq α] [LawfulHashable α] [Inhabited β] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) +theorem get!_inter_of_contains_eq_false_right [EquivBEq α] [LawfulHashable α] [Inhabited β] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} (h : m₂.contains k = false) : Const.get! (m₁.inter m₂) k = default := by revert h - simp_to_model [inter, Const.get!, contains] using List.getValueD_filter_of_not_contains_right + simp_to_model [inter, Const.get!, contains] using List.getValueD_filter_of_containsKey_eq_false_right -theorem get!_inter_of_not_contains_left [EquivBEq α] [LawfulHashable α] [Inhabited β] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) +theorem get!_inter_of_contains_eq_false_left [EquivBEq α] [LawfulHashable α] [Inhabited β] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} (h : m₁.contains k = false) : Const.get! (m₁.inter m₂) k = default := by revert h - simp_to_model [inter, Const.get!, contains] using List.getValueD_filter_of_not_contains_left + simp_to_model [inter, Const.get!, contains] using List.getValueD_filter_of_containsKey_eq_false_left end Const diff --git a/src/Std/Data/Internal/List/Associative.lean b/src/Std/Data/Internal/List/Associative.lean index 847e1d46d2b3..4586651c5b26 100644 --- a/src/Std/Data/Internal/List/Associative.lean +++ b/src/Std/Data/Internal/List/Associative.lean @@ -5277,7 +5277,7 @@ theorem List.getKey?_filter_containsKey_of_containsKey_right [BEq α] [EquivBEq · simp only [↓reduceIte] · exact dl₁ -theorem List.getKey?_filter_containsKey_of_not_contains_left [BEq α] [EquivBEq α] +theorem List.getKey?_filter_containsKey_of_containsKey_eq_false_left [BEq α] [EquivBEq α] {l₁ l₂ : List ((a : α) × β a)} {k : α} (dl₁ : DistinctKeys l₁) : containsKey k l₁ = false → getKey? k (List.filter (fun p => containsKey p.fst l₂) l₁) = none := by @@ -5288,7 +5288,7 @@ theorem List.getKey?_filter_containsKey_of_not_contains_left [BEq α] [EquivBEq · rfl · exact dl₁ -theorem List.getKey?_filter_containsKey_of_not_containsKey_right [BEq α] [EquivBEq α] +theorem List.getKey?_filter_containsKey_of_containsKey_eq_false_right [BEq α] [EquivBEq α] {l₁ l₂ : List ((a : α) × β a)} {k : α} (dl₁ : DistinctKeys l₁) : containsKey k l₂ = false → getKey? k (List.filter (fun p => containsKey p.fst l₂) l₁) = none := by @@ -5307,7 +5307,7 @@ theorem List.getValueCast?_filter_containsKey_of_containsKey_right [BEq α] [Law . simp only [↓reduceIte] . exact dl₁ -theorem List.getValueCast?_filter_containsKey_of_not_containsKey_left [BEq α] [LawfulBEq α] +theorem List.getValueCast?_filter_containsKey_of_containsKey_eq_false_left [BEq α] [LawfulBEq α] {l₁ l₂ : List ((a : α) × β a)} {k : α} (dl₁ : DistinctKeys l₁) : containsKey k l₁ = false → getValueCast? k (List.filter (fun p => containsKey p.fst l₂) l₁) = none := by @@ -5318,7 +5318,7 @@ theorem List.getValueCast?_filter_containsKey_of_not_containsKey_left [BEq α] [ . rfl . exact dl₁ -theorem List.getValueCast?_filter_containsKey_of_not_containsKey_right [BEq α] [LawfulBEq α] +theorem List.getValueCast?_filter_containsKey_of_containsKey_eq_false_right [BEq α] [LawfulBEq α] {l₁ l₂ : List ((a : α) × β a)} {k : α} (dl₁ : DistinctKeys l₁) : containsKey k l₂ = false → getValueCast? k (List.filter (fun p => containsKey p.fst l₂) l₁) = none := by @@ -5368,7 +5368,7 @@ theorem List.getEntry?_filter_containsKey_of_containsKey_right [BEq α] [EquivBE · simp only [↓reduceIte] · exact dl₁ -theorem List.getEntry?_filter_containsKey_of_not_containsKey_left [BEq α] [EquivBEq α] +theorem List.getEntry?_filter_containsKey_of_containsKey_eq_false_left [BEq α] [EquivBEq α] {l₁ l₂ : List ((a : α) × β a)} {k : α} (dl₁ : DistinctKeys l₁) : containsKey k l₁ = false → getEntry? k (List.filter (fun p => containsKey p.fst l₂) l₁) = none := by @@ -5380,7 +5380,7 @@ theorem List.getEntry?_filter_containsKey_of_not_containsKey_left [BEq α] [Equi · rfl · exact dl₁ -theorem List.getEntry?_filter_containsKey_of_not_containsKey_right [BEq α] [EquivBEq α] +theorem List.getEntry?_filter_containsKey_of_containsKey_eq_false_right [BEq α] [EquivBEq α] {l₁ l₂ : List ((a : α) × β a)} {k : α} (dl₁ : DistinctKeys l₁) : containsKey k l₂ = false → getEntry? k (List.filter (fun p => containsKey p.fst l₂) l₁) = none := by @@ -5464,14 +5464,14 @@ theorem List.getKey_filter_containsKey [BEq α] [EquivBEq α] · exact h₁.2 · exact dl₁ -theorem List.not_contains_filter_containsKey_of_not_containsKey_left [BEq α] [EquivBEq α] {l₁ l₂ : List ((a : α) × β a)} {hl₁ : DistinctKeys l₁} {k : α} : +theorem List.not_contains_filter_containsKey_of_containsKey_eq_false_left [BEq α] [EquivBEq α] {l₁ l₂ : List ((a : α) × β a)} {hl₁ : DistinctKeys l₁} {k : α} : containsKey k l₁ = false → containsKey k (List.filter (fun p => containsKey p.fst l₂) l₁) = false := by intro h rw [containsKey_filter_containsKey] . simp [h] . exact hl₁ -theorem List.not_contains_filter_containsKey_of_not_containsKey_right [BEq α] [EquivBEq α] {l₁ l₂ : List ((a : α) × β a)} {hl₁ : DistinctKeys l₁} {k : α} : +theorem List.not_contains_filter_containsKey_of_containsKey_eq_false_right [BEq α] [EquivBEq α] {l₁ l₂ : List ((a : α) × β a)} {hl₁ : DistinctKeys l₁} {k : α} : containsKey k l₂ = false → containsKey k (List.filter (fun p => containsKey p.fst l₂) l₁) = false := by intro h rw [containsKey_filter_containsKey] @@ -5501,7 +5501,7 @@ theorem getValue?_filter_containsKey_eq_if_containsKey {β : Type v} [BEq α] [E case isFalse h => simp only [Option.map_eq_none_iff, getEntry?_eq_none] simp only [Bool.not_eq_true] at h - apply List.not_contains_filter_containsKey_of_not_containsKey_right + apply List.not_contains_filter_containsKey_of_containsKey_eq_false_right . exact dl₁ . exact h @@ -5527,7 +5527,7 @@ theorem getValue?_filter_of_not_contains {β : Type v} [BEq α] [EquivBEq α] · rfl · exact dl₁ -theorem getValue?_filter_of_not_contains_right {β : Type v} [BEq α] [EquivBEq α] +theorem List.getValue?_filter_of_containsKey_eq_false_right {β : Type v} [BEq α] [EquivBEq α] {l₁ l₂ : List ((_ : α) × β)} {k : α} (dl₁ : DistinctKeys l₁) : containsKey k l₂ = false → getValue? k (List.filter (fun p => containsKey p.fst l₂) l₁) = none := by @@ -5548,7 +5548,7 @@ theorem List.getValueCastD_filter_containsKey [BEq α] [LawfulBEq α] apply List.getValueCast?_filter_containsKey_of_containsKey_right dl₁ h case isFalse h => apply getValueCastD_eq_fallback - apply List.not_contains_filter_containsKey_of_not_containsKey_right + apply List.not_contains_filter_containsKey_of_containsKey_eq_false_right . exact dl₁ . simp only [h] @@ -5562,23 +5562,23 @@ theorem List.getValueCastD_filter_containsKey_of_containsKey_right [BEq α] [Law congr 1 apply List.getValueCast?_filter_containsKey_of_containsKey_right dl₁ h -theorem List.getValueCastD_filter_containsKey_of_not_containsKey_left [BEq α] [LawfulBEq α] +theorem List.getValueCastD_filter_containsKey_of_containsKey_eq_false_left [BEq α] [LawfulBEq α] {l₁ l₂ : List ((a : α) × β a)} {k : α} {fallback} (dl₁ : DistinctKeys l₁) : containsKey k l₁ = false → getValueCastD k (List.filter (fun p => containsKey p.fst l₂) l₁) fallback = fallback := by intro h apply getValueCastD_eq_fallback - apply List.not_contains_filter_containsKey_of_not_containsKey_left + apply List.not_contains_filter_containsKey_of_containsKey_eq_false_left . exact dl₁ . exact h -theorem List.getValueCastD_filter_containsKey_of_not_containsKey_right [BEq α] [LawfulBEq α] +theorem List.getValueCastD_filter_containsKey_of_containsKey_eq_false_right [BEq α] [LawfulBEq α] {l₁ l₂ : List ((a : α) × β a)} {k : α} {fallback} (dl₁ : DistinctKeys l₁) : containsKey k l₂ = false → getValueCastD k (List.filter (fun p => containsKey p.fst l₂) l₁) fallback = fallback := by intro h apply getValueCastD_eq_fallback - apply List.not_contains_filter_containsKey_of_not_containsKey_right + apply List.not_contains_filter_containsKey_of_containsKey_eq_false_right . exact dl₁ . exact h @@ -5594,7 +5594,7 @@ theorem List.getKeyD_filter_containsKey [BEq α] [EquivBEq α] apply List.getKey?_filter_containsKey_of_containsKey_right dl₁ h case isFalse h => apply getKeyD_eq_fallback - apply List.not_contains_filter_containsKey_of_not_containsKey_right + apply List.not_contains_filter_containsKey_of_containsKey_eq_false_right · exact dl₁ · simp only [h] @@ -5608,23 +5608,23 @@ theorem List.getKeyD_filter_containsKey_of_containsKey_right [BEq α] [EquivBEq congr 1 apply List.getKey?_filter_containsKey_of_containsKey_right dl₁ h -theorem List.getKeyD_filter_containsKey_of_not_containsKey_left [BEq α] [EquivBEq α] +theorem List.getKeyD_filter_containsKey_of_containsKey_eq_false_left [BEq α] [EquivBEq α] {l₁ l₂ : List ((a : α) × β a)} {k : α} {fallback : α} (dl₁ : DistinctKeys l₁) : containsKey k l₁ = false → getKeyD k (List.filter (fun p => containsKey p.fst l₂) l₁) fallback = fallback := by intro h apply getKeyD_eq_fallback - apply List.not_contains_filter_containsKey_of_not_containsKey_left + apply List.not_contains_filter_containsKey_of_containsKey_eq_false_left · exact dl₁ · exact h -theorem List.getKeyD_filter_containsKey_of_not_containsKey_right [BEq α] [EquivBEq α] +theorem List.getKeyD_filter_containsKey_of_containsKey_eq_false_right [BEq α] [EquivBEq α] {l₁ l₂ : List ((a : α) × β a)} {k : α} {fallback : α} (dl₁ : DistinctKeys l₁) : containsKey k l₂ = false → getKeyD k (List.filter (fun p => containsKey p.fst l₂) l₁) fallback = fallback := by intro h apply getKeyD_eq_fallback - apply List.not_contains_filter_containsKey_of_not_containsKey_right + apply List.not_contains_filter_containsKey_of_containsKey_eq_false_right · exact dl₁ · exact h @@ -5652,7 +5652,7 @@ theorem getValueD_filter_containsKey_eq_if_containsKey {β : Type v} [BEq α] [E apply getValue?_filter_containsKey_eq_containsKey dl₁ h case isFalse h => apply getValueD_eq_fallback - apply List.not_contains_filter_containsKey_of_not_containsKey_right + apply List.not_contains_filter_containsKey_of_containsKey_eq_false_right · exact dl₁ · simp only [h] @@ -5666,23 +5666,23 @@ theorem getValueD_filter_containsKey_eq_containsKey {β : Type v} [BEq α] [Equi congr 1 apply getValue?_filter_containsKey_eq_containsKey dl₁ h -theorem getValueD_filter_of_not_contains_left {β : Type v} [BEq α] [EquivBEq α] +theorem List.getValueD_filter_of_containsKey_eq_false_left {β : Type v} [BEq α] [EquivBEq α] {l₁ l₂ : List ((_ : α) × β)} {k : α} {fallback : β} (dl₁ : DistinctKeys l₁) : containsKey k l₁ = false → getValueD k (List.filter (fun p => containsKey p.fst l₂) l₁) fallback = fallback := by intro h apply getValueD_eq_fallback - apply List.not_contains_filter_containsKey_of_not_containsKey_left + apply List.not_contains_filter_containsKey_of_containsKey_eq_false_left · exact dl₁ · exact h -theorem getValueD_filter_of_not_contains_right {β : Type v} [BEq α] [EquivBEq α] +theorem List.getValueD_filter_of_containsKey_eq_false_right {β : Type v} [BEq α] [EquivBEq α] {l₁ l₂ : List ((_ : α) × β)} {k : α} {fallback : β} (dl₁ : DistinctKeys l₁) : containsKey k l₂ = false → getValueD k (List.filter (fun p => containsKey p.fst l₂) l₁) fallback = fallback := by intro h apply getValueD_eq_fallback - apply List.not_contains_filter_containsKey_of_not_containsKey_right + apply List.not_contains_filter_containsKey_of_containsKey_eq_false_right · exact dl₁ · exact h From dbfe83112342be9f3930964e8ef5acc8a826b269 Mon Sep 17 00:00:00 2001 From: Wojciech Rozowski Date: Wed, 12 Nov 2025 17:58:19 +0100 Subject: [PATCH 39/88] rename size lemma --- src/Std/Data/DHashMap/Internal/RawLemmas.lean | 4 ++-- src/Std/Data/Internal/List/Associative.lean | 2 +- 2 files changed, 3 insertions(+), 3 deletions(-) diff --git a/src/Std/Data/DHashMap/Internal/RawLemmas.lean b/src/Std/Data/DHashMap/Internal/RawLemmas.lean index d81d0a349d4e..00555e7e863e 100644 --- a/src/Std/Data/DHashMap/Internal/RawLemmas.lean +++ b/src/Std/Data/DHashMap/Internal/RawLemmas.lean @@ -2979,10 +2979,10 @@ theorem size_inter_eq_size_right [EquivBEq α] [LawfulHashable α] revert h simp_to_model [inter, size, contains] using List.length_filter_containsKey_of_length_right -theorem size_union_inter_eq_size [EquivBEq α] [LawfulHashable α] +theorem size_add_size_eq_size_union_add_size_inter [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) : m₁.1.size + m₂.1.size = (m₁.union m₂).1.size + (m₁.inter m₂).1.size := by - simp_to_model [union, inter, size] using List.size_insertList_filter_containsKey_eq_size + simp_to_model [union, inter, size] using List.size_add_size_eq_size_insertList_add_size_filter_containsKey /- isEmpty -/ @[simp] diff --git a/src/Std/Data/Internal/List/Associative.lean b/src/Std/Data/Internal/List/Associative.lean index 4586651c5b26..fe712c53e3b9 100644 --- a/src/Std/Data/Internal/List/Associative.lean +++ b/src/Std/Data/Internal/List/Associative.lean @@ -6326,7 +6326,7 @@ theorem List.length_filter_containsKey_of_length_right [BEq α] [EquivBEq α] case w.inr => exact inr -theorem List.size_insertList_filter_containsKey_eq_size [BEq α] [EquivBEq α] +theorem List.size_add_size_eq_size_insertList_add_size_filter_containsKey [BEq α] [EquivBEq α] {l₁ l₂ : List ((a : α) × β a)} (hl₁ : DistinctKeys l₁) (hl₂ : DistinctKeys l₂) : l₁.length + l₂.length = (insertList l₁ l₂).length + From a6c2a75524ddbc4233737e4b3e49b3382e772dd9 Mon Sep 17 00:00:00 2001 From: Wojciech Rozowski Date: Wed, 12 Nov 2025 19:05:45 +0100 Subject: [PATCH 40/88] Add Raw DHashMap inter lemmas --- src/Std/Data/DHashMap/Internal/Raw.lean | 4 + src/Std/Data/DHashMap/Internal/RawLemmas.lean | 12 +- src/Std/Data/DHashMap/RawLemmas.lean | 327 +++++++++++++++++- src/Std/Data/Internal/List/Associative.lean | 24 +- 4 files changed, 348 insertions(+), 19 deletions(-) diff --git a/src/Std/Data/DHashMap/Internal/Raw.lean b/src/Std/Data/DHashMap/Internal/Raw.lean index d7f2ee1a5cb6..b421e0f6caf2 100644 --- a/src/Std/Data/DHashMap/Internal/Raw.lean +++ b/src/Std/Data/DHashMap/Internal/Raw.lean @@ -145,6 +145,10 @@ theorem union_eq [BEq α] [Hashable α] {m₁ m₂ : Raw α β} (h₁ : m₁.WF) m₁.union m₂ = Raw₀.union ⟨m₁, h₁.size_buckets_pos⟩ ⟨m₂, h₂.size_buckets_pos⟩ := by simp [Raw.union, h₁.size_buckets_pos, h₂.size_buckets_pos] +theorem inter_eq [BEq α] [Hashable α] {m₁ m₂ : Raw α β} (h₁ : m₁.WF) (h₂ : m₂.WF) : + m₁.inter m₂ = Raw₀.inter ⟨m₁, h₁.size_buckets_pos⟩ ⟨m₂, h₂.size_buckets_pos⟩ := by + simp [Raw.inter, h₁.size_buckets_pos, h₂.size_buckets_pos] + section variable {β : Type v} diff --git a/src/Std/Data/DHashMap/Internal/RawLemmas.lean b/src/Std/Data/DHashMap/Internal/RawLemmas.lean index 00555e7e863e..68d788fcf149 100644 --- a/src/Std/Data/DHashMap/Internal/RawLemmas.lean +++ b/src/Std/Data/DHashMap/Internal/RawLemmas.lean @@ -2777,19 +2777,19 @@ theorem contains_inter_iff [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF (m₁.inter m₂).contains k ↔ m₁.contains k ∧ m₂.contains k := by simp_to_model [inter, contains] using List.containsKey_filter_containsKey_iff -theorem not_contains_inter_of_contains_eq_false_left [EquivBEq α] [LawfulHashable α] +theorem contains_inter_eq_false_of_contains_eq_false_left [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} (h : m₁.contains k = false) : (m₁.inter m₂).contains k = false := by revert h - simp_to_model [inter, contains] using List.not_contains_filter_containsKey_of_containsKey_eq_false_left + simp_to_model [inter, contains] using List.containsKey_filter_containsKey_eq_false_of_containsKey_eq_false_left -theorem not_contains_inter_of_contains_eq_false_right [EquivBEq α] [LawfulHashable α] +theorem contains_inter_eq_false_of_contains_eq_false_right [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} (h : m₂.contains k = false) : (m₁.inter m₂).contains k = false := by revert h - simp_to_model [inter, contains] using List.not_contains_filter_containsKey_of_containsKey_eq_false_right + simp_to_model [inter, contains] using List.containsKey_filter_containsKey_eq_false_of_containsKey_eq_false_right /- get? -/ theorem get?_inter [LawfulBEq α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} : @@ -2993,10 +2993,10 @@ theorem isEmpty_inter_left [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF /- isEmpty -/ @[simp] -theorem isEmpty_inter_right [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) (h : m₁.1.isEmpty) : +theorem isEmpty_inter_right [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) (h : m₂.1.isEmpty) : (m₁.inter m₂).1.isEmpty = true := by revert h - simp_to_model [isEmpty, inter, contains] using List.isEmpty_filter_containsKey_left + simp_to_model [isEmpty, inter, contains] using List.isEmpty_filter_containsKey_right theorem isEmpty_inter_iff [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) : (m₁.inter m₂).1.isEmpty ↔ ∀ k, m₁.contains k → m₂.contains k = false := by diff --git a/src/Std/Data/DHashMap/RawLemmas.lean b/src/Std/Data/DHashMap/RawLemmas.lean index 7f58c941af36..55e83ca37fdd 100644 --- a/src/Std/Data/DHashMap/RawLemmas.lean +++ b/src/Std/Data/DHashMap/RawLemmas.lean @@ -48,7 +48,7 @@ private meta def baseNames : Array Name := ``getKey?_eq, ``getKey_eq, ``getKey!_eq, ``getKeyD_eq, ``insertMany_eq, ``Const.insertMany_eq, ``Const.insertManyIfNewUnit_eq, ``ofList_eq, ``Const.ofList_eq, ``Const.unitOfList_eq, - ``alter_eq, ``Const.alter_eq, ``modify_eq, ``Const.modify_eq, ``union_eq] + ``alter_eq, ``Const.alter_eq, ``modify_eq, ``Const.modify_eq, ``union_eq, ``inter_eq] /-- Internal implementation detail of the hash map -/ scoped syntax "simp_to_raw" ("using" term)? : tactic @@ -2129,6 +2129,331 @@ theorem isEmpty_union [EquivBEq α] [LawfulHashable α] (h₁ : m₁.WF) (h₂ : end Union +section Inter + +variable {m₁ m₂ : Raw α β} + +@[simp] +theorem inter_eq : m₁.inter m₂ = m₁ ∩ m₂ := by + simp only [Inter.inter] + +/- contains -/ + +@[simp] +theorem contains_inter [EquivBEq α] [LawfulHashable α] {k : α} (h₁ : m₁.WF) (h₂ : m₂.WF) : + (m₁ ∩ m₂).contains k = (m₁.contains k && m₂.contains k) := by + simp only [Inter.inter] + simp_to_raw using Raw₀.contains_inter + +/- mem -/ + +theorem mem_inter_iff [EquivBEq α] [LawfulHashable α] (h₁ : m₁.WF) + (h₂ : m₂.WF) {k : α} : + k ∈ m₁ ∩ m₂ ↔ k ∈ m₁ ∧ k ∈ m₂ := by + simp only [Inter.inter, Membership.mem] + simp_to_raw using Raw₀.contains_inter_iff + +theorem not_mem_inter_of_not_mem_left [EquivBEq α] [LawfulHashable α] + (h₁ : m₁.WF) (h₂ : m₂.WF) {k : α} + (not_mem : k ∉ m₁) : + k ∉ m₁ ∩ m₂ := by + simp only [Inter.inter] + rw [← contains_eq_false_iff_not_mem] at not_mem |- + revert not_mem + simp_to_raw using Raw₀.contains_inter_eq_false_of_contains_eq_false_left + +theorem not_mem_inter_of_not_mem_right [EquivBEq α] [LawfulHashable α] + (h₁ : m₁.WF) (h₂ : m₂.WF) {k : α} + (not_mem : k ∉ m₂) : + k ∉ m₁ ∩ m₂ := by + simp only [Inter.inter] + rw [← contains_eq_false_iff_not_mem] at not_mem |- + revert not_mem + simp_to_raw using Raw₀.contains_inter_eq_false_of_contains_eq_false_right + +/- get? -/ +theorem get?_inter [LawfulBEq α] (h₁ : m₁.WF) (h₂ : m₂.WF) {k : α} : + (m₁ ∩ m₂).get? k = + if k ∈ m₂ then m₁.get? k else none := by + simp only [Membership.mem, Inter.inter] + simp_to_raw using Raw₀.get?_inter + +theorem get?_inter_of_mem_right [LawfulBEq α] (h₁ : m₁.WF) (h₂ : m₂.WF) + {k : α} (mem : k ∈ m₂) : + (m₁ ∩ m₂).get? k = m₁.get? k := by + simp only [Membership.mem] at mem + simp only [Inter.inter] + revert mem + simp_to_raw using Raw₀.get?_inter_of_contains_right + +theorem get?_inter_of_not_mem_left [LawfulBEq α] (h₁ : m₁.WF) (h₂ : m₂.WF) + {k : α} (not_mem : k ∉ m₁) : + (m₁ ∩ m₂).get? k = none := by + rw [← contains_eq_false_iff_not_mem] at not_mem + revert not_mem + simp only [Inter.inter] + simp_to_raw using Raw₀.get?_inter_of_contains_eq_false_left + +theorem get?_inter_of_not_mem_right [LawfulBEq α] (h₁ : m₁.WF) (h₂ : m₂.WF) + {k : α} (not_mem : k ∉ m₂) : + (m₁ ∩ m₂).get? k = none := by + rw [← contains_eq_false_iff_not_mem] at not_mem + revert not_mem + simp only [Inter.inter] + simp_to_raw using Raw₀.get?_inter_of_contains_eq_false_right + +/- get -/ +theorem get_inter [LawfulBEq α] (h₁ : m₁.WF) (h₂ : m₂.WF) + {k : α} {h_mem: k ∈ m₁ ∩ m₂} : + (m₁ ∩ m₂).get k h_mem = + m₁.get k ((mem_inter_iff h₁ h₂).1 h_mem).1 := by + revert h_mem + simp only [Membership.mem, Inter.inter] + simp_to_raw + intro h_mem + apply Raw₀.get_inter + all_goals wf_trivial + +/- getD -/ +theorem getD_inter [LawfulBEq α] (h₁ : m₁.WF) (h₂ : m₂.WF) + {k : α} {fallback : β k} : + (m₁ ∩ m₂).getD k fallback = + if k ∈ m₂ then m₁.getD k fallback else fallback := by + simp only [Inter.inter, Membership.mem] + simp_to_raw using Raw₀.getD_inter + +theorem getD_inter_of_mem_right [LawfulBEq α] (h₁ : m₁.WF) (h₂ : m₂.WF) + {k : α} {fallback : β k} (mem : k ∈ m₂) : + (m₁ ∩ m₂).getD k fallback = m₁.getD k fallback := by + revert mem + simp only [Membership.mem, Inter.inter] + simp_to_raw using Raw₀.getD_inter_of_contains_right + +theorem getD_inter_of_not_mem_right [LawfulBEq α] (h₁ : m₁.WF) (h₂ : m₂.WF) + {k : α} {fallback : β k} (not_mem : k ∉ m₂) : + (m₁ ∩ m₂).getD k fallback = fallback := by + rw [← contains_eq_false_iff_not_mem] at not_mem + simp only [Inter.inter] + revert not_mem + simp_to_raw using Raw₀.getD_inter_of_contains_eq_false_right + +theorem getD_inter_of_not_mem_left [LawfulBEq α] (h₁ : m₁.WF) (h₂ : m₂.WF) + {k : α} {fallback : β k} (not_mem : k ∉ m₁) : + (m₁ ∩ m₂).getD k fallback = fallback := by + rw [← contains_eq_false_iff_not_mem] at not_mem + simp only [Inter.inter] + revert not_mem + simp_to_raw using Raw₀.getD_inter_of_contains_eq_false_left + +/- get! -/ +theorem get!_inter [LawfulBEq α] (h₁ : m₁.WF) (h₂ : m₂.WF) + {k : α} [Inhabited (β k)] : + (m₁ ∩ m₂).get! k = + if k ∈ m₂ then m₁.get! k else default := by + simp only [Inter.inter, Membership.mem] + simp_to_raw using Raw₀.get!_inter + +theorem get!_inter_of_mem_right [LawfulBEq α] (h₁ : m₁.WF) (h₂ : m₂.WF) + {k : α} [Inhabited (β k)] (mem : k ∈ m₂) : + (m₁ ∩ m₂).get! k = m₁.get! k := by + revert mem + simp only [Membership.mem, Inter.inter] + simp_to_raw using Raw₀.get!_inter_of_contains_right + +theorem get!_inter_of_not_mem_right [LawfulBEq α] (h₁ : m₁.WF) (h₂ : m₂.WF) + {k : α} [Inhabited (β k)] (not_mem : k ∉ m₂) : + (m₁ ∩ m₂).get! k = default := by + rw [← contains_eq_false_iff_not_mem] at not_mem + simp only [Inter.inter] + revert not_mem + simp_to_raw using Raw₀.get!_inter_of_contains_eq_false_right + +theorem get!_inter_of_not_mem_left [LawfulBEq α] (h₁ : m₁.WF) (h₂ : m₂.WF) + {k : α} [Inhabited (β k)] (not_mem : k ∉ m₁) : + (m₁ ∩ m₂).get! k = default := by + rw [← contains_eq_false_iff_not_mem] at not_mem + simp only [Inter.inter] + revert not_mem + simp_to_raw using Raw₀.get!_inter_of_contains_eq_false_left + +/- getKey? -/ +theorem getKey?_inter [EquivBEq α] [LawfulHashable α] + (h₁ : m₁.WF) (h₂ : m₂.WF) {k : α} : + (m₁ ∩ m₂).getKey? k = + if k ∈ m₂ then m₁.getKey? k else none := by + simp only [Membership.mem, Inter.inter] + simp_to_raw using Raw₀.getKey?_inter + +theorem getKey?_inter_of_mem_right [EquivBEq α] [LawfulHashable α] + (h₁ : m₁.WF) (h₂ : m₂.WF) {k : α} (mem : k ∈ m₂) : + (m₁ ∩ m₂).getKey? k = m₁.getKey? k := by + simp only [Membership.mem] at mem + simp only [Inter.inter] + revert mem + simp_to_raw using Raw₀.getKey?_inter_of_contains_right + +theorem getKey?_inter_of_not_mem_right [EquivBEq α] [LawfulHashable α] + (h₁ : m₁.WF) (h₂ : m₂.WF) {k : α} (not_mem : k ∉ m₂) : + (m₁ ∩ m₂).getKey? k = none := by + rw [← contains_eq_false_iff_not_mem] at not_mem + simp only [Inter.inter] + revert not_mem + simp_to_raw using Raw₀.getKey?_inter_of_contains_eq_false_right + +theorem getKey?_inter_of_not_mem_left [EquivBEq α] [LawfulHashable α] + (h₁ : m₁.WF) (h₂ : m₂.WF) {k : α} (not_mem : k ∉ m₁) : + (m₁ ∩ m₂).getKey? k = none := by + rw [← contains_eq_false_iff_not_mem] at not_mem + simp only [Inter.inter] + revert not_mem + simp_to_raw using Raw₀.getKey?_inter_of_contains_eq_false_left + +/- getKey -/ +theorem getKey_inter [EquivBEq α] [LawfulHashable α] (h₁ : m₁.WF) (h₂ : m₂.WF) + {k : α} {h_mem : k ∈ m₁ ∩ m₂} : + (m₁ ∩ m₂).getKey k h_mem = + m₁.getKey k ((mem_inter_iff h₁ h₂).1 h_mem).1 := by + revert h_mem + simp only [Membership.mem, Inter.inter] + simp_to_raw + intro h_mem + apply Raw₀.getKey_inter + all_goals wf_trivial + +/- getKeyD -/ +theorem getKeyD_inter [EquivBEq α] [LawfulHashable α] (h₁ : m₁.WF) + (h₂ : m₂.WF) {k fallback : α} : + (m₁ ∩ m₂).getKeyD k fallback = + if k ∈ m₂ then m₁.getKeyD k fallback else fallback := by + simp only [Membership.mem, Inter.inter] + simp_to_raw using Raw₀.getKeyD_inter + +theorem getKeyD_inter_of_mem_right [EquivBEq α] [LawfulHashable α] (h₁ : m₁.WF) + (h₂ : m₂.WF) {k fallback : α} (mem : k ∈ m₂) : + (m₁ ∩ m₂).getKeyD k fallback = m₁.getKeyD k fallback := by + simp only [Membership.mem] at mem + simp only [Inter.inter] + revert mem + simp_to_raw using Raw₀.getKeyD_inter_of_contains_right + +theorem getKeyD_inter_of_not_mem_right [EquivBEq α] [LawfulHashable α] (h₁ : m₁.WF) + (h₂ : m₂.WF) {k fallback : α} (not_mem : k ∉ m₂) : + (m₁ ∩ m₂).getKeyD k fallback = fallback := by + rw [← contains_eq_false_iff_not_mem] at not_mem + simp only [Inter.inter] + revert not_mem + simp_to_raw using Raw₀.getKeyD_inter_of_contains_eq_false_right + +theorem getKeyD_inter_of_not_mem_left [EquivBEq α] [LawfulHashable α] (h₁ : m₁.WF) + (h₂ : m₂.WF) {k fallback : α} (not_mem : k ∉ m₁) : + (m₁ ∩ m₂).getKeyD k fallback = fallback := by + rw [← contains_eq_false_iff_not_mem] at not_mem + simp only [Inter.inter] + revert not_mem + simp_to_raw using Raw₀.getKeyD_inter_of_contains_eq_false_left + +/- getKey! -/ +theorem getKey!_inter [EquivBEq α] [LawfulHashable α] [Inhabited α] (h₁ : m₁.WF) + (h₂ : m₂.WF) {k : α} : + (m₁ ∩ m₂).getKey! k = + if k ∈ m₂ then m₁.getKey! k else default := by + simp only [Membership.mem, Inter.inter] + simp_to_raw using Raw₀.getKey!_inter + +theorem getKey!_inter_of_mem_right [EquivBEq α] [LawfulHashable α] [Inhabited α] (h₁ : m₁.WF) + (h₂ : m₂.WF) {k : α} (mem : k ∈ m₂) : + (m₁ ∩ m₂).getKey! k = m₁.getKey! k := by + simp only [Membership.mem] at mem + simp only [Inter.inter] + revert mem + simp_to_raw using Raw₀.getKey!_inter_of_contains_right + +theorem getKey!_inter_of_not_mem_right [EquivBEq α] [LawfulHashable α] [Inhabited α] (h₁ : m₁.WF) + (h₂ : m₂.WF) {k : α} (not_mem : k ∉ m₂) : + (m₁ ∩ m₂).getKey! k = default := by + rw [← contains_eq_false_iff_not_mem] at not_mem + simp only [Inter.inter] + revert not_mem + simp_to_raw using Raw₀.getKey!_inter_of_contains_eq_false_right + +theorem getKey!_inter_of_not_mem_left [EquivBEq α] [LawfulHashable α] [Inhabited α] (h₁ : m₁.WF) + (h₂ : m₂.WF) {k : α} (not_mem : k ∉ m₁) : + (m₁ ∩ m₂).getKey! k = default := by + rw [← contains_eq_false_iff_not_mem] at not_mem + simp only [Inter.inter] + revert not_mem + simp_to_raw using Raw₀.getKey!_inter_of_contains_eq_false_left + +/- size -/ +theorem size_inter_le_size_left [EquivBEq α] [LawfulHashable α] + (h₁ : m₁.WF) (h₂ : m₂.WF) : + (m₁ ∩ m₂).size ≤ m₁.size := by + simp only [Inter.inter] + simp_to_raw using Raw₀.size_inter_le_size_left + +theorem size_inter_le_size_right [EquivBEq α] [LawfulHashable α] + (h₁ : m₁.WF) (h₂ : m₂.WF) : + (m₁ ∩ m₂).size ≤ m₂.size := by + simp only [Inter.inter] + simp_to_raw using Raw₀.size_inter_le_size_right + +theorem size_inter_eq_size_left [EquivBEq α] [LawfulHashable α] + (h₁ : m₁.WF) (h₂ : m₂.WF) + (h : ∀ (a : α), a ∈ m₁ → a ∈ m₂) : + (m₁ ∩ m₂).size = m₁.size := by + have : ∀ (a : α), m₁.contains a → m₂.contains a := by + intro a ha + rw [contains_iff_mem] at ha ⊢ + exact h a ha + revert this + simp only [Inter.inter] + simp_to_raw using Raw₀.size_inter_eq_size_left + +theorem size_inter_eq_size_right [EquivBEq α] [LawfulHashable α] + (h₁ : m₁.WF) (h₂ : m₂.WF) + (h : ∀ (a : α), a ∈ m₂ → a ∈ m₁) : + (m₁ ∩ m₂).size = m₂.size := by + have : ∀ (a : α), m₂.contains a → m₁.contains a := by + intro a ha + rw [contains_iff_mem] at ha ⊢ + exact h a ha + revert this + simp only [Inter.inter] + simp_to_raw using Raw₀.size_inter_eq_size_right + +theorem size_add_size_eq_size_union_add_size_inter [EquivBEq α] [LawfulHashable α] + (h₁ : m₁.WF) (h₂ : m₂.WF) : + m₁.size + m₂.size = (m₁ ∪ m₂).size + (m₁ ∩ m₂).size := by + simp only [Union.union, Inter.inter] + simp_to_raw using @Raw₀.size_add_size_eq_size_union_add_size_inter _ _ _ _ ⟨m₁, _⟩ ⟨m₂, _⟩ + +/- isEmpty -/ +@[simp] +theorem isEmpty_inter_left [EquivBEq α] [LawfulHashable α] (h₁ : m₁.WF) (h₂ : m₂.WF) (h : m₁.isEmpty) : + (m₁ ∩ m₂).isEmpty = true := by + revert h + simp only [Inter.inter] + simp_to_raw using @Raw₀.isEmpty_inter_left _ _ _ _ ⟨m₁, _⟩ ⟨m₂, _⟩ _ _ + +@[simp] +theorem isEmpty_inter_right [EquivBEq α] [LawfulHashable α] (h₁ : m₁.WF) (h₂ : m₂.WF) (h : m₂.isEmpty) : + (m₁ ∩ m₂).isEmpty = true := by + revert h + simp only [Inter.inter] + simp_to_raw using @Raw₀.isEmpty_inter_right _ _ _ _ ⟨m₁, _⟩ ⟨m₂, _⟩ _ _ + +theorem isEmpty_inter_iff [EquivBEq α] [LawfulHashable α] (h₁ : m₁.WF) (h₂ : m₂.WF) : + (m₁ ∩ m₂).isEmpty ↔ ∀ k, k ∈ m₁ → k ∉ m₂ := by + conv => + rhs + rhs + rhs + rw [← contains_eq_false_iff_not_mem] + simp only [Inter.inter, Membership.mem] + simp_to_raw using @Raw₀.isEmpty_inter_iff _ _ _ _ ⟨m₁, _⟩ ⟨m₂, _⟩ _ _ + +end Inter + namespace Const variable {β : Type v} {m₁ m₂ : Raw α (fun _ => β)} diff --git a/src/Std/Data/Internal/List/Associative.lean b/src/Std/Data/Internal/List/Associative.lean index fe712c53e3b9..b026eecd46cf 100644 --- a/src/Std/Data/Internal/List/Associative.lean +++ b/src/Std/Data/Internal/List/Associative.lean @@ -5464,14 +5464,14 @@ theorem List.getKey_filter_containsKey [BEq α] [EquivBEq α] · exact h₁.2 · exact dl₁ -theorem List.not_contains_filter_containsKey_of_containsKey_eq_false_left [BEq α] [EquivBEq α] {l₁ l₂ : List ((a : α) × β a)} {hl₁ : DistinctKeys l₁} {k : α} : +theorem List.containsKey_filter_containsKey_eq_false_of_containsKey_eq_false_left [BEq α] [EquivBEq α] {l₁ l₂ : List ((a : α) × β a)} {hl₁ : DistinctKeys l₁} {k : α} : containsKey k l₁ = false → containsKey k (List.filter (fun p => containsKey p.fst l₂) l₁) = false := by intro h rw [containsKey_filter_containsKey] . simp [h] . exact hl₁ -theorem List.not_contains_filter_containsKey_of_containsKey_eq_false_right [BEq α] [EquivBEq α] {l₁ l₂ : List ((a : α) × β a)} {hl₁ : DistinctKeys l₁} {k : α} : +theorem List.containsKey_filter_containsKey_eq_false_of_containsKey_eq_false_right [BEq α] [EquivBEq α] {l₁ l₂ : List ((a : α) × β a)} {hl₁ : DistinctKeys l₁} {k : α} : containsKey k l₂ = false → containsKey k (List.filter (fun p => containsKey p.fst l₂) l₁) = false := by intro h rw [containsKey_filter_containsKey] @@ -5501,7 +5501,7 @@ theorem getValue?_filter_containsKey_eq_if_containsKey {β : Type v} [BEq α] [E case isFalse h => simp only [Option.map_eq_none_iff, getEntry?_eq_none] simp only [Bool.not_eq_true] at h - apply List.not_contains_filter_containsKey_of_containsKey_eq_false_right + apply List.containsKey_filter_containsKey_eq_false_of_containsKey_eq_false_right . exact dl₁ . exact h @@ -5548,7 +5548,7 @@ theorem List.getValueCastD_filter_containsKey [BEq α] [LawfulBEq α] apply List.getValueCast?_filter_containsKey_of_containsKey_right dl₁ h case isFalse h => apply getValueCastD_eq_fallback - apply List.not_contains_filter_containsKey_of_containsKey_eq_false_right + apply List.containsKey_filter_containsKey_eq_false_of_containsKey_eq_false_right . exact dl₁ . simp only [h] @@ -5568,7 +5568,7 @@ theorem List.getValueCastD_filter_containsKey_of_containsKey_eq_false_left [BEq containsKey k l₁ = false → getValueCastD k (List.filter (fun p => containsKey p.fst l₂) l₁) fallback = fallback := by intro h apply getValueCastD_eq_fallback - apply List.not_contains_filter_containsKey_of_containsKey_eq_false_left + apply List.containsKey_filter_containsKey_eq_false_of_containsKey_eq_false_left . exact dl₁ . exact h @@ -5578,7 +5578,7 @@ theorem List.getValueCastD_filter_containsKey_of_containsKey_eq_false_right [BEq containsKey k l₂ = false → getValueCastD k (List.filter (fun p => containsKey p.fst l₂) l₁) fallback = fallback := by intro h apply getValueCastD_eq_fallback - apply List.not_contains_filter_containsKey_of_containsKey_eq_false_right + apply List.containsKey_filter_containsKey_eq_false_of_containsKey_eq_false_right . exact dl₁ . exact h @@ -5594,7 +5594,7 @@ theorem List.getKeyD_filter_containsKey [BEq α] [EquivBEq α] apply List.getKey?_filter_containsKey_of_containsKey_right dl₁ h case isFalse h => apply getKeyD_eq_fallback - apply List.not_contains_filter_containsKey_of_containsKey_eq_false_right + apply List.containsKey_filter_containsKey_eq_false_of_containsKey_eq_false_right · exact dl₁ · simp only [h] @@ -5614,7 +5614,7 @@ theorem List.getKeyD_filter_containsKey_of_containsKey_eq_false_left [BEq α] [E containsKey k l₁ = false → getKeyD k (List.filter (fun p => containsKey p.fst l₂) l₁) fallback = fallback := by intro h apply getKeyD_eq_fallback - apply List.not_contains_filter_containsKey_of_containsKey_eq_false_left + apply List.containsKey_filter_containsKey_eq_false_of_containsKey_eq_false_left · exact dl₁ · exact h @@ -5624,7 +5624,7 @@ theorem List.getKeyD_filter_containsKey_of_containsKey_eq_false_right [BEq α] [ containsKey k l₂ = false → getKeyD k (List.filter (fun p => containsKey p.fst l₂) l₁) fallback = fallback := by intro h apply getKeyD_eq_fallback - apply List.not_contains_filter_containsKey_of_containsKey_eq_false_right + apply List.containsKey_filter_containsKey_eq_false_of_containsKey_eq_false_right · exact dl₁ · exact h @@ -5652,7 +5652,7 @@ theorem getValueD_filter_containsKey_eq_if_containsKey {β : Type v} [BEq α] [E apply getValue?_filter_containsKey_eq_containsKey dl₁ h case isFalse h => apply getValueD_eq_fallback - apply List.not_contains_filter_containsKey_of_containsKey_eq_false_right + apply List.containsKey_filter_containsKey_eq_false_of_containsKey_eq_false_right · exact dl₁ · simp only [h] @@ -5672,7 +5672,7 @@ theorem List.getValueD_filter_of_containsKey_eq_false_left {β : Type v} [BEq α containsKey k l₁ = false → getValueD k (List.filter (fun p => containsKey p.fst l₂) l₁) fallback = fallback := by intro h apply getValueD_eq_fallback - apply List.not_contains_filter_containsKey_of_containsKey_eq_false_left + apply List.containsKey_filter_containsKey_eq_false_of_containsKey_eq_false_left · exact dl₁ · exact h @@ -5682,7 +5682,7 @@ theorem List.getValueD_filter_of_containsKey_eq_false_right {β : Type v} [BEq containsKey k l₂ = false → getValueD k (List.filter (fun p => containsKey p.fst l₂) l₁) fallback = fallback := by intro h apply getValueD_eq_fallback - apply List.not_contains_filter_containsKey_of_containsKey_eq_false_right + apply List.containsKey_filter_containsKey_eq_false_of_containsKey_eq_false_right · exact dl₁ · exact h From 95fab3b1d31ca2933c9e46fe2d83ceac22a7b111 Mon Sep 17 00:00:00 2001 From: Wojciech Rozowski Date: Thu, 13 Nov 2025 09:45:50 +0100 Subject: [PATCH 41/88] Add DHashMap lemmas --- src/Std/Data/DHashMap/Basic.lean | 1 + src/Std/Data/DHashMap/Lemmas.lean | 241 ++++++++++++++++++++++++++++++ 2 files changed, 242 insertions(+) diff --git a/src/Std/Data/DHashMap/Basic.lean b/src/Std/Data/DHashMap/Basic.lean index ff3401b36266..3be9b0625450 100644 --- a/src/Std/Data/DHashMap/Basic.lean +++ b/src/Std/Data/DHashMap/Basic.lean @@ -342,6 +342,7 @@ This function always merges the smaller map into the larger map, so the expected wf := Std.DHashMap.Raw.WF.inter₀ m₁.2 m₂.2 instance [BEq α] [Hashable α] : Union (DHashMap α β) := ⟨union⟩ +instance [BEq α] [Hashable α] : Inter (DHashMap α β) := ⟨inter⟩ section Unverified diff --git a/src/Std/Data/DHashMap/Lemmas.lean b/src/Std/Data/DHashMap/Lemmas.lean index 7b396562dacd..435b75cc704f 100644 --- a/src/Std/Data/DHashMap/Lemmas.lean +++ b/src/Std/Data/DHashMap/Lemmas.lean @@ -1932,6 +1932,247 @@ theorem isEmpty_union [EquivBEq α] [LawfulHashable α] : end Union +section Inter + +variable (m₁ m₂ : DHashMap α β) + +variable {m₁ m₂} + +@[simp] +theorem inter_eq : m₁.inter m₂ = m₁ ∩ m₂ := by + simp only [Inter.inter] + +/- contains -/ +@[simp] +theorem contains_inter [EquivBEq α] [LawfulHashable α] {k : α} : + (m₁ ∩ m₂).contains k = (m₁.contains k && m₂.contains k) := + @Raw₀.contains_inter _ _ _ _ ⟨m₁.1, m₁.2.size_buckets_pos⟩ ⟨m₂.1, m₂.2.size_buckets_pos⟩ _ _ m₁.2 m₂.2 k + +/- mem -/ +@[simp] +theorem mem_inter_iff [EquivBEq α] [LawfulHashable α] {k : α} : + k ∈ m₁ ∩ m₂ ↔ k ∈ m₁ ∧ k ∈ m₂ := + @Raw₀.contains_inter_iff _ _ _ _ ⟨m₁.1, _⟩ ⟨m₂.1, _⟩ _ _ m₁.2 m₂.2 k + +theorem not_mem_inter_of_not_mem_left [EquivBEq α] [LawfulHashable α] {k : α} + (not_mem : k ∉ m₁) : + k ∉ m₁ ∩ m₂ := by + rw [← contains_eq_false_iff_not_mem] at not_mem ⊢ + exact @Raw₀.contains_inter_eq_false_of_contains_eq_false_left _ _ _ _ ⟨m₁.1, m₁.2.size_buckets_pos⟩ ⟨m₂.1, m₂.2.size_buckets_pos⟩ _ _ m₁.2 m₂.2 k not_mem + +theorem not_mem_inter_of_not_mem_right [EquivBEq α] [LawfulHashable α] {k : α} + (not_mem : k ∉ m₂) : + k ∉ m₁ ∩ m₂ := by + rw [← contains_eq_false_iff_not_mem] at not_mem ⊢ + exact @Raw₀.contains_inter_eq_false_of_contains_eq_false_right _ _ _ _ ⟨m₁.1, m₁.2.size_buckets_pos⟩ ⟨m₂.1, m₂.2.size_buckets_pos⟩ _ _ m₁.2 m₂.2 k not_mem + +/- get? -/ +theorem get?_inter [LawfulBEq α] {k : α} : + (m₁ ∩ m₂).get? k = + if k ∈ m₂ then m₁.get? k else none := + @Raw₀.get?_inter _ _ _ _ ⟨m₁.1, m₁.2.size_buckets_pos⟩ ⟨m₂.1, m₂.2.size_buckets_pos⟩ _ m₁.2 m₂.2 k + +theorem get?_inter_of_mem_right [LawfulBEq α] + {k : α} (mem : k ∈ m₂) : + (m₁ ∩ m₂).get? k = m₁.get? k := + @Raw₀.get?_inter_of_contains_right _ _ _ _ ⟨m₁.1, m₁.2.size_buckets_pos⟩ ⟨m₂.1, m₂.2.size_buckets_pos⟩ _ m₁.2 m₂.2 k mem + +theorem get?_inter_of_not_mem_left [LawfulBEq α] + {k : α} (not_mem : k ∉ m₁) : + (m₁ ∩ m₂).get? k = none := by + rw [← contains_eq_false_iff_not_mem] at not_mem + exact @Raw₀.get?_inter_of_contains_eq_false_left _ _ _ _ ⟨m₁.1, m₁.2.size_buckets_pos⟩ ⟨m₂.1, m₂.2.size_buckets_pos⟩ _ m₁.2 m₂.2 k not_mem + +theorem get?_inter_of_not_mem_right [LawfulBEq α] + {k : α} (not_mem : k ∉ m₂) : + (m₁ ∩ m₂).get? k = none := by + rw [← contains_eq_false_iff_not_mem] at not_mem + exact @Raw₀.get?_inter_of_contains_eq_false_right _ _ _ _ ⟨m₁.1, m₁.2.size_buckets_pos⟩ ⟨m₂.1, m₂.2.size_buckets_pos⟩ _ m₁.2 m₂.2 k not_mem + +/- get -/ +theorem get_inter [LawfulBEq α] + {k : α} {h_mem : k ∈ m₁ ∩ m₂} : + (m₁ ∩ m₂).get k h_mem = + m₁.get k ((mem_inter_iff.1 h_mem).1) := by + rw [mem_iff_contains] at h_mem + exact @Raw₀.get_inter _ _ _ _ ⟨m₁.1, _⟩ ⟨m₂.1, _⟩ _ m₁.2 m₂.2 k h_mem + +/- getD -/ +theorem getD_inter [LawfulBEq α] {k : α} {fallback : β k} : + (m₁ ∩ m₂).getD k fallback = + if k ∈ m₂ then m₁.getD k fallback else fallback := + @Raw₀.getD_inter _ _ _ _ ⟨m₁.1, m₁.2.size_buckets_pos⟩ ⟨m₂.1, m₂.2.size_buckets_pos⟩ _ m₁.2 m₂.2 k fallback + +theorem getD_inter_of_mem_right [LawfulBEq α] + {k : α} {fallback : β k} (mem : k ∈ m₂) : + (m₁ ∩ m₂).getD k fallback = m₁.getD k fallback := + @Raw₀.getD_inter_of_contains_right _ _ _ _ ⟨m₁.1, m₁.2.size_buckets_pos⟩ ⟨m₂.1, m₂.2.size_buckets_pos⟩ _ m₁.2 m₂.2 k fallback mem + +theorem getD_inter_of_not_mem_right [LawfulBEq α] + {k : α} {fallback : β k} (not_mem : k ∉ m₂) : + (m₁ ∩ m₂).getD k fallback = fallback := by + rw [← contains_eq_false_iff_not_mem] at not_mem + exact @Raw₀.getD_inter_of_contains_eq_false_right _ _ _ _ ⟨m₁.1, m₁.2.size_buckets_pos⟩ ⟨m₂.1, m₂.2.size_buckets_pos⟩ _ m₁.2 m₂.2 k fallback not_mem + +theorem getD_inter_of_not_mem_left [LawfulBEq α] + {k : α} {fallback : β k} (not_mem : k ∉ m₁) : + (m₁ ∩ m₂).getD k fallback = fallback := by + rw [← contains_eq_false_iff_not_mem] at not_mem + exact @Raw₀.getD_inter_of_contains_eq_false_left _ _ _ _ ⟨m₁.1, m₁.2.size_buckets_pos⟩ ⟨m₂.1, m₂.2.size_buckets_pos⟩ _ m₁.2 m₂.2 k fallback not_mem + +/- get! -/ +theorem get!_inter [LawfulBEq α] {k : α} [Inhabited (β k)] : + (m₁ ∩ m₂).get! k = + if k ∈ m₂ then m₁.get! k else default := + @Raw₀.get!_inter _ _ _ _ ⟨m₁.1, m₁.2.size_buckets_pos⟩ ⟨m₂.1, m₂.2.size_buckets_pos⟩ _ m₁.2 m₂.2 k _ + +theorem get!_inter_of_mem_right [LawfulBEq α] + {k : α} [Inhabited (β k)] (mem : k ∈ m₂) : + (m₁ ∩ m₂).get! k = m₁.get! k := + @Raw₀.get!_inter_of_contains_right _ _ _ _ ⟨m₁.1, m₁.2.size_buckets_pos⟩ ⟨m₂.1, m₂.2.size_buckets_pos⟩ _ m₁.2 m₂.2 k _ mem + +theorem get!_inter_of_not_mem_right [LawfulBEq α] + {k : α} [Inhabited (β k)] (not_mem : k ∉ m₂) : + (m₁ ∩ m₂).get! k = default := by + rw [← contains_eq_false_iff_not_mem] at not_mem + exact @Raw₀.get!_inter_of_contains_eq_false_right _ _ _ _ ⟨m₁.1, m₁.2.size_buckets_pos⟩ ⟨m₂.1, m₂.2.size_buckets_pos⟩ _ m₁.2 m₂.2 k _ not_mem + +theorem get!_inter_of_not_mem_left [LawfulBEq α] + {k : α} [Inhabited (β k)] (not_mem : k ∉ m₁) : + (m₁ ∩ m₂).get! k = default := by + rw [← contains_eq_false_iff_not_mem] at not_mem + exact @Raw₀.get!_inter_of_contains_eq_false_left _ _ _ _ ⟨m₁.1, m₁.2.size_buckets_pos⟩ ⟨m₂.1, m₂.2.size_buckets_pos⟩ _ m₁.2 m₂.2 k _ not_mem + +/- getKey? -/ +theorem getKey?_inter [EquivBEq α] [LawfulHashable α] {k : α} : + (m₁ ∩ m₂).getKey? k = + if k ∈ m₂ then m₁.getKey? k else none := + @Raw₀.getKey?_inter _ _ _ _ ⟨m₁.1, m₁.2.size_buckets_pos⟩ ⟨m₂.1, m₂.2.size_buckets_pos⟩ _ _ m₁.2 m₂.2 k + +theorem getKey?_inter_of_mem_right [EquivBEq α] [LawfulHashable α] + {k : α} (mem : k ∈ m₂) : + (m₁ ∩ m₂).getKey? k = m₁.getKey? k := + @Raw₀.getKey?_inter_of_contains_right _ _ _ _ ⟨m₁.1, m₁.2.size_buckets_pos⟩ ⟨m₂.1, m₂.2.size_buckets_pos⟩ _ _ m₁.2 m₂.2 k mem + +theorem getKey?_inter_of_not_mem_right [EquivBEq α] [LawfulHashable α] + {k : α} (not_mem : k ∉ m₂) : + (m₁ ∩ m₂).getKey? k = none := by + rw [← contains_eq_false_iff_not_mem] at not_mem + exact @Raw₀.getKey?_inter_of_contains_eq_false_right _ _ _ _ ⟨m₁.1, m₁.2.size_buckets_pos⟩ ⟨m₂.1, m₂.2.size_buckets_pos⟩ _ _ m₁.2 m₂.2 k not_mem + +theorem getKey?_inter_of_not_mem_left [EquivBEq α] [LawfulHashable α] + {k : α} (not_mem : k ∉ m₁) : + (m₁ ∩ m₂).getKey? k = none := by + rw [← contains_eq_false_iff_not_mem] at not_mem + exact @Raw₀.getKey?_inter_of_contains_eq_false_left _ _ _ _ ⟨m₁.1, m₁.2.size_buckets_pos⟩ ⟨m₂.1, m₂.2.size_buckets_pos⟩ _ _ m₁.2 m₂.2 k not_mem + +/- getKey -/ +theorem getKey_inter [EquivBEq α] [LawfulHashable α] + {k : α} {h_mem : k ∈ m₁ ∩ m₂} : + (m₁ ∩ m₂).getKey k h_mem = + m₁.getKey k ((mem_inter_iff.1 h_mem).1) := by + rw [mem_iff_contains] at h_mem + exact @Raw₀.getKey_inter _ _ _ _ ⟨m₁.1, _⟩ ⟨m₂.1, _⟩ _ _ m₁.2 m₂.2 k h_mem + +/- getKeyD -/ +theorem getKeyD_inter [EquivBEq α] [LawfulHashable α] {k fallback : α} : + (m₁ ∩ m₂).getKeyD k fallback = + if k ∈ m₂ then m₁.getKeyD k fallback else fallback := + @Raw₀.getKeyD_inter _ _ _ _ ⟨m₁.1, m₁.2.size_buckets_pos⟩ ⟨m₂.1, m₂.2.size_buckets_pos⟩ _ _ m₁.2 m₂.2 k fallback + +theorem getKeyD_inter_of_mem_right [EquivBEq α] [LawfulHashable α] + {k fallback : α} (mem : k ∈ m₂) : + (m₁ ∩ m₂).getKeyD k fallback = m₁.getKeyD k fallback := + @Raw₀.getKeyD_inter_of_contains_right _ _ _ _ ⟨m₁.1, m₁.2.size_buckets_pos⟩ ⟨m₂.1, m₂.2.size_buckets_pos⟩ _ _ m₁.2 m₂.2 k fallback mem + +theorem getKeyD_inter_of_not_mem_right [EquivBEq α] [LawfulHashable α] + {k fallback : α} (not_mem : k ∉ m₂) : + (m₁ ∩ m₂).getKeyD k fallback = fallback := by + rw [← contains_eq_false_iff_not_mem] at not_mem + exact @Raw₀.getKeyD_inter_of_contains_eq_false_right _ _ _ _ ⟨m₁.1, m₁.2.size_buckets_pos⟩ ⟨m₂.1, m₂.2.size_buckets_pos⟩ _ _ m₁.2 m₂.2 k fallback not_mem + +theorem getKeyD_inter_of_not_mem_left [EquivBEq α] [LawfulHashable α] + {k fallback : α} (not_mem : k ∉ m₁) : + (m₁ ∩ m₂).getKeyD k fallback = fallback := by + rw [← contains_eq_false_iff_not_mem] at not_mem + exact @Raw₀.getKeyD_inter_of_contains_eq_false_left _ _ _ _ ⟨m₁.1, m₁.2.size_buckets_pos⟩ ⟨m₂.1, m₂.2.size_buckets_pos⟩ _ _ m₁.2 m₂.2 k fallback not_mem + +/- getKey! -/ +theorem getKey!_inter [EquivBEq α] [LawfulHashable α] [Inhabited α] {k : α} : + (m₁ ∩ m₂).getKey! k = + if k ∈ m₂ then m₁.getKey! k else default := + @Raw₀.getKey!_inter _ _ _ _ ⟨m₁.1, m₁.2.size_buckets_pos⟩ ⟨m₂.1, m₂.2.size_buckets_pos⟩ _ _ _ m₁.2 m₂.2 k _ + +theorem getKey!_inter_of_mem_right [EquivBEq α] [LawfulHashable α] [Inhabited α] + {k : α} (mem : k ∈ m₂) : + (m₁ ∩ m₂).getKey! k = m₁.getKey! k := + @Raw₀.getKey!_inter_of_contains_right _ _ _ _ ⟨m₁.1, m₁.2.size_buckets_pos⟩ ⟨m₂.1, m₂.2.size_buckets_pos⟩ _ _ _ m₁.2 m₂.2 k mem + +theorem getKey!_inter_of_not_mem_right [EquivBEq α] [LawfulHashable α] [Inhabited α] + {k : α} (not_mem : k ∉ m₂) : + (m₁ ∩ m₂).getKey! k = default := by + rw [← contains_eq_false_iff_not_mem] at not_mem + exact @Raw₀.getKey!_inter_of_contains_eq_false_right _ _ _ _ ⟨m₁.1, m₁.2.size_buckets_pos⟩ ⟨m₂.1, m₂.2.size_buckets_pos⟩ _ _ _ m₁.2 m₂.2 k not_mem + +theorem getKey!_inter_of_not_mem_left [EquivBEq α] [LawfulHashable α] [Inhabited α] + {k : α} (not_mem : k ∉ m₁) : + (m₁ ∩ m₂).getKey! k = default := by + rw [← contains_eq_false_iff_not_mem] at not_mem + exact @Raw₀.getKey!_inter_of_contains_eq_false_left _ _ _ _ ⟨m₁.1, m₁.2.size_buckets_pos⟩ ⟨m₂.1, m₂.2.size_buckets_pos⟩ _ _ _ m₁.2 m₂.2 k not_mem + +/- size -/ +theorem size_inter_le_size_left [EquivBEq α] [LawfulHashable α] : + (m₁ ∩ m₂).size ≤ m₁.size := + @Raw₀.size_inter_le_size_left _ _ _ _ ⟨m₁.1, m₁.2.size_buckets_pos⟩ ⟨m₂.1, m₂.2.size_buckets_pos⟩ _ _ m₁.2 m₂.2 + +theorem size_inter_le_size_right [EquivBEq α] [LawfulHashable α] : + (m₁ ∩ m₂).size ≤ m₂.size := + @Raw₀.size_inter_le_size_right _ _ _ _ ⟨m₁.1, m₁.2.size_buckets_pos⟩ ⟨m₂.1, m₂.2.size_buckets_pos⟩ _ _ m₁.2 m₂.2 + +theorem size_inter_eq_size_left [EquivBEq α] [LawfulHashable α] + (h : ∀ (a : α), a ∈ m₁ → a ∈ m₂) : + (m₁ ∩ m₂).size = m₁.size := by + have : ∀ (a : α), m₁.contains a → m₂.contains a := by + intro a ha + rw [contains_iff_mem] at ha ⊢ + exact h a ha + exact @Raw₀.size_inter_eq_size_left _ _ _ _ ⟨m₁.1, m₁.2.size_buckets_pos⟩ ⟨m₂.1, m₂.2.size_buckets_pos⟩ _ _ m₁.2 m₂.2 this + +theorem size_inter_eq_size_right [EquivBEq α] [LawfulHashable α] + (h : ∀ (a : α), a ∈ m₂ → a ∈ m₁) : + (m₁ ∩ m₂).size = m₂.size := by + have : ∀ (a : α), m₂.contains a → m₁.contains a := by + intro a ha + rw [contains_iff_mem] at ha ⊢ + exact h a ha + exact @Raw₀.size_inter_eq_size_right _ _ _ _ ⟨m₁.1, m₁.2.size_buckets_pos⟩ ⟨m₂.1, m₂.2.size_buckets_pos⟩ _ _ m₁.2 m₂.2 this + +theorem size_add_size_eq_size_union_add_size_inter [EquivBEq α] [LawfulHashable α] : + m₁.size + m₂.size = (m₁ ∪ m₂).size + (m₁ ∩ m₂).size := + @Raw₀.size_add_size_eq_size_union_add_size_inter _ _ _ _ ⟨m₁.1, m₁.2.size_buckets_pos⟩ ⟨m₂.1, m₂.2.size_buckets_pos⟩ _ _ m₁.2 m₂.2 + +/- isEmpty -/ +@[simp] +theorem isEmpty_inter_left [EquivBEq α] [LawfulHashable α] (h : m₁.isEmpty) : + (m₁ ∩ m₂).isEmpty = true := + @Raw₀.isEmpty_inter_left _ _ _ _ ⟨m₁.1, m₁.2.size_buckets_pos⟩ ⟨m₂.1, m₂.2.size_buckets_pos⟩ _ _ m₁.2 m₂.2 h + +@[simp] +theorem isEmpty_inter_right [EquivBEq α] [LawfulHashable α] (h : m₂.isEmpty) : + (m₁ ∩ m₂).isEmpty = true := + @Raw₀.isEmpty_inter_right _ _ _ _ ⟨m₁.1, m₁.2.size_buckets_pos⟩ ⟨m₂.1, m₂.2.size_buckets_pos⟩ _ _ m₁.2 m₂.2 h + +theorem isEmpty_inter_iff [EquivBEq α] [LawfulHashable α] : + (m₁ ∩ m₂).isEmpty ↔ ∀ k, k ∈ m₁ → k ∉ m₂ := by + conv => + rhs + ext x + rhs + rw [← contains_eq_false_iff_not_mem] + exact @Raw₀.isEmpty_inter_iff _ _ _ _ ⟨m₁.1, m₁.2.size_buckets_pos⟩ ⟨m₂.1, m₂.2.size_buckets_pos⟩ _ _ m₁.wf m₂.wf + +end Inter + namespace Const variable {β : Type v} {m₁ m₂ : DHashMap α (fun _ => β)} From 81efbeefe7cc25723d778dbaa276b2945ae9bf06 Mon Sep 17 00:00:00 2001 From: Wojciech Rozowski Date: Thu, 13 Nov 2025 10:25:20 +0100 Subject: [PATCH 42/88] Add basic support for HashMap raw --- src/Std/Data/DHashMap/Internal/RawLemmas.lean | 1 + src/Std/Data/HashMap/Basic.lean | 5 +++ src/Std/Data/HashMap/Raw.lean | 5 +++ src/Std/Data/HashMap/RawLemmas.lean | 36 +++++++++++++++++++ 4 files changed, 47 insertions(+) diff --git a/src/Std/Data/DHashMap/Internal/RawLemmas.lean b/src/Std/Data/DHashMap/Internal/RawLemmas.lean index 68d788fcf149..547848631cda 100644 --- a/src/Std/Data/DHashMap/Internal/RawLemmas.lean +++ b/src/Std/Data/DHashMap/Internal/RawLemmas.lean @@ -2673,6 +2673,7 @@ theorem size_union_le_size_add_size [EquivBEq α] [LawfulHashable α] theorem isEmpty_union [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) : (m₁.union m₂).1.isEmpty = (m₁.1.isEmpty && m₂.1.isEmpty) := by simp_to_model [isEmpty, union] using List.isEmpty_insertList + end Union namespace Const diff --git a/src/Std/Data/HashMap/Basic.lean b/src/Std/Data/HashMap/Basic.lean index af5950f5e21b..94b929b4f1fb 100644 --- a/src/Std/Data/HashMap/Basic.lean +++ b/src/Std/Data/HashMap/Basic.lean @@ -264,6 +264,11 @@ This function always merges the smaller map into the larger map, so the expected instance [BEq α] [Hashable α] : Union (HashMap α β) := ⟨union⟩ +@[inherit_doc DHashMap.inter, inline] def inter [BEq α] [Hashable α] (m₁ m₂ : HashMap α β) : HashMap α β := + ⟨DHashMap.inter m₁.inner m₂.inner⟩ + +instance [BEq α] [Hashable α] : Inter (HashMap α β) := ⟨inter⟩ + section Unverified /-! We currently do not provide lemmas for the functions below. -/ diff --git a/src/Std/Data/HashMap/Raw.lean b/src/Std/Data/HashMap/Raw.lean index ca58315eab39..fa46e63ddb1e 100644 --- a/src/Std/Data/HashMap/Raw.lean +++ b/src/Std/Data/HashMap/Raw.lean @@ -229,8 +229,13 @@ instance {m : Type w → Type w'} : ForIn m (Raw α β) (α × β) where @[inline] def union [BEq α] [Hashable α] (m₁ m₂ : Raw α β) : Raw α β := ⟨DHashMap.Raw.union m₁.inner m₂.inner⟩ +@[inherit_doc DHashMap.Raw.inter, inline] def inter [BEq α] [Hashable α] (m₁ m₂ : Raw α β) : Raw α β := + ⟨DHashMap.Raw.inter m₁.inner m₂.inner⟩ + instance [BEq α] [Hashable α] : Union (Raw α β) := ⟨union⟩ +instance [BEq α] [Hashable α] : Inter (Raw α β) := ⟨inter⟩ + section Unverified @[inline, inherit_doc DHashMap.Raw.filterMap] def filterMap {γ : Type w} (f : α → β → Option γ) diff --git a/src/Std/Data/HashMap/RawLemmas.lean b/src/Std/Data/HashMap/RawLemmas.lean index c4f505822a8f..55b9634fe192 100644 --- a/src/Std/Data/HashMap/RawLemmas.lean +++ b/src/Std/Data/HashMap/RawLemmas.lean @@ -1428,6 +1428,42 @@ theorem isEmpty_union [EquivBEq α] [LawfulHashable α] (h₁ : m₁.WF) (h₂ : end Union +section Inter + +variable {m₁ m₂ : Raw α β} + +@[simp] +theorem inter_eq : m₁.inter m₂ = m₁ ∩ m₂ := by + simp only [Inter.inter] + +/- contains -/ + +@[simp] +theorem contains_inter [EquivBEq α] [LawfulHashable α] {k : α} (h₁ : m₁.WF) (h₂ : m₂.WF) : + (m₁ ∩ m₂).contains k = (m₁.contains k && m₂.contains k) := + @DHashMap.Raw.contains_inter α (fun _ => β) _ _ m₁.inner m₂.inner _ _ k h₁.out h₂.out + +/- mem -/ + +theorem mem_inter_iff [EquivBEq α] [LawfulHashable α] (h₁ : m₁.WF) + (h₂ : m₂.WF) {k : α} : + k ∈ m₁ ∩ m₂ ↔ k ∈ m₁ ∧ k ∈ m₂ := + @DHashMap.Raw.mem_inter_iff α (fun _ => β) _ _ m₁.inner m₂.inner _ _ h₁.out h₂.out k + +theorem not_mem_inter_of_not_mem_left [EquivBEq α] [LawfulHashable α] + (h₁ : m₁.WF) (h₂ : m₂.WF) {k : α} + (not_mem : k ∉ m₁) : + k ∉ m₁ ∩ m₂ := + @DHashMap.Raw.not_mem_inter_of_not_mem_left α (fun _ => β) _ _ m₁.inner m₂.inner _ _ h₁.out h₂.out k not_mem + +theorem not_mem_inter_of_not_mem_right [EquivBEq α] [LawfulHashable α] + (h₁ : m₁.WF) (h₂ : m₂.WF) {k : α} + (not_mem : k ∉ m₂) : + k ∉ m₁ ∩ m₂ := + @DHashMap.Raw.not_mem_inter_of_not_mem_right α (fun _ => β) _ _ m₁.inner m₂.inner _ _ h₁.out h₂.out k not_mem + +end Inter + 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) : From 3bf6d395ee8d217efd4dee02381d0096ff643998 Mon Sep 17 00:00:00 2001 From: Wojciech Rozowski Date: Thu, 13 Nov 2025 10:28:15 +0100 Subject: [PATCH 43/88] reorganise things in the files --- src/Std/Data/DHashMap/Lemmas.lean | 154 ++++++++++---------- src/Std/Data/DHashMap/RawLemmas.lean | 204 +++++++++++++-------------- 2 files changed, 178 insertions(+), 180 deletions(-) diff --git a/src/Std/Data/DHashMap/Lemmas.lean b/src/Std/Data/DHashMap/Lemmas.lean index 435b75cc704f..8a762c79fccd 100644 --- a/src/Std/Data/DHashMap/Lemmas.lean +++ b/src/Std/Data/DHashMap/Lemmas.lean @@ -1932,11 +1932,84 @@ theorem isEmpty_union [EquivBEq α] [LawfulHashable α] : end Union -section Inter +namespace Const -variable (m₁ m₂ : DHashMap α β) +variable {β : Type v} {m₁ m₂ : DHashMap α (fun _ => β)} -variable {m₁ m₂} +/- get? -/ +theorem get?_union [EquivBEq α] [LawfulHashable α] {k : α} : + Const.get? (m₁.union m₂) k = (Const.get? m₂ k).or (Const.get? m₁ k) := + @Raw₀.Const.get?_union _ _ _ _ ⟨m₁.1, _⟩ ⟨m₂.1, _⟩ _ _ m₁.2 m₂.2 k + +theorem get?_union_of_not_mem_left [EquivBEq α] [LawfulHashable α] + {k : α} (not_mem : ¬k ∈ m₁) : + Const.get? (m₁.union m₂) k = Const.get? m₂ k := by + rw [← contains_eq_false_iff_not_mem] at not_mem + exact @Raw₀.Const.get?_union_of_contains_eq_false_left _ _ _ _ ⟨m₁.1, _⟩ ⟨m₂.1, _⟩ _ _ m₁.2 m₂.2 k not_mem + +theorem get?_union_of_not_mem_right [EquivBEq α] [LawfulHashable α] + {k : α} (not_mem : ¬k ∈ m₂) : + Const.get? (m₁.union m₂) k = Const.get? m₁ k := by + rw [← contains_eq_false_iff_not_mem] at not_mem + exact @Raw₀.Const.get?_union_of_contains_eq_false_right _ _ _ _ ⟨m₁.1, _⟩ ⟨m₂.1, _⟩ _ _ m₁.2 m₂.2 k not_mem + +/- get -/ +theorem get_union_of_mem_right [EquivBEq α] [LawfulHashable α] + {k : α} (mem : m₂.contains k) : + Const.get (m₁.union m₂) k (mem_union_of_right mem) = Const.get m₂ k mem := + @Raw₀.Const.get_union_of_contains_right _ _ _ _ ⟨m₁.1, _⟩ ⟨m₂.1, _⟩ _ _ m₁.2 m₂.2 k mem + +theorem get_union_of_not_mem_left [EquivBEq α] [LawfulHashable α] + {k : α} (not_mem : ¬k ∈ m₁) {h'} : + Const.get (m₁.union m₂) k h' = Const.get m₂ k (mem_of_mem_union_of_not_mem_left h' not_mem) := by + rw [← contains_eq_false_iff_not_mem] at not_mem + exact @Raw₀.Const.get_union_of_contains_eq_false_left _ _ _ _ ⟨m₁.1, _⟩ ⟨m₂.1, _⟩ _ _ m₁.2 m₂.2 k not_mem h' + +theorem get_union_of_not_mem_right [EquivBEq α] [LawfulHashable α] + {k : α} (not_mem : ¬k ∈ m₂) {h'} : + Const.get (m₁.union m₂) k h' = Const.get m₁ k (mem_of_mem_union_of_not_mem_right h' not_mem) := by + rw [← contains_eq_false_iff_not_mem] at not_mem + exact @Raw₀.Const.get_union_of_contains_eq_false_right _ _ _ _ ⟨m₁.1, _⟩ ⟨m₂.1, _⟩ _ _ m₁.2 m₂.2 k not_mem h' + +/- getD -/ +theorem getD_union [EquivBEq α] [LawfulHashable α] {k : α} {fallback : β} : + Const.getD (m₁.union m₂) k fallback = Const.getD m₂ k (Const.getD m₁ k fallback) := + @Raw₀.Const.getD_union _ _ _ _ ⟨m₁.1, m₁.2.size_buckets_pos⟩ ⟨m₂.1, m₂.2.size_buckets_pos⟩ _ _ m₁.2 m₂.2 k fallback + +theorem getD_union_of_not_mem_left [EquivBEq α] [LawfulHashable α] + {k : α} {fallback : β} (not_mem : ¬k ∈ m₁) : + Const.getD (m₁.union m₂) k fallback = Const.getD m₂ k fallback := by + rw [← contains_eq_false_iff_not_mem] at not_mem + exact @Raw₀.Const.getD_union_of_contains_eq_false_left _ _ _ _ ⟨m₁.1, m₁.2.size_buckets_pos⟩ ⟨m₂.1, m₂.2.size_buckets_pos⟩ _ _ m₁.2 m₂.2 k fallback not_mem + +theorem getD_union_of_not_mem_right [EquivBEq α] [LawfulHashable α] + {k : α} {fallback : β} (not_mem : ¬k ∈ m₂) : + Const.getD (m₁.union m₂) k fallback = Const.getD m₁ k fallback := by + rw [← contains_eq_false_iff_not_mem] at not_mem + exact @Raw₀.Const.getD_union_of_contains_eq_false_right _ _ _ _ ⟨m₁.1, m₁.2.size_buckets_pos⟩ ⟨m₂.1, m₂.2.size_buckets_pos⟩ _ _ m₁.2 m₂.2 k fallback not_mem + +/- get! -/ +theorem get!_union [EquivBEq α] [LawfulHashable α] [Inhabited β] {k : α} : + Const.get! (m₁.union m₂) k = Const.getD m₂ k (Const.get! m₁ k) := + @Raw₀.Const.get!_union _ _ _ _ ⟨m₁.1, m₁.2.size_buckets_pos⟩ ⟨m₂.1, m₂.2.size_buckets_pos⟩ _ _ _ m₁.2 m₂.2 k + +theorem get!_union_of_not_mem_left [EquivBEq α] [LawfulHashable α] [Inhabited β] + {k : α} (not_mem : ¬k ∈ m₁) : + Const.get! (m₁.union m₂) k = Const.get! m₂ k := by + rw [← contains_eq_false_iff_not_mem] at not_mem + exact @Raw₀.Const.get!_union_of_contains_eq_false_left _ _ _ _ ⟨m₁.1, m₁.2.size_buckets_pos⟩ ⟨m₂.1, m₂.2.size_buckets_pos⟩ _ _ _ m₁.2 m₂.2 k not_mem + +theorem get!_union_of_not_mem_right [EquivBEq α] [LawfulHashable α] [Inhabited β] + {k : α} (not_mem : ¬k ∈ m₂) : + Const.get! (m₁.union m₂) k = Const.get! m₁ k := by + rw [← contains_eq_false_iff_not_mem] at not_mem + exact @Raw₀.Const.get!_union_of_contains_eq_false_right _ _ _ _ ⟨m₁.1, m₁.2.size_buckets_pos⟩ ⟨m₂.1, m₂.2.size_buckets_pos⟩ _ _ _ m₁.2 m₂.2 k not_mem + +end Const + +section Inter + +variable {m₁ m₂ : DHashMap α β} @[simp] theorem inter_eq : m₁.inter m₂ = m₁ ∩ m₂ := by @@ -2175,81 +2248,6 @@ end Inter namespace Const -variable {β : Type v} {m₁ m₂ : DHashMap α (fun _ => β)} - -/- get? -/ -theorem get?_union [EquivBEq α] [LawfulHashable α] {k : α} : - Const.get? (m₁.union m₂) k = (Const.get? m₂ k).or (Const.get? m₁ k) := - @Raw₀.Const.get?_union _ _ _ _ ⟨m₁.1, _⟩ ⟨m₂.1, _⟩ _ _ m₁.2 m₂.2 k - -theorem get?_union_of_not_mem_left [EquivBEq α] [LawfulHashable α] - {k : α} (not_mem : ¬k ∈ m₁) : - Const.get? (m₁.union m₂) k = Const.get? m₂ k := by - rw [← contains_eq_false_iff_not_mem] at not_mem - exact @Raw₀.Const.get?_union_of_contains_eq_false_left _ _ _ _ ⟨m₁.1, _⟩ ⟨m₂.1, _⟩ _ _ m₁.2 m₂.2 k not_mem - -theorem get?_union_of_not_mem_right [EquivBEq α] [LawfulHashable α] - {k : α} (not_mem : ¬k ∈ m₂) : - Const.get? (m₁.union m₂) k = Const.get? m₁ k := by - rw [← contains_eq_false_iff_not_mem] at not_mem - exact @Raw₀.Const.get?_union_of_contains_eq_false_right _ _ _ _ ⟨m₁.1, _⟩ ⟨m₂.1, _⟩ _ _ m₁.2 m₂.2 k not_mem - -/- get -/ -theorem get_union_of_mem_right [EquivBEq α] [LawfulHashable α] - {k : α} (mem : m₂.contains k) : - Const.get (m₁.union m₂) k (mem_union_of_right mem) = Const.get m₂ k mem := - @Raw₀.Const.get_union_of_contains_right _ _ _ _ ⟨m₁.1, _⟩ ⟨m₂.1, _⟩ _ _ m₁.2 m₂.2 k mem - -theorem get_union_of_not_mem_left [EquivBEq α] [LawfulHashable α] - {k : α} (not_mem : ¬k ∈ m₁) {h'} : - Const.get (m₁.union m₂) k h' = Const.get m₂ k (mem_of_mem_union_of_not_mem_left h' not_mem) := by - rw [← contains_eq_false_iff_not_mem] at not_mem - exact @Raw₀.Const.get_union_of_contains_eq_false_left _ _ _ _ ⟨m₁.1, _⟩ ⟨m₂.1, _⟩ _ _ m₁.2 m₂.2 k not_mem h' - -theorem get_union_of_not_mem_right [EquivBEq α] [LawfulHashable α] - {k : α} (not_mem : ¬k ∈ m₂) {h'} : - Const.get (m₁.union m₂) k h' = Const.get m₁ k (mem_of_mem_union_of_not_mem_right h' not_mem) := by - rw [← contains_eq_false_iff_not_mem] at not_mem - exact @Raw₀.Const.get_union_of_contains_eq_false_right _ _ _ _ ⟨m₁.1, _⟩ ⟨m₂.1, _⟩ _ _ m₁.2 m₂.2 k not_mem h' - -/- getD -/ -theorem getD_union [EquivBEq α] [LawfulHashable α] {k : α} {fallback : β} : - Const.getD (m₁.union m₂) k fallback = Const.getD m₂ k (Const.getD m₁ k fallback) := - @Raw₀.Const.getD_union _ _ _ _ ⟨m₁.1, m₁.2.size_buckets_pos⟩ ⟨m₂.1, m₂.2.size_buckets_pos⟩ _ _ m₁.2 m₂.2 k fallback - -theorem getD_union_of_not_mem_left [EquivBEq α] [LawfulHashable α] - {k : α} {fallback : β} (not_mem : ¬k ∈ m₁) : - Const.getD (m₁.union m₂) k fallback = Const.getD m₂ k fallback := by - rw [← contains_eq_false_iff_not_mem] at not_mem - exact @Raw₀.Const.getD_union_of_contains_eq_false_left _ _ _ _ ⟨m₁.1, m₁.2.size_buckets_pos⟩ ⟨m₂.1, m₂.2.size_buckets_pos⟩ _ _ m₁.2 m₂.2 k fallback not_mem - -theorem getD_union_of_not_mem_right [EquivBEq α] [LawfulHashable α] - {k : α} {fallback : β} (not_mem : ¬k ∈ m₂) : - Const.getD (m₁.union m₂) k fallback = Const.getD m₁ k fallback := by - rw [← contains_eq_false_iff_not_mem] at not_mem - exact @Raw₀.Const.getD_union_of_contains_eq_false_right _ _ _ _ ⟨m₁.1, m₁.2.size_buckets_pos⟩ ⟨m₂.1, m₂.2.size_buckets_pos⟩ _ _ m₁.2 m₂.2 k fallback not_mem - -/- get! -/ -theorem get!_union [EquivBEq α] [LawfulHashable α] [Inhabited β] {k : α} : - Const.get! (m₁.union m₂) k = Const.getD m₂ k (Const.get! m₁ k) := - @Raw₀.Const.get!_union _ _ _ _ ⟨m₁.1, m₁.2.size_buckets_pos⟩ ⟨m₂.1, m₂.2.size_buckets_pos⟩ _ _ _ m₁.2 m₂.2 k - -theorem get!_union_of_not_mem_left [EquivBEq α] [LawfulHashable α] [Inhabited β] - {k : α} (not_mem : ¬k ∈ m₁) : - Const.get! (m₁.union m₂) k = Const.get! m₂ k := by - rw [← contains_eq_false_iff_not_mem] at not_mem - exact @Raw₀.Const.get!_union_of_contains_eq_false_left _ _ _ _ ⟨m₁.1, m₁.2.size_buckets_pos⟩ ⟨m₂.1, m₂.2.size_buckets_pos⟩ _ _ _ m₁.2 m₂.2 k not_mem - -theorem get!_union_of_not_mem_right [EquivBEq α] [LawfulHashable α] [Inhabited β] - {k : α} (not_mem : ¬k ∈ m₂) : - Const.get! (m₁.union m₂) k = Const.get! m₁ k := by - rw [← contains_eq_false_iff_not_mem] at not_mem - exact @Raw₀.Const.get!_union_of_contains_eq_false_right _ _ _ _ ⟨m₁.1, m₁.2.size_buckets_pos⟩ ⟨m₂.1, m₂.2.size_buckets_pos⟩ _ _ _ m₁.2 m₂.2 k not_mem - -end Const - -namespace Const - variable {β : Type v} {m : DHashMap α (fun _ => β)} variable {ρ : Type w} [ForIn Id ρ (α × β)] diff --git a/src/Std/Data/DHashMap/RawLemmas.lean b/src/Std/Data/DHashMap/RawLemmas.lean index 55e83ca37fdd..c0204be0d2e8 100644 --- a/src/Std/Data/DHashMap/RawLemmas.lean +++ b/src/Std/Data/DHashMap/RawLemmas.lean @@ -2129,6 +2129,108 @@ theorem isEmpty_union [EquivBEq α] [LawfulHashable α] (h₁ : m₁.WF) (h₂ : end Union +namespace Const + +variable {β : Type v} {m₁ m₂ : Raw α (fun _ => β)} + +theorem get?_union [EquivBEq α] [LawfulHashable α] (h₁ : m₁.WF) (h₂ : m₂.WF) {k : α} : + Const.get? (m₁ ∪ m₂) k = (Const.get? m₂ k).or (Const.get? m₁ k) := by + simp only [Union.union] + simp_to_raw using Raw₀.Const.get?_union + +theorem get?_union_of_not_mem_left [EquivBEq α] [LawfulHashable α] (h₁ : m₁.WF) (h₂ : m₂.WF) + {k : α} (not_mem : ¬k ∈ m₁) : + Const.get? (m₁ ∪ m₂) k = Const.get? m₂ k := by + rw [← contains_eq_false_iff_not_mem] at not_mem + revert not_mem + simp only [Union.union] + simp_to_raw using Raw₀.Const.get?_union_of_contains_eq_false_left + +theorem get?_union_of_not_mem_right [EquivBEq α] [LawfulHashable α] (h₁ : m₁.WF) (h₂ : m₂.WF) + {k : α} (not_mem : ¬k ∈ m₂) : + Const.get? (m₁ ∪ m₂) k = Const.get? m₁ k := by + rw [← contains_eq_false_iff_not_mem] at not_mem + revert not_mem + simp only [Union.union] + simp_to_raw using Raw₀.Const.get?_union_of_contains_eq_false_right + +/- get -/ +theorem get_union_of_mem_right [EquivBEq α] [LawfulHashable α] (h₁ : m₁.WF) (h₂ : m₂.WF) + {k : α} (mem : k ∈ m₂) : + Const.get (m₁ ∪ m₂) k (mem_union_of_right h₁ h₂ mem) = Const.get m₂ k mem := by + rw [mem_iff_contains] at mem + revert mem + simp only [Union.union] + simp_to_raw using Raw₀.Const.get_union_of_contains_right + +theorem get_union_of_not_mem_left [EquivBEq α] [LawfulHashable α] (h₁ : m₁.WF) (h₂ : m₂.WF) + {k : α} (not_mem : ¬k ∈ m₁) {h'} : + Const.get (m₁ ∪ m₂) k h' = Const.get m₂ k (mem_of_mem_union_of_not_mem_left h₁ h₂ h' not_mem) := by + rw [← contains_eq_false_iff_not_mem] at not_mem + revert not_mem + simp only [Union.union] + simp_to_raw + intro not_mem + apply Raw₀.Const.get_union_of_contains_eq_false_left + all_goals wf_trivial + +theorem get_union_of_not_mem_right [EquivBEq α] [LawfulHashable α] (h₁ : m₁.WF) (h₂ : m₂.WF) + {k : α} (not_mem : ¬k ∈ m₂) {h'} : + Const.get (m₁ ∪ m₂) k h' = Const.get m₁ k (mem_of_mem_union_of_not_mem_right h₁ h₂ h' not_mem) := by + rw [← contains_eq_false_iff_not_mem] at not_mem + revert not_mem + simp only [Union.union] + simp_to_raw + intro not_mem + apply Raw₀.Const.get_union_of_contains_eq_false_right + all_goals wf_trivial + +/- getD -/ +theorem getD_union [EquivBEq α] [LawfulHashable α] (h₁ : m₁.WF) (h₂ : m₂.WF) {k : α} {fallback : β} : + Const.getD (m₁ ∪ m₂) k fallback = Const.getD m₂ k (Const.getD m₁ k fallback) := by + simp only [Union.union] + simp_to_raw using Raw₀.Const.getD_union + +theorem getD_union_of_not_mem_left [EquivBEq α] [LawfulHashable α] (h₁ : m₁.WF) (h₂ : m₂.WF) + {k : α} {fallback : β} (not_mem : ¬k ∈ m₁) : + Const.getD (m₁ ∪ m₂) k fallback = Const.getD m₂ k fallback := by + rw [← contains_eq_false_iff_not_mem] at not_mem + revert not_mem + simp only [Union.union] + simp_to_raw using Raw₀.Const.getD_union_of_contains_eq_false_left + +theorem getD_union_of_not_mem_right [EquivBEq α] [LawfulHashable α] (h₁ : m₁.WF) (h₂ : m₂.WF) + {k : α} {fallback : β} (not_mem : ¬k ∈ m₂) : + Const.getD (m₁ ∪ m₂) k fallback = Const.getD m₁ k fallback := by + rw [← contains_eq_false_iff_not_mem] at not_mem + revert not_mem + simp only [Union.union] + simp_to_raw using Raw₀.Const.getD_union_of_contains_eq_false_right + +/- get! -/ +theorem get!_union [EquivBEq α] [LawfulHashable α] [Inhabited β] (h₁ : m₁.WF) (h₂ : m₂.WF) {k : α} : + Const.get! (m₁ ∪ m₂) k = Const.getD m₂ k (Const.get! m₁ k) := by + simp only [Union.union] + simp_to_raw using Raw₀.Const.get!_union + +theorem get!_union_of_not_mem_left [EquivBEq α] [LawfulHashable α] [Inhabited β] (h₁ : m₁.WF) (h₂ : m₂.WF) + {k : α} (not_mem : ¬k ∈ m₁) : + Const.get! (m₁ ∪ m₂) k = Const.get! m₂ k := by + rw [← contains_eq_false_iff_not_mem] at not_mem + revert not_mem + simp only [Union.union] + simp_to_raw using Raw₀.Const.get!_union_of_contains_eq_false_left + +theorem get!_union_of_not_mem_right [EquivBEq α] [LawfulHashable α] [Inhabited β] (h₁ : m₁.WF) (h₂ : m₂.WF) + {k : α} (not_mem : ¬k ∈ m₂) : + Const.get! (m₁ ∪ m₂) k = Const.get! m₁ k := by + rw [← contains_eq_false_iff_not_mem ] at not_mem + revert not_mem + simp only [Union.union] + simp_to_raw using Raw₀.Const.get!_union_of_contains_eq_false_right + +end Const + section Inter variable {m₁ m₂ : Raw α β} @@ -2456,108 +2558,6 @@ end Inter namespace Const -variable {β : Type v} {m₁ m₂ : Raw α (fun _ => β)} - -theorem get?_union [EquivBEq α] [LawfulHashable α] (h₁ : m₁.WF) (h₂ : m₂.WF) {k : α} : - Const.get? (m₁ ∪ m₂) k = (Const.get? m₂ k).or (Const.get? m₁ k) := by - simp only [Union.union] - simp_to_raw using Raw₀.Const.get?_union - -theorem get?_union_of_not_mem_left [EquivBEq α] [LawfulHashable α] (h₁ : m₁.WF) (h₂ : m₂.WF) - {k : α} (not_mem : ¬k ∈ m₁) : - Const.get? (m₁ ∪ m₂) k = Const.get? m₂ k := by - rw [← contains_eq_false_iff_not_mem] at not_mem - revert not_mem - simp only [Union.union] - simp_to_raw using Raw₀.Const.get?_union_of_contains_eq_false_left - -theorem get?_union_of_not_mem_right [EquivBEq α] [LawfulHashable α] (h₁ : m₁.WF) (h₂ : m₂.WF) - {k : α} (not_mem : ¬k ∈ m₂) : - Const.get? (m₁ ∪ m₂) k = Const.get? m₁ k := by - rw [← contains_eq_false_iff_not_mem] at not_mem - revert not_mem - simp only [Union.union] - simp_to_raw using Raw₀.Const.get?_union_of_contains_eq_false_right - -/- get -/ -theorem get_union_of_mem_right [EquivBEq α] [LawfulHashable α] (h₁ : m₁.WF) (h₂ : m₂.WF) - {k : α} (mem : k ∈ m₂) : - Const.get (m₁ ∪ m₂) k (mem_union_of_right h₁ h₂ mem) = Const.get m₂ k mem := by - rw [mem_iff_contains] at mem - revert mem - simp only [Union.union] - simp_to_raw using Raw₀.Const.get_union_of_contains_right - -theorem get_union_of_not_mem_left [EquivBEq α] [LawfulHashable α] (h₁ : m₁.WF) (h₂ : m₂.WF) - {k : α} (not_mem : ¬k ∈ m₁) {h'} : - Const.get (m₁ ∪ m₂) k h' = Const.get m₂ k (mem_of_mem_union_of_not_mem_left h₁ h₂ h' not_mem) := by - rw [← contains_eq_false_iff_not_mem] at not_mem - revert not_mem - simp only [Union.union] - simp_to_raw - intro not_mem - apply Raw₀.Const.get_union_of_contains_eq_false_left - all_goals wf_trivial - -theorem get_union_of_not_mem_right [EquivBEq α] [LawfulHashable α] (h₁ : m₁.WF) (h₂ : m₂.WF) - {k : α} (not_mem : ¬k ∈ m₂) {h'} : - Const.get (m₁ ∪ m₂) k h' = Const.get m₁ k (mem_of_mem_union_of_not_mem_right h₁ h₂ h' not_mem) := by - rw [← contains_eq_false_iff_not_mem] at not_mem - revert not_mem - simp only [Union.union] - simp_to_raw - intro not_mem - apply Raw₀.Const.get_union_of_contains_eq_false_right - all_goals wf_trivial - -/- getD -/ -theorem getD_union [EquivBEq α] [LawfulHashable α] (h₁ : m₁.WF) (h₂ : m₂.WF) {k : α} {fallback : β} : - Const.getD (m₁ ∪ m₂) k fallback = Const.getD m₂ k (Const.getD m₁ k fallback) := by - simp only [Union.union] - simp_to_raw using Raw₀.Const.getD_union - -theorem getD_union_of_not_mem_left [EquivBEq α] [LawfulHashable α] (h₁ : m₁.WF) (h₂ : m₂.WF) - {k : α} {fallback : β} (not_mem : ¬k ∈ m₁) : - Const.getD (m₁ ∪ m₂) k fallback = Const.getD m₂ k fallback := by - rw [← contains_eq_false_iff_not_mem] at not_mem - revert not_mem - simp only [Union.union] - simp_to_raw using Raw₀.Const.getD_union_of_contains_eq_false_left - -theorem getD_union_of_not_mem_right [EquivBEq α] [LawfulHashable α] (h₁ : m₁.WF) (h₂ : m₂.WF) - {k : α} {fallback : β} (not_mem : ¬k ∈ m₂) : - Const.getD (m₁ ∪ m₂) k fallback = Const.getD m₁ k fallback := by - rw [← contains_eq_false_iff_not_mem] at not_mem - revert not_mem - simp only [Union.union] - simp_to_raw using Raw₀.Const.getD_union_of_contains_eq_false_right - -/- get! -/ -theorem get!_union [EquivBEq α] [LawfulHashable α] [Inhabited β] (h₁ : m₁.WF) (h₂ : m₂.WF) {k : α} : - Const.get! (m₁ ∪ m₂) k = Const.getD m₂ k (Const.get! m₁ k) := by - simp only [Union.union] - simp_to_raw using Raw₀.Const.get!_union - -theorem get!_union_of_not_mem_left [EquivBEq α] [LawfulHashable α] [Inhabited β] (h₁ : m₁.WF) (h₂ : m₂.WF) - {k : α} (not_mem : ¬k ∈ m₁) : - Const.get! (m₁ ∪ m₂) k = Const.get! m₂ k := by - rw [← contains_eq_false_iff_not_mem] at not_mem - revert not_mem - simp only [Union.union] - simp_to_raw using Raw₀.Const.get!_union_of_contains_eq_false_left - -theorem get!_union_of_not_mem_right [EquivBEq α] [LawfulHashable α] [Inhabited β] (h₁ : m₁.WF) (h₂ : m₂.WF) - {k : α} (not_mem : ¬k ∈ m₂) : - Const.get! (m₁ ∪ m₂) k = Const.get! m₁ k := by - rw [← contains_eq_false_iff_not_mem ] at not_mem - revert not_mem - simp only [Union.union] - simp_to_raw using Raw₀.Const.get!_union_of_contains_eq_false_right - -end Const - -namespace Const - variable {β : Type v} {m : Raw α (fun _ => β)} variable {ρ : Type w} [ForIn Id ρ (α × β)] From a6fcc761592d5a977fe218191f201e094f3825f6 Mon Sep 17 00:00:00 2001 From: Wojciech Rozowski Date: Thu, 13 Nov 2025 10:53:15 +0100 Subject: [PATCH 44/88] Add const lemmas --- src/Std/Data/DHashMap/Lemmas.lean | 83 ++++++++++++++++++++ src/Std/Data/DHashMap/RawLemmas.lean | 113 +++++++++++++++++++++++++++ 2 files changed, 196 insertions(+) diff --git a/src/Std/Data/DHashMap/Lemmas.lean b/src/Std/Data/DHashMap/Lemmas.lean index 8a762c79fccd..e738c0df2b0a 100644 --- a/src/Std/Data/DHashMap/Lemmas.lean +++ b/src/Std/Data/DHashMap/Lemmas.lean @@ -2248,6 +2248,89 @@ end Inter namespace Const +variable {β : Type v} {m₁ m₂ : DHashMap α (fun _ => β)} + +/- get? -/ +theorem get?_inter [EquivBEq α] [LawfulHashable α] {k : α} : + Const.get? (m₁.inter m₂) k = + if k ∈ m₂ then Const.get? m₁ k else none := + @Raw₀.Const.get?_inter _ _ _ _ ⟨m₁.1, m₁.2.size_buckets_pos⟩ ⟨m₂.1, m₂.2.size_buckets_pos⟩ _ _ m₁.2 m₂.2 k + +theorem get?_inter_of_mem_right [EquivBEq α] [LawfulHashable α] + {k : α} (mem : k ∈ m₂) : + Const.get? (m₁.inter m₂) k = Const.get? m₁ k := + @Raw₀.Const.get?_inter_of_contains_right _ _ _ _ ⟨m₁.1, m₁.2.size_buckets_pos⟩ ⟨m₂.1, m₂.2.size_buckets_pos⟩ _ _ m₁.2 m₂.2 k mem + +theorem get?_inter_of_not_mem_left [EquivBEq α] [LawfulHashable α] + {k : α} (not_mem : ¬k ∈ m₁) : + Const.get? (m₁.inter m₂) k = none := by + rw [← contains_eq_false_iff_not_mem] at not_mem + exact @Raw₀.Const.get?_inter_of_contains_eq_false_left _ _ _ _ ⟨m₁.1, m₁.2.size_buckets_pos⟩ ⟨m₂.1, m₂.2.size_buckets_pos⟩ _ _ m₁.2 m₂.2 k not_mem + +theorem get?_inter_of_not_mem_right [EquivBEq α] [LawfulHashable α] + {k : α} (not_mem : ¬k ∈ m₂) : + Const.get? (m₁.inter m₂) k = none := by + rw [← contains_eq_false_iff_not_mem] at not_mem + exact @Raw₀.Const.get?_inter_of_contains_eq_false_right _ _ _ _ ⟨m₁.1, m₁.2.size_buckets_pos⟩ ⟨m₂.1, m₂.2.size_buckets_pos⟩ _ _ m₁.2 m₂.2 k not_mem + +/- get -/ +theorem get_inter [EquivBEq α] [LawfulHashable α] + {k : α} {h_mem : k ∈ m₁ ∩ m₂} : + Const.get (m₁.inter m₂) k h_mem = + Const.get m₁ k ((mem_inter_iff.1 h_mem).1) := by + rw [mem_iff_contains] at h_mem + exact @Raw₀.Const.get_inter _ _ _ _ ⟨m₁.1, _⟩ ⟨m₂.1, _⟩ _ _ m₁.2 m₂.2 k h_mem + +/- getD -/ +theorem getD_inter [EquivBEq α] [LawfulHashable α] {k : α} {fallback : β} : + Const.getD (m₁.inter m₂) k fallback = + if k ∈ m₂ then Const.getD m₁ k fallback else fallback := + @Raw₀.Const.getD_inter _ _ _ _ ⟨m₁.1, m₁.2.size_buckets_pos⟩ ⟨m₂.1, m₂.2.size_buckets_pos⟩ _ _ m₁.2 m₂.2 k fallback + +theorem getD_inter_of_mem_right [EquivBEq α] [LawfulHashable α] + {k : α} {fallback : β} (mem : k ∈ m₂) : + Const.getD (m₁.inter m₂) k fallback = Const.getD m₁ k fallback := + @Raw₀.Const.getD_inter_of_contains_right _ _ _ _ ⟨m₁.1, m₁.2.size_buckets_pos⟩ ⟨m₂.1, m₂.2.size_buckets_pos⟩ _ _ m₁.2 m₂.2 k fallback mem + +theorem getD_inter_of_not_mem_right [EquivBEq α] [LawfulHashable α] + {k : α} {fallback : β} (not_mem : ¬k ∈ m₂) : + Const.getD (m₁.inter m₂) k fallback = fallback := by + rw [← contains_eq_false_iff_not_mem] at not_mem + exact @Raw₀.Const.getD_inter_of_contains_eq_false_right _ _ _ _ ⟨m₁.1, m₁.2.size_buckets_pos⟩ ⟨m₂.1, m₂.2.size_buckets_pos⟩ _ _ m₁.2 m₂.2 k fallback not_mem + +theorem getD_inter_of_not_mem_left [EquivBEq α] [LawfulHashable α] + {k : α} {fallback : β} (not_mem : ¬k ∈ m₁) : + Const.getD (m₁.inter m₂) k fallback = fallback := by + rw [← contains_eq_false_iff_not_mem] at not_mem + exact @Raw₀.Const.getD_inter_of_contains_eq_false_left _ _ _ _ ⟨m₁.1, m₁.2.size_buckets_pos⟩ ⟨m₂.1, m₂.2.size_buckets_pos⟩ _ _ m₁.2 m₂.2 k fallback not_mem + +/- get! -/ +theorem get!_inter [EquivBEq α] [LawfulHashable α] [Inhabited β] {k : α} : + Const.get! (m₁.inter m₂) k = + if k ∈ m₂ then Const.get! m₁ k else default := + @Raw₀.Const.get!_inter _ _ _ _ ⟨m₁.1, m₁.2.size_buckets_pos⟩ ⟨m₂.1, m₂.2.size_buckets_pos⟩ _ _ _ m₁.2 m₂.2 k + +theorem get!_inter_of_mem_right [EquivBEq α] [LawfulHashable α] [Inhabited β] + {k : α} (mem : k ∈ m₂) : + Const.get! (m₁.inter m₂) k = Const.get! m₁ k := + @Raw₀.Const.get!_inter_of_contains_right _ _ _ _ ⟨m₁.1, m₁.2.size_buckets_pos⟩ ⟨m₂.1, m₂.2.size_buckets_pos⟩ _ _ _ m₁.2 m₂.2 k mem + +theorem get!_inter_of_not_mem_right [EquivBEq α] [LawfulHashable α] [Inhabited β] + {k : α} (not_mem : ¬k ∈ m₂) : + Const.get! (m₁.inter m₂) k = default := by + rw [← contains_eq_false_iff_not_mem] at not_mem + exact @Raw₀.Const.get!_inter_of_contains_eq_false_right _ _ _ _ ⟨m₁.1, m₁.2.size_buckets_pos⟩ ⟨m₂.1, m₂.2.size_buckets_pos⟩ _ _ _ m₁.2 m₂.2 k not_mem + +theorem get!_inter_of_not_mem_left [EquivBEq α] [LawfulHashable α] [Inhabited β] + {k : α} (not_mem : ¬k ∈ m₁) : + Const.get! (m₁.inter m₂) k = default := by + rw [← contains_eq_false_iff_not_mem] at not_mem + exact @Raw₀.Const.get!_inter_of_contains_eq_false_left _ _ _ _ ⟨m₁.1, m₁.2.size_buckets_pos⟩ ⟨m₂.1, m₂.2.size_buckets_pos⟩ _ _ _ m₁.2 m₂.2 k not_mem + +end Const + +namespace Const + variable {β : Type v} {m : DHashMap α (fun _ => β)} variable {ρ : Type w} [ForIn Id ρ (α × β)] diff --git a/src/Std/Data/DHashMap/RawLemmas.lean b/src/Std/Data/DHashMap/RawLemmas.lean index c0204be0d2e8..f91977b5d9fb 100644 --- a/src/Std/Data/DHashMap/RawLemmas.lean +++ b/src/Std/Data/DHashMap/RawLemmas.lean @@ -2558,6 +2558,119 @@ end Inter namespace Const +variable {β : Type v} {m₁ m₂ : Raw α (fun _ => β)} + +/- get? -/ +theorem get?_inter [EquivBEq α] [LawfulHashable α] (h₁ : m₁.WF) (h₂ : m₂.WF) {k : α} : + Const.get? (m₁ ∩ m₂) k = + if k ∈ m₂ then Const.get? m₁ k else none := by + simp only [Membership.mem, Inter.inter] + simp_to_raw using Raw₀.Const.get?_inter + +theorem get?_inter_of_mem_right [EquivBEq α] [LawfulHashable α] (h₁ : m₁.WF) (h₂ : m₂.WF) + {k : α} (mem : k ∈ m₂) : + Const.get? (m₁ ∩ m₂) k = Const.get? m₁ k := by + rw [mem_iff_contains] at mem + revert mem + simp only [Inter.inter] + simp_to_raw using Raw₀.Const.get?_inter_of_contains_right + +theorem get?_inter_of_not_mem_left [EquivBEq α] [LawfulHashable α] (h₁ : m₁.WF) (h₂ : m₂.WF) + {k : α} (not_mem : ¬k ∈ m₁) : + Const.get? (m₁ ∩ m₂) k = none := by + rw [← contains_eq_false_iff_not_mem] at not_mem + revert not_mem + simp only [Inter.inter] + simp_to_raw using Raw₀.Const.get?_inter_of_contains_eq_false_left + +theorem get?_inter_of_not_mem_right [EquivBEq α] [LawfulHashable α] (h₁ : m₁.WF) (h₂ : m₂.WF) + {k : α} (not_mem : ¬k ∈ m₂) : + Const.get? (m₁ ∩ m₂) k = none := by + rw [← contains_eq_false_iff_not_mem] at not_mem + revert not_mem + simp only [Inter.inter] + simp_to_raw using Raw₀.Const.get?_inter_of_contains_eq_false_right + +/- get -/ +theorem get_inter [EquivBEq α] [LawfulHashable α] (h₁ : m₁.WF) (h₂ : m₂.WF) + {k : α} {h_mem : k ∈ m₁ ∩ m₂} : + Const.get (m₁ ∩ m₂) k h_mem = + Const.get m₁ k ((mem_inter_iff h₁ h₂).1 h_mem).1 := by + revert h_mem + simp only [Membership.mem, Inter.inter] + simp_to_raw + intro h_mem + apply Raw₀.Const.get_inter + all_goals wf_trivial + +/- getD -/ +theorem getD_inter [EquivBEq α] [LawfulHashable α] (h₁ : m₁.WF) (h₂ : m₂.WF) + {k : α} {fallback : β} : + Const.getD (m₁ ∩ m₂) k fallback = + if k ∈ m₂ then Const.getD m₁ k fallback else fallback := by + simp only [Inter.inter, Membership.mem] + simp_to_raw using Raw₀.Const.getD_inter + +theorem getD_inter_of_mem_right [EquivBEq α] [LawfulHashable α] (h₁ : m₁.WF) (h₂ : m₂.WF) + {k : α} {fallback : β} (mem : k ∈ m₂) : + Const.getD (m₁ ∩ m₂) k fallback = Const.getD m₁ k fallback := by + rw [mem_iff_contains] at mem + revert mem + simp only [Inter.inter] + simp_to_raw using Raw₀.Const.getD_inter_of_contains_right + +theorem getD_inter_of_not_mem_right [EquivBEq α] [LawfulHashable α] (h₁ : m₁.WF) (h₂ : m₂.WF) + {k : α} {fallback : β} (not_mem : ¬k ∈ m₂) : + Const.getD (m₁ ∩ m₂) k fallback = fallback := by + rw [← contains_eq_false_iff_not_mem] at not_mem + simp only [Inter.inter] + revert not_mem + simp_to_raw using Raw₀.Const.getD_inter_of_contains_eq_false_right + +theorem getD_inter_of_not_mem_left [EquivBEq α] [LawfulHashable α] (h₁ : m₁.WF) (h₂ : m₂.WF) + {k : α} {fallback : β} (not_mem : ¬k ∈ m₁) : + Const.getD (m₁ ∩ m₂) k fallback = fallback := by + rw [← contains_eq_false_iff_not_mem] at not_mem + simp only [Inter.inter] + revert not_mem + simp_to_raw using Raw₀.Const.getD_inter_of_contains_eq_false_left + +/- get! -/ +theorem get!_inter [EquivBEq α] [LawfulHashable α] [Inhabited β] (h₁ : m₁.WF) (h₂ : m₂.WF) + {k : α} : + Const.get! (m₁ ∩ m₂) k = + if k ∈ m₂ then Const.get! m₁ k else default := by + simp only [Inter.inter, Membership.mem] + simp_to_raw using Raw₀.Const.get!_inter + +theorem get!_inter_of_mem_right [EquivBEq α] [LawfulHashable α] [Inhabited β] (h₁ : m₁.WF) (h₂ : m₂.WF) + {k : α} (mem : k ∈ m₂) : + Const.get! (m₁ ∩ m₂) k = Const.get! m₁ k := by + rw [mem_iff_contains] at mem + revert mem + simp only [Inter.inter] + simp_to_raw using Raw₀.Const.get!_inter_of_contains_right + +theorem get!_inter_of_not_mem_right [EquivBEq α] [LawfulHashable α] [Inhabited β] (h₁ : m₁.WF) (h₂ : m₂.WF) + {k : α} (not_mem : ¬k ∈ m₂) : + Const.get! (m₁ ∩ m₂) k = default := by + rw [← contains_eq_false_iff_not_mem] at not_mem + simp only [Inter.inter] + revert not_mem + simp_to_raw using Raw₀.Const.get!_inter_of_contains_eq_false_right + +theorem get!_inter_of_not_mem_left [EquivBEq α] [LawfulHashable α] [Inhabited β] (h₁ : m₁.WF) (h₂ : m₂.WF) + {k : α} (not_mem : ¬k ∈ m₁) : + Const.get! (m₁ ∩ m₂) k = default := by + rw [← contains_eq_false_iff_not_mem] at not_mem + simp only [Inter.inter] + revert not_mem + simp_to_raw using Raw₀.Const.get!_inter_of_contains_eq_false_left + +end Const + +namespace Const + variable {β : Type v} {m : Raw α (fun _ => β)} variable {ρ : Type w} [ForIn Id ρ (α × β)] From 2c2a9860165dc783d2ef0a269ede12a687abf7a1 Mon Sep 17 00:00:00 2001 From: Wojciech Rozowski Date: Thu, 13 Nov 2025 10:58:38 +0100 Subject: [PATCH 45/88] Add inter lemmas --- src/Std/Data/HashMap/Lemmas.lean | 212 ++++++++++++++++++++++++++++ src/Std/Data/HashMap/RawLemmas.lean | 203 +++++++++++++++++++++++++- 2 files changed, 408 insertions(+), 7 deletions(-) diff --git a/src/Std/Data/HashMap/Lemmas.lean b/src/Std/Data/HashMap/Lemmas.lean index d81532ac5e19..103cd25f0455 100644 --- a/src/Std/Data/HashMap/Lemmas.lean +++ b/src/Std/Data/HashMap/Lemmas.lean @@ -1456,6 +1456,218 @@ theorem isEmpty_union [EquivBEq α] [LawfulHashable α] : end Union +section Inter + +variable {β : Type v} + +variable {m₁ m₂ : HashMap α β} + +@[simp] +theorem inter_eq : m₁.inter m₂ = m₁ ∩ m₂ := by + simp only [Inter.inter] + +/- contains -/ +@[simp] +theorem contains_inter [EquivBEq α] [LawfulHashable α] {k : α} : + (m₁ ∩ m₂).contains k = (m₁.contains k && m₂.contains k) := + @DHashMap.contains_inter _ _ _ _ m₁.inner m₂.inner _ _ k + +/- mem -/ +@[simp] +theorem mem_inter_iff [EquivBEq α] [LawfulHashable α] {k : α} : + k ∈ m₁ ∩ m₂ ↔ k ∈ m₁ ∧ k ∈ m₂ := + @DHashMap.mem_inter_iff _ _ _ _ m₁.inner m₂.inner _ _ k + +theorem not_mem_inter_of_not_mem_left [EquivBEq α] [LawfulHashable α] {k : α} + (not_mem : k ∉ m₁) : + k ∉ m₁ ∩ m₂ := + @DHashMap.not_mem_inter_of_not_mem_left _ _ _ _ m₁.inner m₂.inner _ _ k not_mem + +theorem not_mem_inter_of_not_mem_right [EquivBEq α] [LawfulHashable α] {k : α} + (not_mem : k ∉ m₂) : + k ∉ m₁ ∩ m₂ := + @DHashMap.not_mem_inter_of_not_mem_right _ _ _ _ m₁.inner m₂.inner _ _ k not_mem + +/- get? -/ +theorem get?_inter [EquivBEq α] [LawfulHashable α] {k : α} : + (m₁ ∩ m₂).get? k = + if k ∈ m₂ then m₁.get? k else none := + @DHashMap.Const.get?_inter _ _ _ _ m₁.inner m₂.inner _ _ k + +theorem get?_inter_of_mem_right [EquivBEq α] [LawfulHashable α] + {k : α} (mem : k ∈ m₂) : + (m₁ ∩ m₂).get? k = m₁.get? k := + @DHashMap.Const.get?_inter_of_mem_right _ _ _ _ m₁.inner m₂.inner _ _ k mem + +theorem get?_inter_of_not_mem_left [EquivBEq α] [LawfulHashable α] + {k : α} (not_mem : k ∉ m₁) : + (m₁ ∩ m₂).get? k = none := + @DHashMap.Const.get?_inter_of_not_mem_left _ _ _ _ m₁.inner m₂.inner _ _ k not_mem + +theorem get?_inter_of_not_mem_right [EquivBEq α] [LawfulHashable α] + {k : α} (not_mem : k ∉ m₂) : + (m₁ ∩ m₂).get? k = none := + @DHashMap.Const.get?_inter_of_not_mem_right _ _ _ _ m₁.inner m₂.inner _ _ k not_mem + +/- get -/ +theorem get_inter [EquivBEq α] [LawfulHashable α] + {k : α} {h_mem : k ∈ m₁ ∩ m₂} : + (m₁ ∩ m₂).get k h_mem = + m₁.get k ((mem_inter_iff.1 h_mem).1) := + @DHashMap.Const.get_inter _ _ _ _ m₁.inner m₂.inner _ _ k h_mem + +/- getD -/ +theorem getD_inter [EquivBEq α] [LawfulHashable α] {k : α} {fallback : β} : + (m₁ ∩ m₂).getD k fallback = + if k ∈ m₂ then m₁.getD k fallback else fallback := + @DHashMap.Const.getD_inter _ _ _ _ m₁.inner m₂.inner _ _ k fallback + +theorem getD_inter_of_mem_right [EquivBEq α] [LawfulHashable α] + {k : α} {fallback : β} (mem : k ∈ m₂) : + (m₁ ∩ m₂).getD k fallback = m₁.getD k fallback := + @DHashMap.Const.getD_inter_of_mem_right _ _ _ _ m₁.inner m₂.inner _ _ k fallback mem + +theorem getD_inter_of_not_mem_right [EquivBEq α] [LawfulHashable α] + {k : α} {fallback : β} (not_mem : k ∉ m₂) : + (m₁ ∩ m₂).getD k fallback = fallback := + @DHashMap.Const.getD_inter_of_not_mem_right _ _ _ _ m₁.inner m₂.inner _ _ k fallback not_mem + +theorem getD_inter_of_not_mem_left [EquivBEq α] [LawfulHashable α] + {k : α} {fallback : β} (not_mem : k ∉ m₁) : + (m₁ ∩ m₂).getD k fallback = fallback := + @DHashMap.Const.getD_inter_of_not_mem_left _ _ _ _ m₁.inner m₂.inner _ _ k fallback not_mem + +/- get! -/ +theorem get!_inter [EquivBEq α] [LawfulHashable α] {k : α} [Inhabited β] : + (m₁ ∩ m₂).get! k = + if k ∈ m₂ then m₁.get! k else default := + @DHashMap.Const.get!_inter _ _ _ _ m₁.inner m₂.inner _ _ _ k + +theorem get!_inter_of_mem_right [EquivBEq α] [LawfulHashable α] + {k : α} [Inhabited β] (mem : k ∈ m₂) : + (m₁ ∩ m₂).get! k = m₁.get! k := + @DHashMap.Const.get!_inter_of_mem_right _ _ _ _ m₁.inner m₂.inner _ _ _ k mem + +theorem get!_inter_of_not_mem_right [EquivBEq α] [LawfulHashable α] + {k : α} [Inhabited β] (not_mem : k ∉ m₂) : + (m₁ ∩ m₂).get! k = default := + @DHashMap.Const.get!_inter_of_not_mem_right _ _ _ _ m₁.inner m₂.inner _ _ _ k not_mem + +theorem get!_inter_of_not_mem_left [EquivBEq α] [LawfulHashable α] + {k : α} [Inhabited β] (not_mem : k ∉ m₁) : + (m₁ ∩ m₂).get! k = default := + @DHashMap.Const.get!_inter_of_not_mem_left _ _ _ _ m₁.inner m₂.inner _ _ _ k not_mem + +/- getKey? -/ +theorem getKey?_inter [EquivBEq α] [LawfulHashable α] {k : α} : + (m₁ ∩ m₂).getKey? k = + if k ∈ m₂ then m₁.getKey? k else none := + @DHashMap.getKey?_inter _ _ _ _ m₁.inner m₂.inner _ _ k + +theorem getKey?_inter_of_mem_right [EquivBEq α] [LawfulHashable α] + {k : α} (mem : k ∈ m₂) : + (m₁ ∩ m₂).getKey? k = m₁.getKey? k := + @DHashMap.getKey?_inter_of_mem_right _ _ _ _ m₁.inner m₂.inner _ _ k mem + +theorem getKey?_inter_of_not_mem_right [EquivBEq α] [LawfulHashable α] + {k : α} (not_mem : k ∉ m₂) : + (m₁ ∩ m₂).getKey? k = none := + @DHashMap.getKey?_inter_of_not_mem_right _ _ _ _ m₁.inner m₂.inner _ _ k not_mem + +theorem getKey?_inter_of_not_mem_left [EquivBEq α] [LawfulHashable α] + {k : α} (not_mem : k ∉ m₁) : + (m₁ ∩ m₂).getKey? k = none := + @DHashMap.getKey?_inter_of_not_mem_left _ _ _ _ m₁.inner m₂.inner _ _ k not_mem + +/- getKey -/ +theorem getKey_inter [EquivBEq α] [LawfulHashable α] + {k : α} {h_mem : k ∈ m₁ ∩ m₂} : + (m₁ ∩ m₂).getKey k h_mem = + m₁.getKey k ((mem_inter_iff.1 h_mem).1) := + @DHashMap.getKey_inter _ _ _ _ m₁.inner m₂.inner _ _ k h_mem + +/- getKeyD -/ +theorem getKeyD_inter [EquivBEq α] [LawfulHashable α] {k fallback : α} : + (m₁ ∩ m₂).getKeyD k fallback = + if k ∈ m₂ then m₁.getKeyD k fallback else fallback := + @DHashMap.getKeyD_inter _ _ _ _ m₁.inner m₂.inner _ _ k fallback + +theorem getKeyD_inter_of_mem_right [EquivBEq α] [LawfulHashable α] + {k fallback : α} (mem : k ∈ m₂) : + (m₁ ∩ m₂).getKeyD k fallback = m₁.getKeyD k fallback := + @DHashMap.getKeyD_inter_of_mem_right _ _ _ _ m₁.inner m₂.inner _ _ k fallback mem + +theorem getKeyD_inter_of_not_mem_right [EquivBEq α] [LawfulHashable α] + {k fallback : α} (not_mem : k ∉ m₂) : + (m₁ ∩ m₂).getKeyD k fallback = fallback := + @DHashMap.getKeyD_inter_of_not_mem_right _ _ _ _ m₁.inner m₂.inner _ _ k fallback not_mem + +theorem getKeyD_inter_of_not_mem_left [EquivBEq α] [LawfulHashable α] + {k fallback : α} (not_mem : k ∉ m₁) : + (m₁ ∩ m₂).getKeyD k fallback = fallback := + @DHashMap.getKeyD_inter_of_not_mem_left _ _ _ _ m₁.inner m₂.inner _ _ k fallback not_mem + +/- getKey! -/ +theorem getKey!_inter [EquivBEq α] [LawfulHashable α] [Inhabited α] {k : α} : + (m₁ ∩ m₂).getKey! k = + if k ∈ m₂ then m₁.getKey! k else default := + @DHashMap.getKey!_inter _ _ _ _ m₁.inner m₂.inner _ _ _ k + +theorem getKey!_inter_of_mem_right [EquivBEq α] [LawfulHashable α] [Inhabited α] + {k : α} (mem : k ∈ m₂) : + (m₁ ∩ m₂).getKey! k = m₁.getKey! k := + @DHashMap.getKey!_inter_of_mem_right _ _ _ _ m₁.inner m₂.inner _ _ _ k mem + +theorem getKey!_inter_of_not_mem_right [EquivBEq α] [LawfulHashable α] [Inhabited α] + {k : α} (not_mem : k ∉ m₂) : + (m₁ ∩ m₂).getKey! k = default := + @DHashMap.getKey!_inter_of_not_mem_right _ _ _ _ m₁.inner m₂.inner _ _ _ k not_mem + +theorem getKey!_inter_of_not_mem_left [EquivBEq α] [LawfulHashable α] [Inhabited α] + {k : α} (not_mem : k ∉ m₁) : + (m₁ ∩ m₂).getKey! k = default := + @DHashMap.getKey!_inter_of_not_mem_left _ _ _ _ m₁.inner m₂.inner _ _ _ k not_mem + +/- size -/ +theorem size_inter_le_size_left [EquivBEq α] [LawfulHashable α] : + (m₁ ∩ m₂).size ≤ m₁.size := + @DHashMap.size_inter_le_size_left _ _ _ _ m₁.inner m₂.inner _ _ + +theorem size_inter_le_size_right [EquivBEq α] [LawfulHashable α] : + (m₁ ∩ m₂).size ≤ m₂.size := + @DHashMap.size_inter_le_size_right _ _ _ _ m₁.inner m₂.inner _ _ + +theorem size_inter_eq_size_left [EquivBEq α] [LawfulHashable α] + (h : ∀ (a : α), a ∈ m₁ → a ∈ m₂) : + (m₁ ∩ m₂).size = m₁.size := + @DHashMap.size_inter_eq_size_left _ _ _ _ m₁.inner m₂.inner _ _ h + +theorem size_inter_eq_size_right [EquivBEq α] [LawfulHashable α] + (h : ∀ (a : α), a ∈ m₂ → a ∈ m₁) : + (m₁ ∩ m₂).size = m₂.size := + @DHashMap.size_inter_eq_size_right _ _ _ _ m₁.inner m₂.inner _ _ h + +theorem size_add_size_eq_size_union_add_size_inter [EquivBEq α] [LawfulHashable α] : + m₁.size + m₂.size = (m₁ ∪ m₂).size + (m₁ ∩ m₂).size := + @DHashMap.size_add_size_eq_size_union_add_size_inter _ _ _ _ m₁.inner m₂.inner _ _ + +/- isEmpty -/ +@[simp] +theorem isEmpty_inter_left [EquivBEq α] [LawfulHashable α] (h : m₁.isEmpty) : + (m₁ ∩ m₂).isEmpty = true := + @DHashMap.isEmpty_inter_left α _ _ _ m₁.inner m₂.inner _ _ h + +@[simp] +theorem isEmpty_inter_right [EquivBEq α] [LawfulHashable α] (h : m₂.isEmpty) : + (m₁ ∩ m₂).isEmpty = true := + @DHashMap.isEmpty_inter_right α _ _ _ m₁.inner m₂.inner _ _ h + +theorem isEmpty_inter_iff [EquivBEq α] [LawfulHashable α] : + (m₁ ∩ m₂).isEmpty ↔ ∀ k, k ∈ m₁ → k ∉ m₂ := + @DHashMap.isEmpty_inter_iff α _ _ _ m₁.inner m₂.inner _ _ + +end Inter + -- As `insertManyIfNewUnit` is really an implementation detail for `HashSet.insertMany`, -- we do not add `@[grind]` annotations to any of its lemmas. diff --git a/src/Std/Data/HashMap/RawLemmas.lean b/src/Std/Data/HashMap/RawLemmas.lean index 55b9634fe192..c15e2338315b 100644 --- a/src/Std/Data/HashMap/RawLemmas.lean +++ b/src/Std/Data/HashMap/RawLemmas.lean @@ -1430,6 +1430,8 @@ end Union section Inter +variable {β : Type v} + variable {m₁ m₂ : Raw α β} @[simp] @@ -1437,30 +1439,217 @@ theorem inter_eq : m₁.inter m₂ = m₁ ∩ m₂ := by simp only [Inter.inter] /- contains -/ - @[simp] -theorem contains_inter [EquivBEq α] [LawfulHashable α] {k : α} (h₁ : m₁.WF) (h₂ : m₂.WF) : +theorem contains_inter [EquivBEq α] [LawfulHashable α] (h₁ : m₁.WF) + (h₂ : m₂.WF) {k : α} : (m₁ ∩ m₂).contains k = (m₁.contains k && m₂.contains k) := - @DHashMap.Raw.contains_inter α (fun _ => β) _ _ m₁.inner m₂.inner _ _ k h₁.out h₂.out + @DHashMap.Raw.contains_inter _ _ _ _ m₁.inner m₂.inner _ _ k h₁.out h₂.out /- mem -/ - +@[simp] theorem mem_inter_iff [EquivBEq α] [LawfulHashable α] (h₁ : m₁.WF) (h₂ : m₂.WF) {k : α} : k ∈ m₁ ∩ m₂ ↔ k ∈ m₁ ∧ k ∈ m₂ := - @DHashMap.Raw.mem_inter_iff α (fun _ => β) _ _ m₁.inner m₂.inner _ _ h₁.out h₂.out k + @DHashMap.Raw.mem_inter_iff _ _ _ _ m₁.inner m₂.inner _ _ h₁.out h₂.out k theorem not_mem_inter_of_not_mem_left [EquivBEq α] [LawfulHashable α] (h₁ : m₁.WF) (h₂ : m₂.WF) {k : α} (not_mem : k ∉ m₁) : k ∉ m₁ ∩ m₂ := - @DHashMap.Raw.not_mem_inter_of_not_mem_left α (fun _ => β) _ _ m₁.inner m₂.inner _ _ h₁.out h₂.out k not_mem + @DHashMap.Raw.not_mem_inter_of_not_mem_left _ _ _ _ m₁.inner m₂.inner _ _ h₁.out h₂.out k not_mem theorem not_mem_inter_of_not_mem_right [EquivBEq α] [LawfulHashable α] (h₁ : m₁.WF) (h₂ : m₂.WF) {k : α} (not_mem : k ∉ m₂) : k ∉ m₁ ∩ m₂ := - @DHashMap.Raw.not_mem_inter_of_not_mem_right α (fun _ => β) _ _ m₁.inner m₂.inner _ _ h₁.out h₂.out k not_mem + @DHashMap.Raw.not_mem_inter_of_not_mem_right _ _ _ _ m₁.inner m₂.inner _ _ h₁.out h₂.out k not_mem + +/- get? -/ +theorem get?_inter [EquivBEq α] [LawfulHashable α] (h₁ : m₁.WF) (h₂ : m₂.WF) {k : α} : + get? (m₁ ∩ m₂) k = + if k ∈ m₂ then get? m₁ k else none := + @DHashMap.Raw.Const.get?_inter _ _ _ _ m₁.inner m₂.inner _ _ h₁.out h₂.out k + +theorem get?_inter_of_mem_right [EquivBEq α] [LawfulHashable α] (h₁ : m₁.WF) (h₂ : m₂.WF) + {k : α} (mem : k ∈ m₂) : + get? (m₁ ∩ m₂) k = get? m₁ k := + @DHashMap.Raw.Const.get?_inter_of_mem_right _ _ _ _ m₁.inner m₂.inner _ _ h₁.out h₂.out k mem + +theorem get?_inter_of_not_mem_left [EquivBEq α] [LawfulHashable α] (h₁ : m₁.WF) (h₂ : m₂.WF) + {k : α} (not_mem : k ∉ m₁) : + get? (m₁ ∩ m₂) k = none := + @DHashMap.Raw.Const.get?_inter_of_not_mem_left _ _ _ _ m₁.inner m₂.inner _ _ h₁.out h₂.out k not_mem + +theorem get?_inter_of_not_mem_right [EquivBEq α] [LawfulHashable α] (h₁ : m₁.WF) (h₂ : m₂.WF) + {k : α} (not_mem : k ∉ m₂) : + get? (m₁ ∩ m₂) k = none := + @DHashMap.Raw.Const.get?_inter_of_not_mem_right _ _ _ _ m₁.inner m₂.inner _ _ h₁.out h₂.out k not_mem + +/- get -/ +theorem get_inter [EquivBEq α] [LawfulHashable α] (h₁ : m₁.WF) (h₂ : m₂.WF) + {k : α} {h_mem : k ∈ m₁ ∩ m₂} : + get (m₁ ∩ m₂) k h_mem = + get m₁ k ((mem_inter_iff h₁ h₂).1 h_mem).1 := + @DHashMap.Raw.Const.get_inter _ _ _ _ m₁.inner m₂.inner _ _ h₁.out h₂.out k h_mem + +/- getD -/ +theorem getD_inter [EquivBEq α] [LawfulHashable α] (h₁ : m₁.WF) (h₂ : m₂.WF) + {k : α} {fallback : β} : + getD (m₁ ∩ m₂) k fallback = + if k ∈ m₂ then getD m₁ k fallback else fallback := + @DHashMap.Raw.Const.getD_inter _ _ _ _ m₁.inner m₂.inner _ _ h₁.out h₂.out k fallback + +theorem getD_inter_of_mem_right [EquivBEq α] [LawfulHashable α] (h₁ : m₁.WF) (h₂ : m₂.WF) + {k : α} {fallback : β} (mem : k ∈ m₂) : + getD (m₁ ∩ m₂) k fallback = getD m₁ k fallback := + @DHashMap.Raw.Const.getD_inter_of_mem_right _ _ _ _ m₁.inner m₂.inner _ _ h₁.out h₂.out k fallback mem + +theorem getD_inter_of_not_mem_right [EquivBEq α] [LawfulHashable α] (h₁ : m₁.WF) (h₂ : m₂.WF) + {k : α} {fallback : β} (not_mem : k ∉ m₂) : + getD (m₁ ∩ m₂) k fallback = fallback := + @DHashMap.Raw.Const.getD_inter_of_not_mem_right _ _ _ _ m₁.inner m₂.inner _ _ h₁.out h₂.out k fallback not_mem + +theorem getD_inter_of_not_mem_left [EquivBEq α] [LawfulHashable α] (h₁ : m₁.WF) (h₂ : m₂.WF) + {k : α} {fallback : β} (not_mem : k ∉ m₁) : + getD (m₁ ∩ m₂) k fallback = fallback := + @DHashMap.Raw.Const.getD_inter_of_not_mem_left _ _ _ _ m₁.inner m₂.inner _ _ h₁.out h₂.out k fallback not_mem + +/- get! -/ +theorem get!_inter [EquivBEq α] [LawfulHashable α] [Inhabited β] (h₁ : m₁.WF) (h₂ : m₂.WF) {k : α} : + get! (m₁ ∩ m₂) k = + if k ∈ m₂ then get! m₁ k else default := + @DHashMap.Raw.Const.get!_inter _ _ _ _ m₁.inner m₂.inner _ _ _ h₁.out h₂.out k + +theorem get!_inter_of_mem_right [EquivBEq α] [LawfulHashable α] [Inhabited β] (h₁ : m₁.WF) (h₂ : m₂.WF) + {k : α} (mem : k ∈ m₂) : + get! (m₁ ∩ m₂) k = get! m₁ k := + @DHashMap.Raw.Const.get!_inter_of_mem_right _ _ _ _ m₁.inner m₂.inner _ _ _ h₁.out h₂.out k mem + +theorem get!_inter_of_not_mem_right [EquivBEq α] [LawfulHashable α] [Inhabited β] (h₁ : m₁.WF) (h₂ : m₂.WF) + {k : α} (not_mem : k ∉ m₂) : + get! (m₁ ∩ m₂) k = default := + @DHashMap.Raw.Const.get!_inter_of_not_mem_right _ _ _ _ m₁.inner m₂.inner _ _ _ h₁.out h₂.out k not_mem + +theorem get!_inter_of_not_mem_left [EquivBEq α] [LawfulHashable α] [Inhabited β] (h₁ : m₁.WF) (h₂ : m₂.WF) + {k : α} (not_mem : k ∉ m₁) : + get! (m₁ ∩ m₂) k = default := + @DHashMap.Raw.Const.get!_inter_of_not_mem_left _ _ _ _ m₁.inner m₂.inner _ _ _ h₁.out h₂.out k not_mem + +/- getKey? -/ +theorem getKey?_inter [EquivBEq α] [LawfulHashable α] + (h₁ : m₁.WF) (h₂ : m₂.WF) {k : α} : + (m₁ ∩ m₂).getKey? k = + if k ∈ m₂ then m₁.getKey? k else none := + @DHashMap.Raw.getKey?_inter _ _ _ _ m₁.inner m₂.inner _ _ h₁.out h₂.out k + +theorem getKey?_inter_of_mem_right [EquivBEq α] [LawfulHashable α] + (h₁ : m₁.WF) (h₂ : m₂.WF) {k : α} (mem : k ∈ m₂) : + (m₁ ∩ m₂).getKey? k = m₁.getKey? k := + @DHashMap.Raw.getKey?_inter_of_mem_right _ _ _ _ m₁.inner m₂.inner _ _ h₁.out h₂.out k mem + +theorem getKey?_inter_of_not_mem_right [EquivBEq α] [LawfulHashable α] + (h₁ : m₁.WF) (h₂ : m₂.WF) {k : α} (not_mem : k ∉ m₂) : + (m₁ ∩ m₂).getKey? k = none := + @DHashMap.Raw.getKey?_inter_of_not_mem_right _ _ _ _ m₁.inner m₂.inner _ _ h₁.out h₂.out k not_mem + +theorem getKey?_inter_of_not_mem_left [EquivBEq α] [LawfulHashable α] + (h₁ : m₁.WF) (h₂ : m₂.WF) {k : α} (not_mem : k ∉ m₁) : + (m₁ ∩ m₂).getKey? k = none := + @DHashMap.Raw.getKey?_inter_of_not_mem_left _ _ _ _ m₁.inner m₂.inner _ _ h₁.out h₂.out k not_mem + +/- getKey -/ +theorem getKey_inter [EquivBEq α] [LawfulHashable α] (h₁ : m₁.WF) (h₂ : m₂.WF) + {k : α} {h_mem : k ∈ m₁ ∩ m₂} : + (m₁ ∩ m₂).getKey k h_mem = + m₁.getKey k ((mem_inter_iff h₁ h₂).1 h_mem).1 := + @DHashMap.Raw.getKey_inter _ _ _ _ m₁.inner m₂.inner _ _ h₁.out h₂.out k h_mem + +/- getKeyD -/ +theorem getKeyD_inter [EquivBEq α] [LawfulHashable α] (h₁ : m₁.WF) + (h₂ : m₂.WF) {k fallback : α} : + (m₁ ∩ m₂).getKeyD k fallback = + if k ∈ m₂ then m₁.getKeyD k fallback else fallback := + @DHashMap.Raw.getKeyD_inter _ _ _ _ m₁.inner m₂.inner _ _ h₁.out h₂.out k fallback + +theorem getKeyD_inter_of_mem_right [EquivBEq α] [LawfulHashable α] (h₁ : m₁.WF) + (h₂ : m₂.WF) {k fallback : α} (mem : k ∈ m₂) : + (m₁ ∩ m₂).getKeyD k fallback = m₁.getKeyD k fallback := + @DHashMap.Raw.getKeyD_inter_of_mem_right _ _ _ _ m₁.inner m₂.inner _ _ h₁.out h₂.out k fallback mem + +theorem getKeyD_inter_of_not_mem_right [EquivBEq α] [LawfulHashable α] (h₁ : m₁.WF) + (h₂ : m₂.WF) {k fallback : α} (not_mem : k ∉ m₂) : + (m₁ ∩ m₂).getKeyD k fallback = fallback := + @DHashMap.Raw.getKeyD_inter_of_not_mem_right _ _ _ _ m₁.inner m₂.inner _ _ h₁.out h₂.out k fallback not_mem + +theorem getKeyD_inter_of_not_mem_left [EquivBEq α] [LawfulHashable α] (h₁ : m₁.WF) + (h₂ : m₂.WF) {k fallback : α} (not_mem : k ∉ m₁) : + (m₁ ∩ m₂).getKeyD k fallback = fallback := + @DHashMap.Raw.getKeyD_inter_of_not_mem_left _ _ _ _ m₁.inner m₂.inner _ _ h₁.out h₂.out k fallback not_mem + +/- getKey! -/ +theorem getKey!_inter [EquivBEq α] [LawfulHashable α] [Inhabited α] (h₁ : m₁.WF) + (h₂ : m₂.WF) {k : α} : + (m₁ ∩ m₂).getKey! k = + if k ∈ m₂ then m₁.getKey! k else default := + @DHashMap.Raw.getKey!_inter _ _ _ _ m₁.inner m₂.inner _ _ _ h₁.out h₂.out k + +theorem getKey!_inter_of_mem_right [EquivBEq α] [LawfulHashable α] [Inhabited α] (h₁ : m₁.WF) + (h₂ : m₂.WF) {k : α} (mem : k ∈ m₂) : + (m₁ ∩ m₂).getKey! k = m₁.getKey! k := + @DHashMap.Raw.getKey!_inter_of_mem_right _ _ _ _ m₁.inner m₂.inner _ _ _ h₁.out h₂.out k mem + +theorem getKey!_inter_of_not_mem_right [EquivBEq α] [LawfulHashable α] [Inhabited α] (h₁ : m₁.WF) + (h₂ : m₂.WF) {k : α} (not_mem : k ∉ m₂) : + (m₁ ∩ m₂).getKey! k = default := + @DHashMap.Raw.getKey!_inter_of_not_mem_right _ _ _ _ m₁.inner m₂.inner _ _ _ h₁.out h₂.out k not_mem + +theorem getKey!_inter_of_not_mem_left [EquivBEq α] [LawfulHashable α] [Inhabited α] (h₁ : m₁.WF) + (h₂ : m₂.WF) {k : α} (not_mem : k ∉ m₁) : + (m₁ ∩ m₂).getKey! k = default := + @DHashMap.Raw.getKey!_inter_of_not_mem_left _ _ _ _ m₁.inner m₂.inner _ _ _ h₁.out h₂.out k not_mem + +/- size -/ +theorem size_inter_le_size_left [EquivBEq α] [LawfulHashable α] + (h₁ : m₁.WF) (h₂ : m₂.WF) : + (m₁ ∩ m₂).size ≤ m₁.size := + @DHashMap.Raw.size_inter_le_size_left _ _ _ _ m₁.inner m₂.inner _ _ h₁.out h₂.out + +theorem size_inter_le_size_right [EquivBEq α] [LawfulHashable α] + (h₁ : m₁.WF) (h₂ : m₂.WF) : + (m₁ ∩ m₂).size ≤ m₂.size := + @DHashMap.Raw.size_inter_le_size_right _ _ _ _ m₁.inner m₂.inner _ _ h₁.out h₂.out + +theorem size_inter_eq_size_left [EquivBEq α] [LawfulHashable α] + (h₁ : m₁.WF) (h₂ : m₂.WF) + (h : ∀ (a : α), a ∈ m₁ → a ∈ m₂) : + (m₁ ∩ m₂).size = m₁.size := + @DHashMap.Raw.size_inter_eq_size_left _ _ _ _ m₁.inner m₂.inner _ _ h₁.out h₂.out h + +theorem size_inter_eq_size_right [EquivBEq α] [LawfulHashable α] + (h₁ : m₁.WF) (h₂ : m₂.WF) + (h : ∀ (a : α), a ∈ m₂ → a ∈ m₁) : + (m₁ ∩ m₂).size = m₂.size := + @DHashMap.Raw.size_inter_eq_size_right _ _ _ _ m₁.inner m₂.inner _ _ h₁.out h₂.out h + +theorem size_add_size_eq_size_union_add_size_inter [EquivBEq α] [LawfulHashable α] + (h₁ : m₁.WF) (h₂ : m₂.WF) : + m₁.size + m₂.size = (m₁ ∪ m₂).size + (m₁ ∩ m₂).size := + @DHashMap.Raw.size_add_size_eq_size_union_add_size_inter _ _ _ _ m₁.inner m₂.inner _ _ h₁.out h₂.out + +/- isEmpty -/ +@[simp] +theorem isEmpty_inter_left [EquivBEq α] [LawfulHashable α] (h₁ : m₁.WF) (h₂ : m₂.WF) (h : m₁.isEmpty) : + (m₁ ∩ m₂).isEmpty = true := + @DHashMap.Raw.isEmpty_inter_left _ _ _ _ m₁.inner m₂.inner _ _ h₁.out h₂.out h + +@[simp] +theorem isEmpty_inter_right [EquivBEq α] [LawfulHashable α] (h₁ : m₁.WF) (h₂ : m₂.WF) (h : m₂.isEmpty) : + (m₁ ∩ m₂).isEmpty = true := + @DHashMap.Raw.isEmpty_inter_right _ _ _ _ m₁.inner m₂.inner _ _ h₁.out h₂.out h + +theorem isEmpty_inter_iff [EquivBEq α] [LawfulHashable α] (h₁ : m₁.WF) (h₂ : m₂.WF) : + (m₁ ∩ m₂).isEmpty ↔ ∀ k, k ∈ m₁ → k ∉ m₂ := + @DHashMap.Raw.isEmpty_inter_iff _ _ _ _ m₁.inner m₂.inner _ _ h₁.out h₂.out end Inter From b5a421b1663097d10cc711f6117219324d3266fd Mon Sep 17 00:00:00 2001 From: Wojciech Rozowski Date: Thu, 13 Nov 2025 11:06:45 +0100 Subject: [PATCH 46/88] Add hashSet lemmas --- src/Std/Data/HashSet/Basic.lean | 11 ++ src/Std/Data/HashSet/Lemmas.lean | 140 ++++++++++++++++++++++++ src/Std/Data/HashSet/Raw.lean | 11 ++ src/Std/Data/HashSet/RawLemmas.lean | 161 ++++++++++++++++++++++++++++ 4 files changed, 323 insertions(+) diff --git a/src/Std/Data/HashSet/Basic.lean b/src/Std/Data/HashSet/Basic.lean index 380722b10f8a..c74fa8b0d9c6 100644 --- a/src/Std/Data/HashSet/Basic.lean +++ b/src/Std/Data/HashSet/Basic.lean @@ -243,6 +243,17 @@ This function always merges the smaller set into the larger set, so the expected instance [BEq α] [Hashable α] : Union (HashSet α) := ⟨union⟩ +/-- +Computes the intersection of the given hash sets. The result will only contain entries from the first map. + +This function always merges the smaller set into the larger set, so the expected runtime is +`O(min(m₁.size, m₂.size))`. +-/ +@[inline] def inter [BEq α] [Hashable α] (m₁ m₂ : HashSet α) : HashSet α := + ⟨HashMap.inter m₁.inner m₂.inner⟩ + +instance [BEq α] [Hashable α] : Inter (HashSet α) := ⟨inter⟩ + section Unverified /-! We currently do not provide lemmas for the functions below. -/ diff --git a/src/Std/Data/HashSet/Lemmas.lean b/src/Std/Data/HashSet/Lemmas.lean index 3ee84bdaac41..02a6d11679d1 100644 --- a/src/Std/Data/HashSet/Lemmas.lean +++ b/src/Std/Data/HashSet/Lemmas.lean @@ -849,6 +849,146 @@ theorem isEmpty_union [EquivBEq α] [LawfulHashable α] : end Union +section Inter + +variable {m₁ m₂ : HashSet α} + +@[simp] +theorem inter_eq : m₁.inter m₂ = m₁ ∩ m₂ := by + simp only [Inter.inter] + +/- contains -/ +@[simp] +theorem contains_inter [EquivBEq α] [LawfulHashable α] {k : α} : + (m₁ ∩ m₂).contains k = (m₁.contains k && m₂.contains k) := + @HashMap.contains_inter _ _ _ _ m₁.inner m₂.inner _ _ k + +/- mem -/ +@[simp] +theorem mem_inter_iff [EquivBEq α] [LawfulHashable α] {k : α} : + (m₁ ∩ m₂).contains k ↔ m₁.contains k ∧ m₂.contains k := + @HashMap.mem_inter_iff _ _ _ _ m₁.inner m₂.inner _ _ k + +theorem not_mem_inter_of_not_mem_left [EquivBEq α] [LawfulHashable α] {k : α} + (not_mem : k ∉ m₁) : + k ∉ m₁ ∩ m₂ := + @HashMap.not_mem_inter_of_not_mem_left _ _ _ _ m₁.inner m₂.inner _ _ k not_mem + +theorem not_mem_inter_of_not_mem_right [EquivBEq α] [LawfulHashable α] {k : α} + (not_mem : k ∉ m₂) : + k ∉ m₁ ∩ m₂ := + @HashMap.not_mem_inter_of_not_mem_right _ _ _ _ m₁.inner m₂.inner _ _ k not_mem + +/- get? -/ +theorem get?_inter [EquivBEq α] [LawfulHashable α] {k : α} : + (m₁ ∩ m₂).get? k = + if k ∈ m₂ then m₁.get? k else none := + @HashMap.getKey?_inter _ _ _ _ m₁.inner m₂.inner _ _ k + +theorem get?_inter_of_mem_right [EquivBEq α] [LawfulHashable α] + {k : α} (mem : k ∈ m₂) : + (m₁ ∩ m₂).get? k = m₁.get? k := + @HashMap.getKey?_inter_of_mem_right _ _ _ _ m₁.inner m₂.inner _ _ k mem + +theorem get?_inter_of_not_mem_left [EquivBEq α] [LawfulHashable α] + {k : α} (not_mem : k ∉ m₁) : + (m₁ ∩ m₂).get? k = none := + @HashMap.getKey?_inter_of_not_mem_left _ _ _ _ m₁.inner m₂.inner _ _ k not_mem + +theorem get?_inter_of_not_mem_right [EquivBEq α] [LawfulHashable α] + {k : α} (not_mem : k ∉ m₂) : + (m₁ ∩ m₂).get? k = none := + @HashMap.getKey?_inter_of_not_mem_right _ _ _ _ m₁.inner m₂.inner _ _ k not_mem + +/- get -/ +theorem get_inter [EquivBEq α] [LawfulHashable α] + {k : α} {h_mem : k ∈ m₁ ∩ m₂} : + (m₁ ∩ m₂).get k h_mem = + m₁.get k ((mem_inter_iff.1 h_mem).1) := + @HashMap.getKey_inter _ _ _ _ m₁.inner m₂.inner _ _ k h_mem + +/- getD -/ +theorem getD_inter [EquivBEq α] [LawfulHashable α] {k fallback : α} : + (m₁ ∩ m₂).getD k fallback = + if k ∈ m₂ then m₁.getD k fallback else fallback := + @HashMap.getKeyD_inter _ _ _ _ m₁.inner m₂.inner _ _ k fallback + +theorem getD_inter_of_mem_right [EquivBEq α] [LawfulHashable α] + {k fallback : α} (mem : k ∈ m₂) : + (m₁ ∩ m₂).getD k fallback = m₁.getD k fallback := + @HashMap.getKeyD_inter_of_mem_right _ _ _ _ m₁.inner m₂.inner _ _ k fallback mem + +theorem getD_inter_of_not_mem_right [EquivBEq α] [LawfulHashable α] + {k fallback : α} (not_mem : k ∉ m₂) : + (m₁ ∩ m₂).getD k fallback = fallback := + @HashMap.getKeyD_inter_of_not_mem_right _ _ _ _ m₁.inner m₂.inner _ _ k fallback not_mem + +theorem getD_inter_of_not_mem_left [EquivBEq α] [LawfulHashable α] + {k fallback : α} (not_mem : k ∉ m₁) : + (m₁ ∩ m₂).getD k fallback = fallback := + @HashMap.getKeyD_inter_of_not_mem_left _ _ _ _ m₁.inner m₂.inner _ _ k fallback not_mem + +/- get! -/ +theorem get!_inter [EquivBEq α] [LawfulHashable α] [Inhabited α] {k : α} : + (m₁ ∩ m₂).get! k = + if k ∈ m₂ then m₁.get! k else default := + @HashMap.getKey!_inter _ _ _ _ m₁.inner m₂.inner _ _ _ k + +theorem get!_inter_of_mem_right [EquivBEq α] [LawfulHashable α] [Inhabited α] + {k : α} (mem : k ∈ m₂) : + (m₁ ∩ m₂).get! k = m₁.get! k := + @HashMap.getKey!_inter_of_mem_right _ _ _ _ m₁.inner m₂.inner _ _ _ k mem + +theorem get!_inter_of_not_mem_right [EquivBEq α] [LawfulHashable α] [Inhabited α] + {k : α} (not_mem : k ∉ m₂) : + (m₁ ∩ m₂).get! k = default := + @HashMap.getKey!_inter_of_not_mem_right _ _ _ _ m₁.inner m₂.inner _ _ _ k not_mem + +theorem get!_inter_of_not_mem_left [EquivBEq α] [LawfulHashable α] [Inhabited α] + {k : α} (not_mem : k ∉ m₁) : + (m₁ ∩ m₂).get! k = default := + @HashMap.getKey!_inter_of_not_mem_left _ _ _ _ m₁.inner m₂.inner _ _ _ k not_mem + +/- size -/ +theorem size_inter_le_size_left [EquivBEq α] [LawfulHashable α] : + (m₁ ∩ m₂).size ≤ m₁.size := + @HashMap.size_inter_le_size_left _ _ _ _ m₁.inner m₂.inner _ _ + +theorem size_inter_le_size_right [EquivBEq α] [LawfulHashable α] : + (m₁ ∩ m₂).size ≤ m₂.size := + @HashMap.size_inter_le_size_right _ _ _ _ m₁.inner m₂.inner _ _ + +theorem size_inter_eq_size_left [EquivBEq α] [LawfulHashable α] + (h : ∀ (a : α), a ∈ m₁ → a ∈ m₂) : + (m₁ ∩ m₂).size = m₁.size := + @HashMap.size_inter_eq_size_left _ _ _ _ m₁.inner m₂.inner _ _ h + +theorem size_inter_eq_size_right [EquivBEq α] [LawfulHashable α] + (h : ∀ (a : α), a ∈ m₂ → a ∈ m₁) : + (m₁ ∩ m₂).size = m₂.size := + @HashMap.size_inter_eq_size_right _ _ _ _ m₁.inner m₂.inner _ _ h + +theorem size_add_size_eq_size_union_add_size_inter [EquivBEq α] [LawfulHashable α] : + m₁.size + m₂.size = (m₁ ∪ m₂).size + (m₁ ∩ m₂).size := + @HashMap.size_add_size_eq_size_union_add_size_inter _ _ _ _ m₁.inner m₂.inner _ _ + +/- isEmpty -/ +@[simp] +theorem isEmpty_inter_left [EquivBEq α] [LawfulHashable α] (h : m₁.isEmpty) : + (m₁ ∩ m₂).isEmpty = true := + @HashMap.isEmpty_inter_left α _ _ _ m₁.inner m₂.inner _ _ h + +@[simp] +theorem isEmpty_inter_right [EquivBEq α] [LawfulHashable α] (h : m₂.isEmpty) : + (m₁ ∩ m₂).isEmpty = true := + @HashMap.isEmpty_inter_right α _ _ _ m₁.inner m₂.inner _ _ h + +theorem isEmpty_inter_iff [EquivBEq α] [LawfulHashable α] : + (m₁ ∩ m₂).isEmpty ↔ ∀ k, k ∈ m₁ → k ∉ m₂ := + @HashMap.isEmpty_inter_iff α _ _ _ m₁.inner m₂.inner _ _ + +end Inter + section @[simp, grind =] diff --git a/src/Std/Data/HashSet/Raw.lean b/src/Std/Data/HashSet/Raw.lean index cc634d2218ac..fef0430e5652 100644 --- a/src/Std/Data/HashSet/Raw.lean +++ b/src/Std/Data/HashSet/Raw.lean @@ -232,6 +232,17 @@ This function always merges the smaller set into the larger set, so the expected instance [BEq α] [Hashable α] : Union (Raw α) := ⟨union⟩ +/-- +Computes the intersection of the given hash sets. The result will only contain entries from the first map. + +This function always merges the smaller set into the larger set, so the expected runtime is +`O(min(m₁.size, m₂.size))`. +-/ +@[inline] def inter [BEq α] [Hashable α] (m₁ m₂ : Raw α) : Raw α := + ⟨HashMap.Raw.inter m₁.inner m₂.inner⟩ + +instance [BEq α] [Hashable α] : Inter (Raw α) := ⟨inter⟩ + section Unverified /-! We currently do not provide lemmas for the functions below. -/ diff --git a/src/Std/Data/HashSet/RawLemmas.lean b/src/Std/Data/HashSet/RawLemmas.lean index 7455f3e98532..eaeff1bfe46a 100644 --- a/src/Std/Data/HashSet/RawLemmas.lean +++ b/src/Std/Data/HashSet/RawLemmas.lean @@ -888,6 +888,167 @@ theorem isEmpty_union [EquivBEq α] [LawfulHashable α] (h₁ : m₁.WF) (h₂ : end Union +section Inter + +variable {m₁ m₂ : Raw α} + +@[simp] +theorem inter_eq : m₁.inter m₂ = m₁ ∩ m₂ := by + simp only [Inter.inter] + +/- contains -/ +@[simp] +theorem contains_inter [EquivBEq α] [LawfulHashable α] (h₁ : m₁.WF) + (h₂ : m₂.WF) {k : α} : + (m₁ ∩ m₂).contains k = (m₁.contains k && m₂.contains k) := + @HashMap.Raw.contains_inter _ _ _ _ m₁.inner m₂.inner _ _ h₁.out h₂.out k + +/- mem -/ +@[simp] +theorem mem_inter_iff [EquivBEq α] [LawfulHashable α] (h₁ : m₁.WF) + (h₂ : m₂.WF) {k : α} : + k ∈ m₁ ∩ m₂ ↔ k ∈ m₁ ∧ k ∈ m₂ := + @HashMap.Raw.mem_inter_iff _ _ _ _ m₁.inner m₂.inner _ _ h₁.out h₂.out k + +theorem not_mem_inter_of_not_mem_left [EquivBEq α] [LawfulHashable α] + (h₁ : m₁.WF) (h₂ : m₂.WF) {k : α} + (not_mem : k ∉ m₁) : + k ∉ m₁ ∩ m₂ := + @HashMap.Raw.not_mem_inter_of_not_mem_left _ _ _ _ m₁.inner m₂.inner _ _ h₁.out h₂.out k not_mem + +theorem not_mem_inter_of_not_mem_right [EquivBEq α] [LawfulHashable α] + (h₁ : m₁.WF) (h₂ : m₂.WF) {k : α} + (not_mem : k ∉ m₂) : + k ∉ m₁ ∩ m₂ := + @HashMap.Raw.not_mem_inter_of_not_mem_right _ _ _ _ m₁.inner m₂.inner _ _ h₁.out h₂.out k not_mem + +/- get? -/ +theorem get?_inter [EquivBEq α] [LawfulHashable α] + (h₁ : m₁.WF) (h₂ : m₂.WF) + {k : α} : + (m₁ ∩ m₂).get? k = + if k ∈ m₂ then m₁.get? k else none := + @HashMap.Raw.getKey?_inter _ _ _ _ m₁.inner m₂.inner _ _ h₁.out h₂.out k + +theorem get?_inter_of_mem_right [EquivBEq α] [LawfulHashable α] + (h₁ : m₁.WF) (h₂ : m₂.WF) + {k : α} (mem : k ∈ m₂) : + (m₁ ∩ m₂).get? k = m₁.get? k := + @HashMap.Raw.getKey?_inter_of_mem_right _ _ _ _ m₁.inner m₂.inner _ _ h₁.out h₂.out k mem + +theorem get?_inter_of_not_mem_left [EquivBEq α] [LawfulHashable α] + (h₁ : m₁.WF) (h₂ : m₂.WF) + {k : α} (not_mem : k ∉ m₁) : + (m₁ ∩ m₂).get? k = none := + @HashMap.Raw.getKey?_inter_of_not_mem_left _ _ _ _ m₁.inner m₂.inner _ _ h₁.out h₂.out k not_mem + +theorem get?_inter_of_not_mem_right [EquivBEq α] [LawfulHashable α] + (h₁ : m₁.WF) (h₂ : m₂.WF) + {k : α} (not_mem : k ∉ m₂) : + (m₁ ∩ m₂).get? k = none := + @HashMap.Raw.getKey?_inter_of_not_mem_right _ _ _ _ m₁.inner m₂.inner _ _ h₁.out h₂.out k not_mem + +/- get -/ +theorem get_inter [EquivBEq α] [LawfulHashable α] (h₁ : m₁.WF) (h₂ : m₂.WF) + {k : α} {h_mem : k ∈ m₁ ∩ m₂} : + (m₁ ∩ m₂).get k h_mem = + m₁.get k ((mem_inter_iff h₁ h₂).1 h_mem).1 := + @HashMap.Raw.getKey_inter _ _ _ _ m₁.inner m₂.inner _ _ h₁.out h₂.out k h_mem + +/- getD -/ +theorem getD_inter [EquivBEq α] + [LawfulHashable α] (h₁ : m₁.WF) + (h₂ : m₂.WF) {k fallback : α} : + (m₁ ∩ m₂).getD k fallback = + if k ∈ m₂ then m₁.getD k fallback else fallback := + @HashMap.Raw.getKeyD_inter _ _ _ _ m₁.inner m₂.inner _ _ h₁.out h₂.out k fallback + +theorem getD_inter_of_mem_right [EquivBEq α] [LawfulHashable α] (h₁ : m₁.WF) + (h₂ : m₂.WF) {k fallback : α} (mem : k ∈ m₂) : + (m₁ ∩ m₂).getD k fallback = m₁.getD k fallback := + @HashMap.Raw.getKeyD_inter_of_mem_right _ _ _ _ m₁.inner m₂.inner _ _ h₁.out h₂.out k fallback mem + +theorem getD_inter_of_not_mem_right [EquivBEq α] [LawfulHashable α] (h₁ : m₁.WF) + (h₂ : m₂.WF) {k fallback : α} (not_mem : k ∉ m₂) : + (m₁ ∩ m₂).getD k fallback = fallback := + @HashMap.Raw.getKeyD_inter_of_not_mem_right _ _ _ _ m₁.inner m₂.inner _ _ h₁.out h₂.out k fallback not_mem + +theorem getD_inter_of_not_mem_left [EquivBEq α] [LawfulHashable α] (h₁ : m₁.WF) + (h₂ : m₂.WF) {k fallback : α} (not_mem : k ∉ m₁) : + (m₁ ∩ m₂).getD k fallback = fallback := + @HashMap.Raw.getKeyD_inter_of_not_mem_left _ _ _ _ m₁.inner m₂.inner _ _ h₁.out h₂.out k fallback not_mem + +/- get! -/ +theorem get!_inter [EquivBEq α] [LawfulHashable α] [Inhabited α] + (h₁ : m₁.WF) + (h₂ : m₂.WF) {k : α} : + (m₁ ∩ m₂).get! k = + if k ∈ m₂ then m₁.get! k else default := + @HashMap.Raw.getKey!_inter _ _ _ _ m₁.inner m₂.inner _ _ _ h₁.out h₂.out k + +theorem get!_inter_of_mem_right [Inhabited α] + [EquivBEq α] [LawfulHashable α] (h₁ : m₁.WF) (h₂ : m₂.WF) {k : α} + (mem : k ∈ m₂) : + (m₁ ∩ m₂).get! k = m₁.get! k := + @HashMap.Raw.getKey!_inter_of_mem_right _ _ _ _ m₁.inner m₂.inner _ _ _ h₁.out h₂.out k mem + +theorem get!_inter_of_not_mem_right [Inhabited α] + [EquivBEq α] [LawfulHashable α] (h₁ : m₁.WF) (h₂ : m₂.WF) {k : α} + (not_mem : k ∉ m₂) : + (m₁ ∩ m₂).get! k = default := + @HashMap.Raw.getKey!_inter_of_not_mem_right _ _ _ _ m₁.inner m₂.inner _ _ _ h₁.out h₂.out k not_mem + +theorem get!_inter_of_not_mem_left [Inhabited α] + [EquivBEq α] [LawfulHashable α] (h₁ : m₁.WF) (h₂ : m₂.WF) {k : α} + (not_mem : k ∉ m₁) : + (m₁ ∩ m₂).get! k = default := + @HashMap.Raw.getKey!_inter_of_not_mem_left _ _ _ _ m₁.inner m₂.inner _ _ _ h₁.out h₂.out k not_mem + +/- size -/ +theorem size_inter_le_size_left [EquivBEq α] [LawfulHashable α] (h₁ : m₁.WF) + (h₂ : m₂.WF) : + (m₁ ∩ m₂).size ≤ m₁.size := + @HashMap.Raw.size_inter_le_size_left _ _ _ _ m₁.inner m₂.inner _ _ h₁.out h₂.out + +theorem size_inter_le_size_right [EquivBEq α] [LawfulHashable α] (h₁ : m₁.WF) + (h₂ : m₂.WF) : + (m₁ ∩ m₂).size ≤ m₂.size := + @HashMap.Raw.size_inter_le_size_right _ _ _ _ m₁.inner m₂.inner _ _ h₁.out h₂.out + +theorem size_inter_eq_size_left [EquivBEq α] [LawfulHashable α] (h₁ : m₁.WF) + (h₂ : m₂.WF) + (h : ∀ (a : α), a ∈ m₁ → a ∈ m₂) : + (m₁ ∩ m₂).size = m₁.size := + @HashMap.Raw.size_inter_eq_size_left _ _ _ _ m₁.inner m₂.inner _ _ h₁.out h₂.out h + +theorem size_inter_eq_size_right [EquivBEq α] [LawfulHashable α] (h₁ : m₁.WF) + (h₂ : m₂.WF) + (h : ∀ (a : α), a ∈ m₂ → a ∈ m₁) : + (m₁ ∩ m₂).size = m₂.size := + @HashMap.Raw.size_inter_eq_size_right _ _ _ _ m₁.inner m₂.inner _ _ h₁.out h₂.out h + +theorem size_add_size_eq_size_union_add_size_inter [EquivBEq α] [LawfulHashable α] + (h₁ : m₁.WF) (h₂ : m₂.WF) : + m₁.size + m₂.size = (m₁ ∪ m₂).size + (m₁ ∩ m₂).size := + @HashMap.Raw.size_add_size_eq_size_union_add_size_inter _ _ _ _ m₁.inner m₂.inner _ _ h₁.out h₂.out + +/- isEmpty -/ +@[simp] +theorem isEmpty_inter_left [EquivBEq α] [LawfulHashable α] (h₁ : m₁.WF) (h₂ : m₂.WF) (h : m₁.isEmpty) : + (m₁ ∩ m₂).isEmpty = true := + @HashMap.Raw.isEmpty_inter_left _ _ _ _ m₁.inner m₂.inner _ _ h₁.out h₂.out h + +@[simp] +theorem isEmpty_inter_right [EquivBEq α] [LawfulHashable α] (h₁ : m₁.WF) (h₂ : m₂.WF) (h : m₂.isEmpty) : + (m₁ ∩ m₂).isEmpty = true := + @HashMap.Raw.isEmpty_inter_right _ _ _ _ m₁.inner m₂.inner _ _ h₁.out h₂.out h + +theorem isEmpty_inter_iff [EquivBEq α] [LawfulHashable α] (h₁ : m₁.WF) (h₂ : m₂.WF) : + (m₁ ∩ m₂).isEmpty ↔ ∀ k, k ∈ m₁ → k ∉ m₂ := + @HashMap.Raw.isEmpty_inter_iff _ _ _ _ m₁.inner m₂.inner _ _ h₁.out h₂.out + +end Inter + @[simp, grind =] theorem ofList_nil : ofList ([] : List α) = ∅ := From 8be273fa1003a4c86c1a71f1fcfd2e8f7c1750e1 Mon Sep 17 00:00:00 2001 From: Wojciech Rozowski Date: Thu, 13 Nov 2025 19:21:27 +0100 Subject: [PATCH 47/88] Slight reorganisation --- src/Std/Data/DHashMap/Internal/WF.lean | 91 -------------------- src/Std/Data/Internal/List/Associative.lean | 95 +++++++++++++++++++++ 2 files changed, 95 insertions(+), 91 deletions(-) diff --git a/src/Std/Data/DHashMap/Internal/WF.lean b/src/Std/Data/DHashMap/Internal/WF.lean index aa7dbf9c1caa..6c4140cd1576 100644 --- a/src/Std/Data/DHashMap/Internal/WF.lean +++ b/src/Std/Data/DHashMap/Internal/WF.lean @@ -1375,45 +1375,6 @@ theorem foldl_perm_cong [BEq α] {init₁ init₂ : List ((a : α) × β a)} {l · exact (h₂ h init₁ init₂ h₃ h₁).1 · exact (h₂ h init₁ init₂ h₃ h₁).2 -/-- Internal implementation detail of the hash map -/ -def List.interSmallerFn [BEq α] (l sofar : List ((a : α) × β a)) (k : α) : List ((a : α) × β a) := - match List.getEntry? k l with - | some kv' => List.insertEntry kv'.1 kv'.2 sofar - | none => sofar - -theorem List.distinctKeys_interSmallerFn [BEq α] [PartialEquivBEq α] - (l sofar : List ((a : α) × β a)) (k : α) (hs : DistinctKeys sofar) : - DistinctKeys (interSmallerFn l sofar k) := by - rw [interSmallerFn] - split - · exact hs.insertEntry - · exact hs - -theorem List.getEntry?_interSmallerFn [BEq α] [PartialEquivBEq α] (l sofar : List ((a : α) × β a)) (k k' : α) : - List.getEntry? k' (interSmallerFn l sofar k) = - ((List.getEntry? k' l).filter (fun kv => k == kv.1)).or (sofar.getEntry? k') := by - rw [interSmallerFn] - split - · rename_i kv hkv - rw [getEntry?_insertEntry] - split <;> rename_i hk - · have hk' : k == k' := (BEq.trans (BEq.symm (List.beq_of_getEntry?_eq_some hkv)) hk) - simp [← List.getEntry?_congr hk', hkv, Option.filter_some, BEq.trans hk' (BEq.symm hk)] - · rw [Option.or_eq_right_of_none] - apply Option.filter_eq_none_iff.2 - intro p hp - have hp' := List.beq_of_getEntry?_eq_some hp - have hkv' := List.beq_of_getEntry?_eq_some hkv - exact fun h => hk (BEq.trans hkv' (BEq.trans h hp')) - · rename_i hk - rw [Option.or_eq_right_of_none] - apply Option.filter_eq_none_iff.2 - intro p hp - have := List.beq_of_getEntry?_eq_some hp - intro hkp - have := BEq.trans hkp this - simp [List.getEntry?_congr this, hp] at hk - theorem toListModel_interSmallerFnₘ [BEq α] [EquivBEq α] [Hashable α] [LawfulHashable α] (m sofar : Raw₀ α β) (l : List ((a : α) × β a)) (hm : Raw.WFImp m.1) (hs : Raw.WFImp sofar.1) (k : α) (hml : toListModel sofar.1.buckets ~ l) : @@ -1424,58 +1385,6 @@ theorem toListModel_interSmallerFnₘ [BEq α] [EquivBEq α] [Hashable α] [Lawf · simpa [*] using (toListModel_insertₘ hs).trans (List.insertEntry_of_perm hs.distinct hml) · simp [*] -/-- Internal implementation detail of the hash map -/ -def List.interSmaller [BEq α] (l₁ l₂ : List ((a : α) × β a)) : List ((a : α) × β a) := - l₂.foldl (fun sofar kv => List.interSmallerFn l₁ sofar kv.1) [] - -@[simp] -theorem Option.filter_false {o : Option α} : o.filter (fun _ => false) = none := by - cases o <;> simp - -theorem Option.filter_or {o : Option α} {p q : α → Bool} : o.filter (fun a => p a || q a) = - (o.filter p).or (o.filter q) := by - cases o with - | none => simp - | some a => - simp [Option.filter_some] - cases p a <;> cases q a <;> simp - -theorem List.getEntry?_interSmaller [BEq α] [PartialEquivBEq α] (l₁ l₂ : List ((a : α) × β a)) (k : α) : - List.getEntry? k (interSmaller l₁ l₂) = (List.getEntry? k l₁).filter (fun kv => containsKey kv.1 l₂) := by - rw [interSmaller] - suffices ∀ l₃, - List.getEntry? k (foldl (fun sofar kv => interSmallerFn l₁ sofar kv.fst) l₃ l₂) = - (Option.filter (fun kv => containsKey kv.fst l₂) (List.getEntry? k l₁)).or (List.getEntry? k l₃) by - simpa using this [] - intro l₃ - induction l₂ using assoc_induction generalizing l₃ with - | nil => simp - | cons k v tl ih => - rw [List.foldl_cons, ih] - simp only [List.containsKey_cons, Bool.or_comm (k == _), Option.filter_or, Option.or_assoc, - List.getEntry?_interSmallerFn] - -theorem List.distinctKeys_interSmaller [BEq α] [PartialEquivBEq α] {l₁ l₂ : List ((a : α) × β a)} : - DistinctKeys (interSmaller l₁ l₂) := by - rw [interSmaller] - suffices ∀ l, DistinctKeys l → DistinctKeys (l₂.foldl (fun sofar kv => List.interSmallerFn l₁ sofar kv.1) l) by - simpa using this [] (by simp) - intro l hl - induction l₂ generalizing l with - | nil => simpa - | cons ht tl ih => - rw [List.foldl_cons] - apply ih - exact distinctKeys_interSmallerFn _ _ _ hl - -theorem List.interSmaller_perm_filter [BEq α] [EquivBEq α] (l₁ l₂ : List ((a : α) × β a)) (h₁ : DistinctKeys l₁) : - List.interSmaller l₁ l₂ ~ l₁.filter (fun kv => containsKey kv.1 l₂) := by - apply List.getEntry?_ext - · exact List.distinctKeys_interSmaller - · exact h₁.filter (f := fun k v => containsKey k l₂) - · intro k' - rw [List.getEntry?_filter h₁, List.getEntry?_interSmaller] - theorem toListModel_interSmallerₘ [BEq α] [EquivBEq α] [Hashable α] [LawfulHashable α] (m₁ : Raw₀ α β) (m₂ : Raw α β) (hm₁ : Raw.WFImp m₁.1) : toListModel (m₁.interSmallerₘ m₂).1.buckets ~ diff --git a/src/Std/Data/Internal/List/Associative.lean b/src/Std/Data/Internal/List/Associative.lean index b026eecd46cf..d16d183287f9 100644 --- a/src/Std/Data/Internal/List/Associative.lean +++ b/src/Std/Data/Internal/List/Associative.lean @@ -9048,4 +9048,99 @@ end Const end Max +section InterSmaller + +/-- Internal implementation detail of the hash map -/ +def interSmallerFn [BEq α] (l sofar : List ((a : α) × β a)) (k : α) : List ((a : α) × β a) := + match List.getEntry? k l with + | some kv' => List.insertEntry kv'.1 kv'.2 sofar + | none => sofar + +theorem List.distinctKeys_interSmallerFn [BEq α] [PartialEquivBEq α] + (l sofar : List ((a : α) × β a)) (k : α) (hs : DistinctKeys sofar) : + DistinctKeys (List.interSmallerFn l sofar k) := by + rw [List.interSmallerFn] + split + · exact hs.insertEntry + · exact hs + +theorem getEntry?_interSmallerFn [BEq α] [PartialEquivBEq α] (l sofar : List ((a : α) × β a)) (k k' : α) : + List.getEntry? k' (List.interSmallerFn l sofar k) = + ((List.getEntry? k' l).filter (fun kv => k == kv.1)).or (List.getEntry? k' sofar) := by + rw [List.interSmallerFn] + split + · rename_i kv hkv + rw [getEntry?_insertEntry] + split <;> rename_i hk + · have hk' : k == k' := (BEq.trans (BEq.symm (List.beq_of_getEntry?_eq_some hkv)) hk) + simp [← List.getEntry?_congr hk', hkv, Option.filter_some, BEq.trans hk' (BEq.symm hk)] + · rw [Option.or_eq_right_of_none] + apply Option.filter_eq_none_iff.2 + intro p hp + have hp' := List.beq_of_getEntry?_eq_some hp + have hkv' := List.beq_of_getEntry?_eq_some hkv + exact fun h => hk (BEq.trans hkv' (BEq.trans h hp')) + · rename_i hk + rw [Option.or_eq_right_of_none] + apply Option.filter_eq_none_iff.2 + intro p hp + have := List.beq_of_getEntry?_eq_some hp + intro hkp + have := BEq.trans hkp this + simp [List.getEntry?_congr this, hp] at hk + +/-- Internal implementation detail of the hash map -/ +def interSmaller [BEq α] (l₁ l₂ : List ((a : α) × β a)) : List ((a : α) × β a) := + l₂.foldl (fun sofar kv => List.interSmallerFn l₁ sofar kv.1) [] + +@[simp] +theorem Option.filter_false {o : Option α} : o.filter (fun _ => false) = none := by + cases o <;> simp + +theorem Option.filter_or {o : Option α} {p q : α → Bool} : o.filter (fun a => p a || q a) = + (o.filter p).or (o.filter q) := by + cases o with + | none => simp + | some a => + simp [Option.filter_some] + cases p a <;> cases q a <;> simp + + +theorem getEntry?_interSmaller [BEq α] [PartialEquivBEq α] (l₁ l₂ : List ((a : α) × β a)) (k : α) : + List.getEntry? k (interSmaller l₁ l₂) = (List.getEntry? k l₁).filter (fun kv => containsKey kv.1 l₂) := by + rw [interSmaller] + suffices ∀ l₃, + List.getEntry? k (List.foldl (fun sofar kv => List.interSmallerFn l₁ sofar kv.fst) l₃ l₂) = + (Option.filter (fun kv => containsKey kv.fst l₂) (List.getEntry? k l₁)).or (List.getEntry? k l₃) by + simpa using this [] + intro l₃ + induction l₂ using assoc_induction generalizing l₃ with + | nil => simp + | cons k v tl ih => + rw [List.foldl_cons, ih] + simp only [List.containsKey_cons, Bool.or_comm (k == _), Option.filter_or, Option.or_assoc, + List.getEntry?_interSmallerFn] + +theorem List.distinctKeys_interSmaller [BEq α] [PartialEquivBEq α] {l₁ l₂ : List ((a : α) × β a)} : + DistinctKeys (interSmaller l₁ l₂) := by + rw [interSmaller] + suffices ∀ l, DistinctKeys l → DistinctKeys (l₂.foldl (fun sofar kv => List.interSmallerFn l₁ sofar kv.1) l) by + simpa using this [] (by simp) + intro l hl + induction l₂ generalizing l with + | nil => simpa + | cons ht tl ih => + rw [List.foldl_cons] + apply ih + exact distinctKeys_interSmallerFn _ _ _ hl + +theorem List.interSmaller_perm_filter [BEq α] [EquivBEq α] (l₁ l₂ : List ((a : α) × β a)) (h₁ : DistinctKeys l₁) : + List.Perm (List.interSmaller l₁ l₂) (l₁.filter (fun kv => containsKey kv.1 l₂)) := by + apply List.getEntry?_ext + · exact List.distinctKeys_interSmaller + · exact h₁.filter (f := fun k v => containsKey k l₂) + · intro k' + rw [List.getEntry?_filter h₁, List.getEntry?_interSmaller] + +end InterSmaller end Std.Internal.List From 8b7cfa1faae26c55904ba50383d10c7eb1081595 Mon Sep 17 00:00:00 2001 From: Wojciech Rozowski Date: Thu, 13 Nov 2025 19:42:47 +0100 Subject: [PATCH 48/88] Add `eraseMany` and `wf_eraseMany` --- src/Std/Data/DHashMap/Internal/Defs.lean | 10 ++++++++++ src/Std/Data/DHashMap/Internal/WF.lean | 5 +++++ 2 files changed, 15 insertions(+) diff --git a/src/Std/Data/DHashMap/Internal/Defs.lean b/src/Std/Data/DHashMap/Internal/Defs.lean index effe67fb01a6..49b4f8cb3ecc 100644 --- a/src/Std/Data/DHashMap/Internal/Defs.lean +++ b/src/Std/Data/DHashMap/Internal/Defs.lean @@ -451,6 +451,16 @@ def insertMany {ρ : Type w} [ForIn Id ρ ((a : α) × β a)] [BEq α] [Hashable r := ⟨r.1.insert a b, fun _ h hm => h (r.2 _ h hm)⟩ return r +/-- Internal implementation detail of the hash map -/ +def eraseMany {ρ : Type w} [ForIn Id ρ ((a : α) × β a)] [BEq α] [Hashable α] + (m : Raw₀ α β) (l : ρ) : { m' : Raw₀ α β // ∀ (P : Raw₀ α β → Prop), + (∀ {m'' a}, P m'' → P (m''.erase a)) → P m → P m' } := Id.run do + let mut r : { m' : Raw₀ α β // ∀ (P : Raw₀ α β → Prop), + (∀ {m'' a}, P m'' → P (m''.erase a)) → P m → P m' } := ⟨m, fun _ _ => id⟩ + for ⟨a, _⟩ in l do + r := ⟨r.1.erase a, fun _ h hm => h (r.2 _ h hm)⟩ + return r + /-- Internal implementation detail of the hash map -/ @[inline] def insertManyIfNew {ρ : Type w} [ForIn Id ρ ((a : α) × β a)] [BEq α] [Hashable α] (m : Raw₀ α β) (l : ρ) : { m' : Raw₀ α β // ∀ (P : Raw₀ α β → Prop), diff --git a/src/Std/Data/DHashMap/Internal/WF.lean b/src/Std/Data/DHashMap/Internal/WF.lean index 6c4140cd1576..fdcaf46de201 100644 --- a/src/Std/Data/DHashMap/Internal/WF.lean +++ b/src/Std/Data/DHashMap/Internal/WF.lean @@ -1295,6 +1295,11 @@ theorem wf_insertMany₀ [BEq α] [Hashable α] [EquivBEq α] [LawfulHashable α (Raw₀.insertMany ⟨m, h⟩ l).1.1.WF := (Raw₀.insertMany ⟨m, h⟩ l).2 _ Raw.WF.insert₀ h' +theorem wf_eraseMany₀ [BEq α] [Hashable α] [EquivBEq α] [LawfulHashable α] {ρ : Type w} + [ForIn Id ρ ((a : α) × β a)] {m : Raw α β} {h : 0 < m.buckets.size} {l : ρ} (h' : m.WF) : + (Raw₀.eraseMany ⟨m, h⟩ l).1.1.WF := + (Raw₀.eraseMany ⟨m, h⟩ l).2 _ Raw.WF.erase₀ h' + 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) From 3766be5f96f786732d86ad7edc8ad88a175f0c49 Mon Sep 17 00:00:00 2001 From: Wojciech Rozowski Date: Thu, 13 Nov 2025 19:47:40 +0100 Subject: [PATCH 49/88] Add Raw.erase --- src/Std/Data/DHashMap/Internal/WF.lean | 5 +++++ src/Std/Data/DHashMap/Raw.lean | 12 ++++++++++++ 2 files changed, 17 insertions(+) diff --git a/src/Std/Data/DHashMap/Internal/WF.lean b/src/Std/Data/DHashMap/Internal/WF.lean index fdcaf46de201..845880fc1ba2 100644 --- a/src/Std/Data/DHashMap/Internal/WF.lean +++ b/src/Std/Data/DHashMap/Internal/WF.lean @@ -1290,6 +1290,11 @@ 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 WF.eraseMany [BEq α] [Hashable α] {ρ : Type w} [ForIn Id ρ ((a : α) × β a)] {m : Raw α β} + {l : ρ} (h : m.WF) : (m.eraseMany l).WF := by + simpa [Raw.eraseMany, h.size_buckets_pos] using + (Raw₀.eraseMany ⟨m, h.size_buckets_pos⟩ l).2 _ Raw.WF.erase₀ h + theorem wf_insertMany₀ [BEq α] [Hashable α] [EquivBEq α] [LawfulHashable α] {ρ : Type w} [ForIn Id ρ ((a : α) × β a)] {m : Raw α β} {h : 0 < m.buckets.size} {l : ρ} (h' : m.WF) : (Raw₀.insertMany ⟨m, h⟩ l).1.1.WF := diff --git a/src/Std/Data/DHashMap/Raw.lean b/src/Std/Data/DHashMap/Raw.lean index 6c82c30613b1..dbc84edfa187 100644 --- a/src/Std/Data/DHashMap/Raw.lean +++ b/src/Std/Data/DHashMap/Raw.lean @@ -538,6 +538,18 @@ appearance. (Raw₀.insertMany ⟨m, h⟩ l).1 else m -- will never happen for well-formed inputs +/-- +Erases multiple keys from the hash map by iterating over the given collection and calling +`erase` on each key. The values in the collection are ignored; only the keys are used for erasure. +If the same key appears multiple times in the collection, subsequent erasures have no effect after +the first one removes the key. +-/ +@[inline] def eraseMany [BEq α] [Hashable α] {ρ : Type w} [ForIn Id ρ ((a : α) × β a)] + (m : Raw α β) (l : ρ) : Raw α β := + if h : 0 < m.buckets.size then + (Raw₀.eraseMany ⟨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 a9c0b35870bcbe3446c4c3597e91f27e52954216 Mon Sep 17 00:00:00 2001 From: Wojciech Rozowski Date: Thu, 13 Nov 2025 19:57:07 +0100 Subject: [PATCH 50/88] Add further sublemmas --- src/Std/Data/DHashMap/Internal/Model.lean | 20 ++++++++++++++++++++ src/Std/Data/DHashMap/Internal/WF.lean | 12 ++++++++++++ src/Std/Data/Internal/List/Associative.lean | 15 +++++++++++++++ 3 files changed, 47 insertions(+) diff --git a/src/Std/Data/DHashMap/Internal/Model.lean b/src/Std/Data/DHashMap/Internal/Model.lean index c70aaf7c3d47..21ea8e7bc56a 100644 --- a/src/Std/Data/DHashMap/Internal/Model.lean +++ b/src/Std/Data/DHashMap/Internal/Model.lean @@ -407,6 +407,12 @@ def insertListₘ [BEq α] [Hashable α] (m : Raw₀ α β) (l : List ((a : α) | .nil => m | .cons hd tl => insertListₘ (m.insert hd.1 hd.2) tl +/-- Internal implementation detail of the hash map -/ +def eraseListₘ [BEq α] [Hashable α] (m : Raw₀ α β) (l : List ((a : α) × β a)) : Raw₀ α β := + match l with + | .nil => m + | .cons hd tl => eraseListₘ (m.erase hd.1) tl + /-- Internal implementation detail of the hash map -/ def insertListIfNewₘ [BEq α] [Hashable α] (m : Raw₀ α β) (l : List ((a : α) × β a)) : Raw₀ α β := match l with @@ -656,6 +662,20 @@ theorem insertMany_eq_insertListₘ [BEq α] [Hashable α] (m : Raw₀ α β) (l simp only [List.foldl_cons, insertListₘ] apply ih +theorem eraseMany_eq_eraseListₘ [BEq α] [Hashable α] (m : Raw₀ α β) (l : List ((a : α) × β a)) : + eraseMany m l = eraseListₘ m l := by + simp only [eraseMany, Id.run_pure, pure_bind, List.forIn_pure_yield_eq_foldl] + suffices ∀ (t : { m' // ∀ (P : Raw₀ α β → Prop), + (∀ {m'' : Raw₀ α β} {a : α}, P m'' → P (m''.erase a)) → P m → P m' }), + (List.foldl (fun m' p => ⟨m'.val.erase p.1, fun P h₁ h₂ => h₁ (m'.2 _ h₁ h₂)⟩) t l).val = + t.val.eraseListₘ l from this _ + intro t + induction l generalizing m with + | nil => simp [eraseListₘ] + | cons hd tl ih => + simp only [List.foldl_cons, eraseListₘ] + apply ih + theorem insertManyIfNew_eq_insertListIfNewₘ [BEq α] [Hashable α] (m : Raw₀ α β) (l : List ((a : α) × β a)) : insertManyIfNew m l = insertListIfNewₘ m l := by simp only [insertManyIfNew, Id.run_pure, pure_bind, List.forIn_pure_yield_eq_foldl] diff --git a/src/Std/Data/DHashMap/Internal/WF.lean b/src/Std/Data/DHashMap/Internal/WF.lean index 845880fc1ba2..691e97cd695b 100644 --- a/src/Std/Data/DHashMap/Internal/WF.lean +++ b/src/Std/Data/DHashMap/Internal/WF.lean @@ -1082,6 +1082,18 @@ theorem toListModel_insertListₘ [BEq α] [Hashable α] [EquivBEq α] [LawfulHa apply Perm.trans (ih (wfImp_insert h)) apply List.insertList_perm_of_perm_first (toListModel_insert h) (wfImp_insert h).distinct +theorem toListModel_eraseListₘ [BEq α] [Hashable α] [EquivBEq α] [LawfulHashable α] + {m : Raw₀ α β} {l : List ((a : α) × β a)} (h : Raw.WFImp m.1) : + Perm (toListModel (eraseListₘ m l).1.buckets) + (List.eraseList (toListModel m.1.buckets) l) := by + induction l generalizing m with + | nil => + simp [eraseListₘ, List.eraseList] + | cons hd tl ih => + simp only [eraseListₘ, List.eraseList] + apply Perm.trans (ih (wfImp_erase h)) + apply List.eraseList_perm_of_perm_first (toListModel_erase h) (wfImp_erase h).distinct + /-! # `insertListₘ` -/ theorem toListModel_insertListIfNewₘ [BEq α] [Hashable α] [EquivBEq α] [LawfulHashable α] diff --git a/src/Std/Data/Internal/List/Associative.lean b/src/Std/Data/Internal/List/Associative.lean index d16d183287f9..af5349bc7873 100644 --- a/src/Std/Data/Internal/List/Associative.lean +++ b/src/Std/Data/Internal/List/Associative.lean @@ -2815,6 +2815,12 @@ def insertList [BEq α] (l toInsert : List ((a : α) × β a)) : List ((a : α) | .nil => l | .cons ⟨k, v⟩ toInsert => insertList (insertEntry k v l) toInsert +/-- Internal implementation detail of the hash map -/ +def eraseList [BEq α] (l toErase : List ((a : α) × β a)) : List ((a : α) × β a) := + match toErase with + | .nil => l + | .cons ⟨k, _⟩ toErase => eraseList (eraseKey k l) toErase + theorem DistinctKeys.insertList [BEq α] [PartialEquivBEq α] {l₁ l₂ : List ((a : α) × β a)} (h : DistinctKeys l₁) : DistinctKeys (insertList l₁ l₂) := by @@ -2833,6 +2839,15 @@ theorem insertList_perm_of_perm_first [BEq α] [EquivBEq α] {l1 l2 toInsert : L simp only [insertList] apply ih (insertEntry_of_perm distinct h) (DistinctKeys.insertEntry distinct) +theorem eraseList_perm_of_perm_first [BEq α] [EquivBEq α] {l1 l2 toErase : List ((a : α) × β a)} + (h : Perm l1 l2) (distinct : DistinctKeys l1) : + Perm (eraseList l1 toErase) (eraseList l2 toErase) := by + induction toErase generalizing l1 l2 with + | nil => simp [eraseList, h] + | cons hd tl ih => + simp only [eraseList] + apply ih (eraseKey_of_perm distinct h) (DistinctKeys.eraseKey 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 From 922beb52a503bab2dc44a0980fcda4fa6e40d04f Mon Sep 17 00:00:00 2001 From: Wojciech Rozowski Date: Fri, 14 Nov 2025 10:53:28 +0100 Subject: [PATCH 51/88] add new correspondence lemma --- src/Std/Data/DHashMap/Internal/WF.lean | 33 ++++++-- src/Std/Data/Internal/List/Associative.lean | 93 +++++++++++++++++---- 2 files changed, 104 insertions(+), 22 deletions(-) diff --git a/src/Std/Data/DHashMap/Internal/WF.lean b/src/Std/Data/DHashMap/Internal/WF.lean index 691e97cd695b..1d79f830f471 100644 --- a/src/Std/Data/DHashMap/Internal/WF.lean +++ b/src/Std/Data/DHashMap/Internal/WF.lean @@ -1302,6 +1302,16 @@ 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_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 + +/-! # `eraseMany` -/ + theorem WF.eraseMany [BEq α] [Hashable α] {ρ : Type w} [ForIn Id ρ ((a : α) × β a)] {m : Raw α β} {l : ρ} (h : m.WF) : (m.eraseMany l).WF := by simpa [Raw.eraseMany, h.size_buckets_pos] using @@ -1317,13 +1327,22 @@ theorem wf_eraseMany₀ [BEq α] [Hashable α] [EquivBEq α] [LawfulHashable α] (Raw₀.eraseMany ⟨m, h⟩ l).1.1.WF := (Raw₀.eraseMany ⟨m, h⟩ l).2 _ Raw.WF.erase₀ h' -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 +theorem eraseMany_eq_eraseListₘ_toListModel [BEq α] [Hashable α] (m m₂ : Raw₀ α β) : + eraseMany m m₂.1 = eraseListₘ m (toListModel m₂.1.buckets) := by + simp only [eraseMany, bind_pure_comp, map_pure, bind_pure] + simp only [ForIn.forIn] + simp only [Raw.forIn_eq_forIn_toListModel, forIn_pure_yield_eq_foldl, Id.run_pure] + generalize toListModel m₂.val.buckets = l + suffices ∀ (t : { m' // ∀ (P : Raw₀ α β → Prop), + (∀ {m'' : Raw₀ α β} {a : α}, P m'' → P (m''.erase a)) → P m → P m' }), + (List.foldl (fun m' p => ⟨m'.val.erase p.1, fun P h₁ h₂ => h₁ (m'.2 _ h₁ h₂)⟩) t l).val = + t.val.eraseListₘ l from this _ + intro t + induction l generalizing m with + | nil => simp [eraseListₘ] + | cons hd tl ih => + simp only [List.foldl_cons, eraseListₘ] + apply ih /-! # `insertManyIfNew` -/ diff --git a/src/Std/Data/Internal/List/Associative.lean b/src/Std/Data/Internal/List/Associative.lean index af5349bc7873..156f3941d6ea 100644 --- a/src/Std/Data/Internal/List/Associative.lean +++ b/src/Std/Data/Internal/List/Associative.lean @@ -2815,12 +2815,6 @@ def insertList [BEq α] (l toInsert : List ((a : α) × β a)) : List ((a : α) | .nil => l | .cons ⟨k, v⟩ toInsert => insertList (insertEntry k v l) toInsert -/-- Internal implementation detail of the hash map -/ -def eraseList [BEq α] (l toErase : List ((a : α) × β a)) : List ((a : α) × β a) := - match toErase with - | .nil => l - | .cons ⟨k, _⟩ toErase => eraseList (eraseKey k l) toErase - theorem DistinctKeys.insertList [BEq α] [PartialEquivBEq α] {l₁ l₂ : List ((a : α) × β a)} (h : DistinctKeys l₁) : DistinctKeys (insertList l₁ l₂) := by @@ -2839,15 +2833,6 @@ theorem insertList_perm_of_perm_first [BEq α] [EquivBEq α] {l1 l2 toInsert : L simp only [insertList] apply ih (insertEntry_of_perm distinct h) (DistinctKeys.insertEntry distinct) -theorem eraseList_perm_of_perm_first [BEq α] [EquivBEq α] {l1 l2 toErase : List ((a : α) × β a)} - (h : Perm l1 l2) (distinct : DistinctKeys l1) : - Perm (eraseList l1 toErase) (eraseList l2 toErase) := by - induction toErase generalizing l1 l2 with - | nil => simp [eraseList, h] - | cons hd tl ih => - simp only [eraseList] - apply ih (eraseKey_of_perm distinct h) (DistinctKeys.eraseKey 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 @@ -5292,6 +5277,84 @@ theorem List.getKey?_filter_containsKey_of_containsKey_right [BEq α] [EquivBEq · simp only [↓reduceIte] · exact dl₁ + +/-- Internal implementation detail of the hash map -/ +def eraseList [BEq α] (l toErase : List ((a : α) × β a)) : List ((a : α) × β a) := + match toErase with + | .nil => l + | .cons ⟨k, _⟩ toErase => eraseList (eraseKey k l) toErase + +theorem eraseList_perm_filter_containsKey [BEq α] [EquivBEq α] (l toErase : List ((a : α) × β a)) (hl : DistinctKeys l): + List.Perm (eraseList l toErase) (l.filter (fun p => !containsKey p.fst toErase)) := by + induction toErase generalizing l with + | nil => + simp [eraseList, containsKey] + clear hl + induction l + case nil => simp + case cons h t t_ih => + simp only [List.filter_cons_of_pos, List.perm_cons] + apply t_ih + | cons kv tl ih => + simp only [eraseList] + apply Perm.trans + · apply ih + · apply DistinctKeys.eraseKey hl + · apply getEntry?_ext + · apply @DistinctKeys.filter α β _ (eraseKey kv.fst l) (fun p _ => !containsKey p tl) + · apply DistinctKeys.eraseKey hl + · apply @DistinctKeys.filter α β _ l (fun p _ => !containsKey p (kv :: tl)) hl + · intro a + rw [getEntry?_filter, getEntry?_filter] + rw [getEntry?_eraseKey] + split + · rename_i heq + conv => + rhs + lhs + ext p + congr + rw [containsKey_cons] + simp only [Option.filter_none, Bool.not_or] + generalize heq2 : getEntry? a l = x + cases x + · simp + · rename_i val + rw [getEntry?_eq_some_iff] at heq2 + · simp [Option.filter] + intro hyp + rw [PartialEquivBEq.trans heq heq2.1] at hyp + contradiction + · exact hl + · rename_i heq + simp only [Bool.not_eq_true] at heq + generalize heq2 : getEntry? a l = x + cases x + · simp + · rename_i val + simp only [Option.filter, Bool.not_eq_eq_eq_not, Bool.not_true] + rw [containsKey_cons] + rw [getEntry?_eq_some_iff] at heq2 + · suffices (kv.fst == val.fst) = false by simp [this] + apply Classical.byContradiction + intro hyp + simp only [Bool.not_eq_false] at hyp + rw [BEq.trans hyp (BEq.symm heq2.1)] at heq + contradiction + · exact hl + · exact hl + · exact hl + · apply DistinctKeys.eraseKey hl + +theorem eraseList_perm_of_perm_first [BEq α] [EquivBEq α] {l1 l2 toErase : List ((a : α) × β a)} + (h : Perm l1 l2) (distinct : DistinctKeys l1) : + Perm (eraseList l1 toErase) (eraseList l2 toErase) := by + induction toErase generalizing l1 l2 with + | nil => simp [eraseList, h] + | cons hd tl ih => + simp only [eraseList] + apply ih (eraseKey_of_perm distinct h) (DistinctKeys.eraseKey distinct) + theorem List.getKey?_filter_containsKey_of_containsKey_eq_false_left [BEq α] [EquivBEq α] {l₁ l₂ : List ((a : α) × β a)} {k : α} (dl₁ : DistinctKeys l₁) : From b652b7b66aad48892037a11e55894783dd2be198 Mon Sep 17 00:00:00 2001 From: Wojciech Rozowski Date: Fri, 14 Nov 2025 11:31:05 +0100 Subject: [PATCH 52/88] add to list model --- src/Std/Data/DHashMap/Internal/Defs.lean | 5 ++++ src/Std/Data/DHashMap/Internal/Model.lean | 7 +++++ src/Std/Data/DHashMap/Internal/WF.lean | 35 +++++++++++++++++++++++ 3 files changed, 47 insertions(+) diff --git a/src/Std/Data/DHashMap/Internal/Defs.lean b/src/Std/Data/DHashMap/Internal/Defs.lean index 49b4f8cb3ecc..b027ec0e6407 100644 --- a/src/Std/Data/DHashMap/Internal/Defs.lean +++ b/src/Std/Data/DHashMap/Internal/Defs.lean @@ -490,6 +490,11 @@ def interSmaller [BEq α] [Hashable α] (m₁ : Raw₀ α β) (m₂ : Raw α β) def inter [BEq α] [Hashable α] (m₁ m₂ : Raw₀ α β) : Raw₀ α β := if m₁.1.size ≤ m₂.1.size then m₁.filter fun k _ => m₂.contains k else interSmaller m₁ m₂ +/-- Internal implementation detail of the hash map -/ +@[inline] def diff [BEq α] [Hashable α] (m₁ m₂ : Raw₀ α β) : Raw₀ α β := + if m₁.1.size ≤ m₂.1.size then m₁.filter (fun k _ => !m₂.contains k) else (eraseMany m₁ m₂.1).1 + + section variable {β : Type v} diff --git a/src/Std/Data/DHashMap/Internal/Model.lean b/src/Std/Data/DHashMap/Internal/Model.lean index 21ea8e7bc56a..d4ecf9175410 100644 --- a/src/Std/Data/DHashMap/Internal/Model.lean +++ b/src/Std/Data/DHashMap/Internal/Model.lean @@ -413,6 +413,13 @@ def eraseListₘ [BEq α] [Hashable α] (m : Raw₀ α β) (l : List ((a : α) | .nil => m | .cons hd tl => eraseListₘ (m.erase hd.1) tl +/-- Internal implementation detail of the hash map -/ +def diffₘ [BEq α] [Hashable α] (m₁ m₂ : Raw₀ α β) : Raw₀ α β := + if m₁.1.size ≤ m₂.1.size then + filterₘ m₁ (fun k _ => !containsₘ m₂ k) + else + eraseMany m₁ (toListModel m₂.1.buckets) + /-- Internal implementation detail of the hash map -/ def insertListIfNewₘ [BEq α] [Hashable α] (m : Raw₀ α β) (l : List ((a : α) × β a)) : Raw₀ α β := match l with diff --git a/src/Std/Data/DHashMap/Internal/WF.lean b/src/Std/Data/DHashMap/Internal/WF.lean index 1d79f830f471..5a15b550fe74 100644 --- a/src/Std/Data/DHashMap/Internal/WF.lean +++ b/src/Std/Data/DHashMap/Internal/WF.lean @@ -1344,6 +1344,41 @@ theorem eraseMany_eq_eraseListₘ_toListModel [BEq α] [Hashable α] (m m₂ : R simp only [List.foldl_cons, eraseListₘ] apply ih +theorem toListModel_diffₘ [BEq α] [Hashable α] [EquivBEq α] [LawfulHashable α] + {m₁ m₂ : Raw₀ α β} (h₁ : Raw.WFImp m₁.1) (h₂ : Raw.WFImp m₂.1) : + Perm (toListModel (diffₘ m₁ m₂).1.buckets) + (List.filter (fun k => !containsKey k.fst (toListModel m₂.1.buckets)) (toListModel m₁.1.buckets)) := by + rw [diffₘ] + split + · apply Perm.trans + · apply toListModel_filterₘ + · conv => + lhs + lhs + ext p + congr + rw [containsₘ_eq_containsKey h₂] + · rw [eraseMany_eq_eraseListₘ] + apply Perm.trans (toListModel_eraseListₘ h₁) + · apply eraseList_perm_filter_containsKey + · exact h₁.distinct + +theorem diff_eq_diffₘ [BEq α] [Hashable α] (m₁ m₂ : Raw₀ α β) : + diff m₁ m₂ = diffₘ m₁ m₂ := by + rw [diff, diffₘ] + split + · rw [filter_eq_filterₘ] + congr + · rw [eraseMany_eq_eraseListₘ_toListModel] + rw [eraseMany_eq_eraseListₘ] + +theorem toListModel_diff [BEq α] [Hashable α] [EquivBEq α] [LawfulHashable α] {m₁ m₂ : Raw₀ α β} + (h₁ : Raw.WFImp m₁.1) (h₂ : Raw.WFImp m₂.1) : + Perm (toListModel (m₁.diff m₂).1.buckets) + (List.filter (fun k => !containsKey k.fst (toListModel m₂.1.buckets)) (toListModel m₁.1.buckets)) := by + rw [diff_eq_diffₘ] + exact toListModel_diffₘ h₁ h₂ + /-! # `insertManyIfNew` -/ theorem wfImp_insertManyIfNew [BEq α] [Hashable α] [EquivBEq α] [LawfulHashable α] {ρ : Type w} From ed7b23ac9f907767e37feca47cf6b3e6755dedf0 Mon Sep 17 00:00:00 2001 From: Wojciech Rozowski Date: Fri, 14 Nov 2025 11:47:13 +0100 Subject: [PATCH 53/88] add lemmas --- src/Std/Data/DHashMap/Internal/RawLemmas.lean | 245 +++++++++++++++++- src/Std/Data/DHashMap/Internal/WF.lean | 9 + 2 files changed, 253 insertions(+), 1 deletion(-) diff --git a/src/Std/Data/DHashMap/Internal/RawLemmas.lean b/src/Std/Data/DHashMap/Internal/RawLemmas.lean index 547848631cda..5570bb512182 100644 --- a/src/Std/Data/DHashMap/Internal/RawLemmas.lean +++ b/src/Std/Data/DHashMap/Internal/RawLemmas.lean @@ -91,7 +91,7 @@ macro_rules | apply Raw₀.wf_insertMany₀ | apply Raw₀.Const.wf_insertMany₀ | apply Raw₀.Const.wf_insertManyIfNewUnit₀ | apply Raw₀.wf_union₀ | apply Raw.WF.filter₀ | apply Raw₀.wf_map₀ | apply Raw₀.wf_filterMap₀ - | apply Raw.WF.emptyWithCapacity₀ | apply Raw₀.wf_inter₀) <;> wf_trivial) + | apply Raw.WF.emptyWithCapacity₀ | apply Raw₀.wf_inter₀ | apply Raw₀.wf_diff₀) <;> wf_trivial) /-- Internal implementation detail of the hash map -/ scoped macro "empty" : tactic => `(tactic| { intros; simp_all [List.isEmpty_iff] } ) @@ -106,6 +106,7 @@ private meta def modifyMap : Std.DHashMap Name (fun _ => Name) := ⟨`insertMany, ``toListModel_insertMany_list⟩, ⟨`union, ``toListModel_union⟩, ⟨`inter, ``toListModel_inter⟩, + ⟨`diff, ``toListModel_diff⟩, ⟨`Const.insertMany, ``Const.toListModel_insertMany_list⟩, ⟨`Const.insertManyIfNewUnit, ``Const.toListModel_insertManyIfNewUnit_list⟩, ⟨`alter, ``toListModel_alter⟩, @@ -3091,6 +3092,248 @@ theorem get!_inter_of_contains_eq_false_left [EquivBEq α] [LawfulHashable α] [ end Const +section Diff + +variable {m₁ m₂ : Raw₀ α β} + +/- contains -/ +theorem contains_diff [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) + (h₂ : m₂.val.WF) {k : α} : + (m₁.diff m₂).contains k = (m₁.contains k && !m₂.contains k) := by + sorry + +theorem contains_diff_iff [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) + (h₂ : m₂.val.WF) {k : α} : + (m₁.diff m₂).contains k ↔ m₁.contains k ∧ !m₂.contains k = false := by + sorry + +theorem contains_diff_eq_false_of_contains_eq_false_left [EquivBEq α] [LawfulHashable α] + (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} + (h : m₁.contains k = false) : + (m₁.diff m₂).contains k = false := by + sorry + +theorem contains_diff_eq_false_of_contains_right [EquivBEq α] [LawfulHashable α] + (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} + (h : m₂.contains k) : + (m₁.diff m₂).contains k = false := by + sorry + +/- Equiv -/ +theorem diff_erase_right_equiv_erase_diff [EquivBEq α] [LawfulHashable α] {k : α} + (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) : + (m₁.diff (m₂.erase k)).1.Equiv ((m₁.diff m₂).erase k).1 := by + sorry + +/- get? -/ +theorem get?_diff [LawfulBEq α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} : + (m₁.diff m₂).get? k = if m₂.contains k then none else m₁.get? k := by + sorry + +theorem get?_diff_of_contains_left [LawfulBEq α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) + {k : α} (h : m₁.contains k) (h' : m₂.contains k = false) : + (m₁.diff m₂).get? k = m₁.get? k := by sorry + +theorem get?_diff_of_contains_eq_false_left [LawfulBEq α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) + {k : α} (h : m₁.contains k = false) : + (m₁.diff m₂).get? k = none := by sorry + +theorem get?_diff_of_contains_right [LawfulBEq α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) + {k : α} (h : m₂.contains k) : + (m₁.diff m₂).get? k = none := by sorry + +/- get -/ +theorem get_diff [LawfulBEq α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) + {k : α} {h_contains : (m₁.diff m₂).contains k} : + (m₁.diff m₂).get k h_contains = + m₁.get k ((contains_diff_iff h₁ h₂).1 h_contains).1 := by sorry + +/- getD -/ +theorem getD_diff [LawfulBEq α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) + {k : α} {fallback : β k} : + (m₁.diff m₂).getD k fallback = + if m₂.contains k then fallback else m₁.getD k fallback := by sorry + +theorem getD_diff_of_contains_left [LawfulBEq α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) + {k : α} {fallback : β k} (h : m₁.contains k) (h' : m₂.contains k = false) : + (m₁.diff m₂).getD k fallback = m₁.getD k fallback := by sorry + +theorem getD_diff_of_contains_right [LawfulBEq α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) + {k : α} {fallback : β k} (h : m₂.contains k) : + (m₁.diff m₂).getD k fallback = fallback := by sorry + +theorem getD_diff_of_contains_eq_false_left [LawfulBEq α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) + {k : α} {fallback : β k} (h : m₁.contains k = false) : + (m₁.diff m₂).getD k fallback = fallback := by sorry + +/- get! -/ +theorem get!_diff [LawfulBEq α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) + {k : α} [Inhabited (β k)] : + (m₁.diff m₂).get! k = if m₂.contains k then default else m₁.get! k := by sorry + +theorem get!_diff_of_contains_left [LawfulBEq α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) + {k : α} [Inhabited (β k)] (h : m₁.contains k) (h' : m₂.contains k = false) : + (m₁.diff m₂).get! k = m₁.get! k := by sorry + +theorem get!_diff_of_contains_right [LawfulBEq α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) + {k : α} [Inhabited (β k)] (h : m₂.contains k) : + (m₁.diff m₂).get! k = default := by sorry + +theorem get!_diff_of_contains_eq_false_left [LawfulBEq α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) + {k : α} [Inhabited (β k)] (h : m₁.contains k = false) : + (m₁.diff m₂).get! k = default := by sorry + +/- getKey? -/ +theorem getKey?_diff [EquivBEq α] [LawfulHashable α] + (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} : + (m₁.diff m₂).getKey? k = if m₂.contains k then none else m₁.getKey? k := by sorry + +theorem getKey?_diff_of_contains_left [EquivBEq α] [LawfulHashable α] + (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} + (h : m₁.contains k) (h' : m₂.contains k = false) : + (m₁.diff m₂).getKey? k = m₁.getKey? k := by sorry + +theorem getKey?_diff_of_contains_eq_false_left [EquivBEq α] [LawfulHashable α] + (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} (h : m₁.contains k = false) : + (m₁.diff m₂).getKey? k = none := by sorry + +theorem getKey?_diff_of_contains_right [EquivBEq α] [LawfulHashable α] + (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} (h : m₂.contains k) : + (m₁.diff m₂).getKey? k = none := by sorry + +/- getKey -/ +theorem getKey_diff [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) + {k : α} {h_contains : (m₁.diff m₂).contains k} : + (m₁.diff m₂).getKey k h_contains = + m₁.getKey k ((contains_diff_iff h₁ h₂).1 h_contains).1 := by sorry + +/- getKeyD -/ +theorem getKeyD_diff [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) + (h₂ : m₂.val.WF) {k fallback : α} : + (m₁.diff m₂).getKeyD k fallback = + if m₂.contains k then fallback else m₁.getKeyD k fallback := by sorry + +theorem getKeyD_diff_of_contains_left [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) + (h₂ : m₂.val.WF) {k fallback : α} + (h : m₁.contains k) (h' : m₂.contains k = false) : + (m₁.diff m₂).getKeyD k fallback = m₁.getKeyD k fallback := by sorry + +theorem getKeyD_diff_of_contains_right [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) + (h₂ : m₂.val.WF) {k fallback : α} (h : m₂.contains k) : + (m₁.diff m₂).getKeyD k fallback = fallback := by sorry + +theorem getKeyD_diff_of_contains_eq_false_left [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) + (h₂ : m₂.val.WF) {k fallback : α} (h : m₁.contains k = false) : + (m₁.diff m₂).getKeyD k fallback = fallback := by sorry + +/- getKey! -/ +theorem getKey!_diff [EquivBEq α] [LawfulHashable α] [Inhabited α] (h₁ : m₁.val.WF) + (h₂ : m₂.val.WF) {k : α} : + (m₁.diff m₂).getKey! k = if m₂.contains k then default else m₁.getKey! k := by sorry + +theorem getKey!_diff_of_contains_left [EquivBEq α] [LawfulHashable α] [Inhabited α] (h₁ : m₁.val.WF) + (h₂ : m₂.val.WF) {k : α} (h : m₁.contains k) (h' : m₂.contains k = false) : + (m₁.diff m₂).getKey! k = m₁.getKey! k := by sorry + +theorem getKey!_diff_of_contains_right [EquivBEq α] [LawfulHashable α] [Inhabited α] (h₁ : m₁.val.WF) + (h₂ : m₂.val.WF) {k : α} (h : m₂.contains k) : + (m₁.diff m₂).getKey! k = default := by sorry + +theorem getKey!_diff_of_contains_eq_false_left [EquivBEq α] [LawfulHashable α] [Inhabited α] (h₁ : m₁.val.WF) + (h₂ : m₂.val.WF) {k : α} (h : m₁.contains k = false) : + (m₁.diff m₂).getKey! k = default := by sorry + +/- size -/ +theorem size_diff_le_size_left [EquivBEq α] [LawfulHashable α] + (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) : + (m₁.diff m₂).1.size ≤ m₁.1.size := by sorry + +theorem size_diff_eq_size_left_of_disjoint [EquivBEq α] [LawfulHashable α] + (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) + (h : ∀ (a : α), m₁.contains a → m₂.contains a = false) : + (m₁.diff m₂).1.size = m₁.1.size := by sorry + +theorem size_diff_add_size_inter_eq_size_left [EquivBEq α] [LawfulHashable α] + (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) : + (m₁.diff m₂).1.size + (m₁.inter m₂).1.size = m₁.1.size := by sorry + +/- isEmpty -/ +@[simp] +theorem isEmpty_diff_of_isEmpty_left [EquivBEq α] [LawfulHashable α] + (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) (h : m₁.1.isEmpty) : + (m₁.diff m₂).1.isEmpty = true := by sorry + +theorem isEmpty_diff_iff [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) : + (m₁.diff m₂).1.isEmpty ↔ ∀ k, m₁.contains k → m₂.contains k := by sorry + +theorem isEmpty_diff_of_subset [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) + (h : ∀ k, m₁.contains k → m₂.contains k) : + (m₁.diff m₂).1.isEmpty = true := by sorry + +end Diff + +namespace Const + +variable {β : Type v} {m₁ m₂ : Raw₀ α (fun _ => β)} + +/- get? -/ +theorem get?_diff [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} : + Const.get? (m₁.diff m₂) k = if m₂.contains k then none else Const.get? m₁ k := by sorry + +theorem get?_diff_of_contains_left [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) + {k : α} (h : m₁.contains k) (h' : m₂.contains k = false) : + Const.get? (m₁.diff m₂) k = Const.get? m₁ k := by sorry + +theorem get?_diff_of_contains_eq_false_left [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) + {k : α} (h : m₁.contains k = false) : + Const.get? (m₁.diff m₂) k = none := by sorry + +theorem get?_diff_of_contains_right [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) + {k : α} (h : m₂.contains k) : + Const.get? (m₁.diff m₂) k = none := by sorry + +/- get -/ +theorem get_diff [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) + {k : α} {h_contains : (m₁.diff m₂).contains k} : + Const.get (m₁.diff m₂) k h_contains = + Const.get m₁ k ((contains_diff_iff h₁ h₂).1 h_contains).1 := by sorry + +/- getD -/ +theorem getD_diff [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) + {k : α} {fallback : β} : + Const.getD (m₁.diff m₂) k fallback = + if m₂.contains k then fallback else Const.getD m₁ k fallback := by sorry + +theorem getD_diff_of_contains_left [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) + {k : α} {fallback : β} (h : m₁.contains k) (h' : m₂.contains k = false) : + Const.getD (m₁.diff m₂) k fallback = Const.getD m₁ k fallback := by sorry + +theorem getD_diff_of_contains_right [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) + {k : α} {fallback : β} (h : m₂.contains k) : + Const.getD (m₁.diff m₂) k fallback = fallback := by sorry + +theorem getD_diff_of_contains_eq_false_left [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) + {k : α} {fallback : β} (h : m₁.contains k = false) : + Const.getD (m₁.diff m₂) k fallback = fallback := by sorry + +/- get! -/ +theorem get!_diff [EquivBEq α] [LawfulHashable α] [Inhabited β] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} : + Const.get! (m₁.diff m₂) k = if m₂.contains k then default else Const.get! m₁ k := by sorry + +theorem get!_diff_of_contains_left [EquivBEq α] [LawfulHashable α] [Inhabited β] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) + {k : α} (h : m₁.contains k) (h' : m₂.contains k = false) : + Const.get! (m₁.diff m₂) k = Const.get! m₁ k := by sorry + +theorem get!_diff_of_contains_right [EquivBEq α] [LawfulHashable α] [Inhabited β] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) + {k : α} (h : m₂.contains k) : + Const.get! (m₁.diff m₂) k = default := by sorry + +theorem get!_diff_of_contains_eq_false_left [EquivBEq α] [LawfulHashable α] [Inhabited β] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) + {k : α} (h : m₁.contains k = false) : + Const.get! (m₁.diff m₂) k = default := by sorry + +end Const + section Alter theorem isEmpty_alter_eq_isEmpty_erase [LawfulBEq α] (h : m.1.WF) {k : α} diff --git a/src/Std/Data/DHashMap/Internal/WF.lean b/src/Std/Data/DHashMap/Internal/WF.lean index 5a15b550fe74..336e21c29cb7 100644 --- a/src/Std/Data/DHashMap/Internal/WF.lean +++ b/src/Std/Data/DHashMap/Internal/WF.lean @@ -1379,6 +1379,15 @@ theorem toListModel_diff [BEq α] [Hashable α] [EquivBEq α] [LawfulHashable α rw [diff_eq_diffₘ] exact toListModel_diffₘ h₁ h₂ +theorem wf_diff₀ [BEq α] [Hashable α] [EquivBEq α] [LawfulHashable α] + {m₁ m₂ : Raw α β} {h₁ : 0 < m₁.buckets.size} {h₂ : 0 < m₂.buckets.size} (h'₁ : m₁.WF) : + (Raw₀.diff ⟨m₁, h₁⟩ ⟨m₂, h₂⟩).1.WF := by + rw [diff] + split + · apply Raw.WF.filter₀ h'₁ + · exact wf_eraseMany₀ ‹_› + + /-! # `insertManyIfNew` -/ theorem wfImp_insertManyIfNew [BEq α] [Hashable α] [EquivBEq α] [LawfulHashable α] {ρ : Type w} From 7f4590368e2b2f0d1e649ae3b0e422cee4d5295a Mon Sep 17 00:00:00 2001 From: Wojciech Rozowski Date: Fri, 14 Nov 2025 12:07:40 +0100 Subject: [PATCH 54/88] chore: add sublemmas --- src/Std/Data/DHashMap/Internal/RawLemmas.lean | 6 ++--- src/Std/Data/Internal/List/Associative.lean | 22 +++++++++++++++++++ 2 files changed, 25 insertions(+), 3 deletions(-) diff --git a/src/Std/Data/DHashMap/Internal/RawLemmas.lean b/src/Std/Data/DHashMap/Internal/RawLemmas.lean index 5570bb512182..7129b5852b54 100644 --- a/src/Std/Data/DHashMap/Internal/RawLemmas.lean +++ b/src/Std/Data/DHashMap/Internal/RawLemmas.lean @@ -3100,12 +3100,12 @@ variable {m₁ m₂ : Raw₀ α β} theorem contains_diff [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} : (m₁.diff m₂).contains k = (m₁.contains k && !m₂.contains k) := by - sorry + simp_to_model [diff, contains] using List.containsKey_filter_containsKey_eq_false theorem contains_diff_iff [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} : - (m₁.diff m₂).contains k ↔ m₁.contains k ∧ !m₂.contains k = false := by - sorry + (m₁.diff m₂).contains k ↔ m₁.contains k ∧ ¬m₂.contains k := by + simp_to_model [diff, contains] using List.containsKey_filter_containsKey_eq_false_iff theorem contains_diff_eq_false_of_contains_eq_false_left [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} diff --git a/src/Std/Data/Internal/List/Associative.lean b/src/Std/Data/Internal/List/Associative.lean index 156f3941d6ea..326f126c92fa 100644 --- a/src/Std/Data/Internal/List/Associative.lean +++ b/src/Std/Data/Internal/List/Associative.lean @@ -5507,6 +5507,28 @@ theorem List.containsKey_filter_containsKey_iff [BEq α] [EquivBEq α] {l₁ l . simp . exact hl₁ +theorem containsKey_filter_containsKey_eq_false [BEq α] [EquivBEq α] {l₁ l₂ : List ((a : α) × β a)} {hl₁ : DistinctKeys l₁} {k : α} : + (containsKey k (List.filter (fun p => !containsKey p.fst (l₂)) l₁)) = + (containsKey k l₁ && !containsKey k l₂) := by + rw [containsKey_eq_isSome_getEntry?, containsKey_eq_isSome_getEntry?] + rw [getEntry?_filter] + generalize heq : (getEntry? k l₁) = x + cases x + case none => + simp + case some kv => + simp only [Option.isSome_filter, Option.any_some, Option.isSome_some, Bool.true_and] + rw [containsKey_congr] + apply List.beq_of_getEntry?_eq_some heq + . exact hl₁ + +theorem containsKey_filter_containsKey_eq_false_iff [BEq α] [EquivBEq α] {l₁ l₂ : List ((a : α) × β a)} {hl₁ : DistinctKeys l₁} {k : α} : + (containsKey k (List.filter (fun p => !containsKey p.fst l₂) l₁)) ↔ + (containsKey k l₁ ∧ ¬ containsKey k l₂) := by + rw [containsKey_filter_containsKey_eq_false] + . simp + . exact hl₁ + theorem List.getValueCast_filter_containsKey [BEq α] [LawfulBEq α] {l₁ l₂ : List ((a : α) × β a)} {k : α} (dl₁ : DistinctKeys l₁) {h₁ h₂} : From b1435c13ade109c7a19bfb7ba49cfa23207052c6 Mon Sep 17 00:00:00 2001 From: Wojciech Rozowski Date: Fri, 14 Nov 2025 12:54:23 +0100 Subject: [PATCH 55/88] further progress on lemmas --- src/Std/Data/DHashMap/Internal/RawLemmas.lean | 29 ++++--- src/Std/Data/Internal/List/Associative.lean | 83 ++++++++++++++++++- 2 files changed, 99 insertions(+), 13 deletions(-) diff --git a/src/Std/Data/DHashMap/Internal/RawLemmas.lean b/src/Std/Data/DHashMap/Internal/RawLemmas.lean index 7129b5852b54..0dab3d9ce0c9 100644 --- a/src/Std/Data/DHashMap/Internal/RawLemmas.lean +++ b/src/Std/Data/DHashMap/Internal/RawLemmas.lean @@ -3100,24 +3100,26 @@ variable {m₁ m₂ : Raw₀ α β} theorem contains_diff [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} : (m₁.diff m₂).contains k = (m₁.contains k && !m₂.contains k) := by - simp_to_model [diff, contains] using List.containsKey_filter_containsKey_eq_false + simp_to_model [diff, contains] using List.containsKey_diff_eq_false theorem contains_diff_iff [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} : (m₁.diff m₂).contains k ↔ m₁.contains k ∧ ¬m₂.contains k := by - simp_to_model [diff, contains] using List.containsKey_filter_containsKey_eq_false_iff + simp_to_model [diff, contains] using List.containsKey_diff_eq_false_iff theorem contains_diff_eq_false_of_contains_eq_false_left [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} (h : m₁.contains k = false) : (m₁.diff m₂).contains k = false := by - sorry + revert h + simp_to_model [diff, contains] using List.containsKey_diff_eq_false_of_containsKey_eq_false_left theorem contains_diff_eq_false_of_contains_right [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} (h : m₂.contains k) : (m₁.diff m₂).contains k = false := by - sorry + revert h + simp_to_model [diff, contains] using List.containsKey_diff_of_containsKey_eq_false_right /- Equiv -/ theorem diff_erase_right_equiv_erase_diff [EquivBEq α] [LawfulHashable α] {k : α} @@ -3128,19 +3130,26 @@ theorem diff_erase_right_equiv_erase_diff [EquivBEq α] [LawfulHashable α] {k : /- get? -/ theorem get?_diff [LawfulBEq α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} : (m₁.diff m₂).get? k = if m₂.contains k then none else m₁.get? k := by - sorry + simp_to_model [diff, get?, contains] using List.getValueCast?_diff -theorem get?_diff_of_contains_left [LawfulBEq α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) - {k : α} (h : m₁.contains k) (h' : m₂.contains k = false) : - (m₁.diff m₂).get? k = m₁.get? k := by sorry +theorem get?_diff_of_contains_eq_false_right [LawfulBEq α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) + {k : α} (h : m₂.contains k = false) : + (m₁.diff m₂).get? k = m₁.get? k := by + revert h + simp_to_model [diff, contains, get?] using List.getValueCast?_diff_of_containsKey_eq_false_right theorem get?_diff_of_contains_eq_false_left [LawfulBEq α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} (h : m₁.contains k = false) : - (m₁.diff m₂).get? k = none := by sorry + (m₁.diff m₂).get? k = none := by + revert h + simp_to_model [diff, contains, get?] using List.getValueCast?_diff_of_containsKey_eq_false_left theorem get?_diff_of_contains_right [LawfulBEq α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} (h : m₂.contains k) : - (m₁.diff m₂).get? k = none := by sorry + (m₁.diff m₂).get? k = none := by + revert h + simp_to_model [diff, get?, contains] using List.getValueCast?_diff_of_containsKey_right + /- get -/ theorem get_diff [LawfulBEq α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) diff --git a/src/Std/Data/Internal/List/Associative.lean b/src/Std/Data/Internal/List/Associative.lean index 326f126c92fa..bf59714a74a8 100644 --- a/src/Std/Data/Internal/List/Associative.lean +++ b/src/Std/Data/Internal/List/Associative.lean @@ -5233,6 +5233,68 @@ theorem List.getValueCast?_filter_containsKey [BEq α] [LawfulBEq α] case h_2 => rfl exact dl₁ +theorem List.getValueCast?_diff [BEq α] [LawfulBEq α] + {l₁ l₂ : List ((a : α) × β a)} {k : α} + (dl₁ : DistinctKeys l₁) : + getValueCast? k (List.filter (fun p => !containsKey p.fst l₂) l₁) = + if containsKey k l₂ = true then none else getValueCast? k l₁ := by + simp only [getValueCast?_eq_getEntry?] + split + case isTrue h => + suffices (getEntry? k (List.filter (fun p => !containsKey p.fst l₂) l₁)) = none by simp [this] + rw [getEntry?_filter] + simp [Option.filter] + split + case h_1 _ val heq => + rw [@containsKey_congr _ _ _ _ l₂ _ _ (PartialEquivBEq.symm <| @beq_of_getEntry?_eq_some α β _ l₁ k val heq)] at h + simp [h] + case h_2 => rfl + exact dl₁ + case isFalse h => + apply Option.dmap_congr + . simp + . rw [getEntry?_filter] + . simp [Option.filter] + split + case h_1 _ val heq => + rw [@containsKey_congr _ _ _ _ l₂ _ _ (PartialEquivBEq.symm <| @beq_of_getEntry?_eq_some α β _ l₁ k val heq)] at h + simp [h, heq] + case h_2 _ heq => + simp [heq] + . exact dl₁ + +theorem getValueCast?_diff_of_containsKey_eq_false_right [BEq α] [LawfulBEq α] + {l₁ l₂ : List ((a : α) × β a)} {k : α} + (dl₁ : DistinctKeys l₁) : + containsKey k l₂ = false → getValueCast? k (List.filter (fun p => !containsKey p.fst l₂) l₁) = + getValueCast? k l₁ := by + intro h₁ + rw [@List.getValueCast?_diff, h₁] + all_goals simp [dl₁] + +theorem getValueCast?_diff_of_containsKey_eq_false_left [BEq α] [LawfulBEq α] + {l₁ l₂ : List ((a : α) × β a)} {k : α} + (dl₁ : DistinctKeys l₁) : + containsKey k l₁ = false → getValueCast? k (List.filter (fun p => !containsKey p.fst l₂) l₁) = + none := by + intro h₁ + rw [List.getValueCast?_diff] + · simp [getValueCast?_eq_none h₁] + · exact dl₁ + +theorem List.getValueCast?_diff_of_containsKey_right [BEq α] [LawfulBEq α] + {l₁ l₂ : List ((a : α) × β a)} {k : α} + (dl₁ : DistinctKeys l₁) : + containsKey k l₂→ getValueCast? k (List.filter (fun p => !containsKey p.fst l₂) l₁) = + none := by + intro h₁ + rw [List.getValueCast?_diff] + · simp only [ite_eq_left_iff, Bool.not_eq_true] + intro h₂ + rw [h₂] at h₁ + contradiction + · exact dl₁ + theorem List.getKey?_filter_containsKey [BEq α] [EquivBEq α] {l₁ l₂ : List ((a : α) × β a)} {k : α} (dl₁ : DistinctKeys l₁) : @@ -5278,6 +5340,7 @@ theorem List.getKey?_filter_containsKey_of_containsKey_right [BEq α] [EquivBEq · exact dl₁ + /-- Internal implementation detail of the hash map -/ def eraseList [BEq α] (l toErase : List ((a : α) × β a)) : List ((a : α) × β a) := match toErase with @@ -5507,7 +5570,7 @@ theorem List.containsKey_filter_containsKey_iff [BEq α] [EquivBEq α] {l₁ l . simp . exact hl₁ -theorem containsKey_filter_containsKey_eq_false [BEq α] [EquivBEq α] {l₁ l₂ : List ((a : α) × β a)} {hl₁ : DistinctKeys l₁} {k : α} : +theorem List.containsKey_diff_eq_false [BEq α] [EquivBEq α] {l₁ l₂ : List ((a : α) × β a)} {hl₁ : DistinctKeys l₁} {k : α} : (containsKey k (List.filter (fun p => !containsKey p.fst (l₂)) l₁)) = (containsKey k l₁ && !containsKey k l₂) := by rw [containsKey_eq_isSome_getEntry?, containsKey_eq_isSome_getEntry?] @@ -5522,10 +5585,10 @@ theorem containsKey_filter_containsKey_eq_false [BEq α] [EquivBEq α] {l₁ l apply List.beq_of_getEntry?_eq_some heq . exact hl₁ -theorem containsKey_filter_containsKey_eq_false_iff [BEq α] [EquivBEq α] {l₁ l₂ : List ((a : α) × β a)} {hl₁ : DistinctKeys l₁} {k : α} : +theorem List.containsKey_diff_eq_false_iff [BEq α] [EquivBEq α] {l₁ l₂ : List ((a : α) × β a)} {hl₁ : DistinctKeys l₁} {k : α} : (containsKey k (List.filter (fun p => !containsKey p.fst l₂) l₁)) ↔ (containsKey k l₁ ∧ ¬ containsKey k l₂) := by - rw [containsKey_filter_containsKey_eq_false] + rw [List.containsKey_diff_eq_false] . simp . exact hl₁ @@ -5571,6 +5634,13 @@ theorem List.containsKey_filter_containsKey_eq_false_of_containsKey_eq_false_lef . simp [h] . exact hl₁ +theorem List.containsKey_diff_eq_false_of_containsKey_eq_false_left [BEq α] [EquivBEq α] {l₁ l₂ : List ((a : α) × β a)} {hl₁ : DistinctKeys l₁} {k : α} : + containsKey k l₁ = false → containsKey k (List.filter (fun p => !containsKey p.fst l₂) l₁) = false := by + intro h + rw [List.containsKey_diff_eq_false] + . simp [h] + . exact hl₁ + theorem List.containsKey_filter_containsKey_eq_false_of_containsKey_eq_false_right [BEq α] [EquivBEq α] {l₁ l₂ : List ((a : α) × β a)} {hl₁ : DistinctKeys l₁} {k : α} : containsKey k l₂ = false → containsKey k (List.filter (fun p => containsKey p.fst l₂) l₁) = false := by intro h @@ -5578,6 +5648,13 @@ theorem List.containsKey_filter_containsKey_eq_false_of_containsKey_eq_false_rig . simp [h] . exact hl₁ +theorem List.containsKey_diff_of_containsKey_eq_false_right [BEq α] [EquivBEq α] {l₁ l₂ : List ((a : α) × β a)} {hl₁ : DistinctKeys l₁} {k : α} : + containsKey k l₂ = true → containsKey k (List.filter (fun p => !containsKey p.fst l₂) l₁) = false := by + intro h + rw [List.containsKey_diff_eq_false] + . simp [h] + . exact hl₁ + theorem getValue?_filter_containsKey_eq_if_containsKey {β : Type v} [BEq α] [EquivBEq α] {l₁ l₂ : List ((_ : α) × β)} {k : α} (dl₁ : DistinctKeys l₁) : From 002545f37ce6ec57c6c54a06443e7ab62419ced4 Mon Sep 17 00:00:00 2001 From: Wojciech Rozowski Date: Fri, 14 Nov 2025 13:15:55 +0100 Subject: [PATCH 56/88] push intermediate work --- src/Std/Data/DHashMap/Internal/RawLemmas.lean | 185 ++++++++++++------ src/Std/Data/Internal/List/Associative.lean | 13 +- 2 files changed, 137 insertions(+), 61 deletions(-) diff --git a/src/Std/Data/DHashMap/Internal/RawLemmas.lean b/src/Std/Data/DHashMap/Internal/RawLemmas.lean index 0dab3d9ce0c9..2ebd9406de9d 100644 --- a/src/Std/Data/DHashMap/Internal/RawLemmas.lean +++ b/src/Std/Data/DHashMap/Internal/RawLemmas.lean @@ -3150,134 +3150,177 @@ theorem get?_diff_of_contains_right [LawfulBEq α] (h₁ : m₁.val.WF) (h₂ : revert h simp_to_model [diff, get?, contains] using List.getValueCast?_diff_of_containsKey_right - /- get -/ theorem get_diff [LawfulBEq α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} {h_contains : (m₁.diff m₂).contains k} : (m₁.diff m₂).get k h_contains = - m₁.get k ((contains_diff_iff h₁ h₂).1 h_contains).1 := by sorry + m₁.get k ((contains_diff_iff h₁ h₂).1 h_contains).1 := by + simp_to_model [diff, get, contains] using List.getValueCast_diff /- getD -/ theorem getD_diff [LawfulBEq α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} {fallback : β k} : (m₁.diff m₂).getD k fallback = - if m₂.contains k then fallback else m₁.getD k fallback := by sorry + if m₂.contains k then fallback else m₁.getD k fallback := by + simp_to_model [diff, getD, contains] using List.getValueCastD_diff -theorem getD_diff_of_contains_left [LawfulBEq α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) - {k : α} {fallback : β k} (h : m₁.contains k) (h' : m₂.contains k = false) : - (m₁.diff m₂).getD k fallback = m₁.getD k fallback := by sorry +theorem getD_diff_of_contains_eq_false_right [LawfulBEq α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) + {k : α} {fallback : β k} (h : m₂.contains k = false) : + (m₁.diff m₂).getD k fallback = m₁.getD k fallback := by + revert h + simp_to_model [diff, contains, getD] using List.getValueCastD_diff_of_containsKey_eq_false_right theorem getD_diff_of_contains_right [LawfulBEq α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} {fallback : β k} (h : m₂.contains k) : - (m₁.diff m₂).getD k fallback = fallback := by sorry + (m₁.diff m₂).getD k fallback = fallback := by + revert h + simp_to_model [diff, getD, contains] using List.getValueCastD_diff_of_containsKey_right theorem getD_diff_of_contains_eq_false_left [LawfulBEq α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} {fallback : β k} (h : m₁.contains k = false) : - (m₁.diff m₂).getD k fallback = fallback := by sorry + (m₁.diff m₂).getD k fallback = fallback := by + revert h + simp_to_model [diff, getD, contains] using List.getValueCastD_diff_of_containsKey_eq_false_left /- get! -/ theorem get!_diff [LawfulBEq α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} [Inhabited (β k)] : - (m₁.diff m₂).get! k = if m₂.contains k then default else m₁.get! k := by sorry + (m₁.diff m₂).get! k = if m₂.contains k then default else m₁.get! k := by + simp_to_model [diff, get!, contains] using List.getValueCastD_diff -theorem get!_diff_of_contains_left [LawfulBEq α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) - {k : α} [Inhabited (β k)] (h : m₁.contains k) (h' : m₂.contains k = false) : - (m₁.diff m₂).get! k = m₁.get! k := by sorry +theorem get!_diff_of_contains_eq_false_right [LawfulBEq α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) + {k : α} [Inhabited (β k)] (h : m₂.contains k = false) : + (m₁.diff m₂).get! k = m₁.get! k := by + revert h + simp_to_model [diff, contains, get!] using List.getValueCastD_diff_of_containsKey_eq_false_right theorem get!_diff_of_contains_right [LawfulBEq α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} [Inhabited (β k)] (h : m₂.contains k) : - (m₁.diff m₂).get! k = default := by sorry + (m₁.diff m₂).get! k = default := by + revert h + simp_to_model [diff, get!, contains] using List.getValueCastD_diff_of_containsKey_right theorem get!_diff_of_contains_eq_false_left [LawfulBEq α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} [Inhabited (β k)] (h : m₁.contains k = false) : - (m₁.diff m₂).get! k = default := by sorry + (m₁.diff m₂).get! k = default := by + revert h + simp_to_model [diff, get!, contains] using List.getValueCastD_diff_of_containsKey_eq_false_left /- getKey? -/ theorem getKey?_diff [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} : - (m₁.diff m₂).getKey? k = if m₂.contains k then none else m₁.getKey? k := by sorry + (m₁.diff m₂).getKey? k = if m₂.contains k then none else m₁.getKey? k := by + simp_to_model [diff, contains, getKey?] using List.getKey?_diff -theorem getKey?_diff_of_contains_left [EquivBEq α] [LawfulHashable α] - (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} - (h : m₁.contains k) (h' : m₂.contains k = false) : - (m₁.diff m₂).getKey? k = m₁.getKey? k := by sorry +theorem getKey?_diff_of_contains_eq_false_right [EquivBEq α] [LawfulHashable α] + (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} (h : m₂.contains k = false) : + (m₁.diff m₂).getKey? k = m₁.getKey? k := by + revert h + simp_to_model [contains, getKey?, diff] using List.getKey?_diff_of_containsKey_eq_false_right theorem getKey?_diff_of_contains_eq_false_left [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} (h : m₁.contains k = false) : - (m₁.diff m₂).getKey? k = none := by sorry + (m₁.diff m₂).getKey? k = none := by + revert h + simp_to_model [contains, getKey?, diff] using List.getKey?_diff_of_containsKey_eq_false_left theorem getKey?_diff_of_contains_right [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} (h : m₂.contains k) : - (m₁.diff m₂).getKey? k = none := by sorry + (m₁.diff m₂).getKey? k = none := by + revert h + simp_to_model [contains, getKey?, diff] using List.getKey?_diff_of_containsKey_right /- getKey -/ theorem getKey_diff [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} {h_contains : (m₁.diff m₂).contains k} : (m₁.diff m₂).getKey k h_contains = - m₁.getKey k ((contains_diff_iff h₁ h₂).1 h_contains).1 := by sorry + m₁.getKey k ((contains_diff_iff h₁ h₂).1 h_contains).1 := by + simp_to_model [diff, contains, getKey] using List.getKey_diff /- getKeyD -/ theorem getKeyD_diff [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k fallback : α} : (m₁.diff m₂).getKeyD k fallback = - if m₂.contains k then fallback else m₁.getKeyD k fallback := by sorry + if m₂.contains k then fallback else m₁.getKeyD k fallback := by + simp_to_model [diff, getKeyD, contains] using List.getKeyD_diff -theorem getKeyD_diff_of_contains_left [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) - (h₂ : m₂.val.WF) {k fallback : α} - (h : m₁.contains k) (h' : m₂.contains k = false) : - (m₁.diff m₂).getKeyD k fallback = m₁.getKeyD k fallback := by sorry +theorem getKeyD_diff_of_contains_eq_false_right [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) + (h₂ : m₂.val.WF) {k fallback : α} (h : m₂.contains k = false) : + (m₁.diff m₂).getKeyD k fallback = m₁.getKeyD k fallback := by + revert h + simp_to_model [contains, diff, getKeyD] using List.getKeyD_diff_of_containsKey_eq_false_right theorem getKeyD_diff_of_contains_right [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k fallback : α} (h : m₂.contains k) : - (m₁.diff m₂).getKeyD k fallback = fallback := by sorry + (m₁.diff m₂).getKeyD k fallback = fallback := by + revert h + simp_to_model [diff, getKeyD, contains] using List.getKeyD_diff_of_containsKey_right theorem getKeyD_diff_of_contains_eq_false_left [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k fallback : α} (h : m₁.contains k = false) : - (m₁.diff m₂).getKeyD k fallback = fallback := by sorry + (m₁.diff m₂).getKeyD k fallback = fallback := by + revert h + simp_to_model [diff, getKeyD, contains] using List.getKeyD_diff_of_containsKey_eq_false_left /- getKey! -/ theorem getKey!_diff [EquivBEq α] [LawfulHashable α] [Inhabited α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} : - (m₁.diff m₂).getKey! k = if m₂.contains k then default else m₁.getKey! k := by sorry + (m₁.diff m₂).getKey! k = if m₂.contains k then default else m₁.getKey! k := by + simp_to_model [diff, getKey!, contains] using List.getKeyD_diff -theorem getKey!_diff_of_contains_left [EquivBEq α] [LawfulHashable α] [Inhabited α] (h₁ : m₁.val.WF) - (h₂ : m₂.val.WF) {k : α} (h : m₁.contains k) (h' : m₂.contains k = false) : - (m₁.diff m₂).getKey! k = m₁.getKey! k := by sorry +theorem getKey!_diff_of_contains_eq_false_right [EquivBEq α] [LawfulHashable α] [Inhabited α] (h₁ : m₁.val.WF) + (h₂ : m₂.val.WF) {k : α} (h : m₂.contains k = false) : + (m₁.diff m₂).getKey! k = m₁.getKey! k := by + revert h + simp_to_model [diff, getKey!, contains] using List.getKeyD_diff_of_containsKey_eq_false_right theorem getKey!_diff_of_contains_right [EquivBEq α] [LawfulHashable α] [Inhabited α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} (h : m₂.contains k) : - (m₁.diff m₂).getKey! k = default := by sorry + (m₁.diff m₂).getKey! k = default := by + revert h + simp_to_model [diff, getKey!, contains] using List.getKeyD_diff_of_containsKey_right theorem getKey!_diff_of_contains_eq_false_left [EquivBEq α] [LawfulHashable α] [Inhabited α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} (h : m₁.contains k = false) : - (m₁.diff m₂).getKey! k = default := by sorry + (m₁.diff m₂).getKey! k = default := by + revert h + simp_to_model [diff, getKey!, contains] using List.getKeyD_diff_of_containsKey_eq_false_left /- size -/ theorem size_diff_le_size_left [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) : - (m₁.diff m₂).1.size ≤ m₁.1.size := by sorry + (m₁.diff m₂).1.size ≤ m₁.1.size := by + simp_to_model [diff, size] using List.length_diff_le theorem size_diff_eq_size_left_of_disjoint [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) (h : ∀ (a : α), m₁.contains a → m₂.contains a = false) : - (m₁.diff m₂).1.size = m₁.1.size := by sorry + (m₁.diff m₂).1.size = m₁.1.size := by + revert h + simp_to_model [diff, size, contains] using List.length_diff_eq_length_left theorem size_diff_add_size_inter_eq_size_left [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) : - (m₁.diff m₂).1.size + (m₁.inter m₂).1.size = m₁.1.size := by sorry + (m₁.diff m₂).1.size + (m₁.inter m₂).1.size = m₁.1.size := by + simp_to_model [diff, inter, size] using List.size_diff_add_size_inter_eq_size_left /- isEmpty -/ @[simp] theorem isEmpty_diff_of_isEmpty_left [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) (h : m₁.1.isEmpty) : - (m₁.diff m₂).1.isEmpty = true := by sorry + (m₁.diff m₂).1.isEmpty = true := by + revert h + simp_to_model [isEmpty, diff] using List.isEmpty_diff_left theorem isEmpty_diff_iff [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) : - (m₁.diff m₂).1.isEmpty ↔ ∀ k, m₁.contains k → m₂.contains k := by sorry + (m₁.diff m₂).1.isEmpty ↔ ∀ k, m₁.contains k → m₂.contains k := by + simp_to_model [diff, contains, isEmpty] using List.isEmpty_diff_iff theorem isEmpty_diff_of_subset [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) (h : ∀ k, m₁.contains k → m₂.contains k) : - (m₁.diff m₂).1.isEmpty = true := by sorry + (m₁.diff m₂).1.isEmpty = true := by + revert h + simp_to_model [diff, contains, isEmpty] using List.isEmpty_diff_of_subset end Diff @@ -3287,59 +3330,81 @@ variable {β : Type v} {m₁ m₂ : Raw₀ α (fun _ => β)} /- get? -/ theorem get?_diff [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} : - Const.get? (m₁.diff m₂) k = if m₂.contains k then none else Const.get? m₁ k := by sorry + Const.get? (m₁.diff m₂) k = if m₂.contains k then none else Const.get? m₁ k := by + simp_to_model [diff, Const.get?, contains] using List.getValue?_diff -theorem get?_diff_of_contains_left [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) - {k : α} (h : m₁.contains k) (h' : m₂.contains k = false) : - Const.get? (m₁.diff m₂) k = Const.get? m₁ k := by sorry +theorem get?_diff_of_contains_eq_false_right [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) + {k : α} (h : m₂.contains k = false) : + Const.get? (m₁.diff m₂) k = Const.get? m₁ k := by + revert h + simp_to_model [diff, contains, Const.get?] using List.getValue?_diff_of_containsKey_eq_false_right theorem get?_diff_of_contains_eq_false_left [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} (h : m₁.contains k = false) : - Const.get? (m₁.diff m₂) k = none := by sorry + Const.get? (m₁.diff m₂) k = none := by + revert h + simp_to_model [diff, Const.get?, contains] using List.getValue?_diff_of_containsKey_eq_false_left theorem get?_diff_of_contains_right [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} (h : m₂.contains k) : - Const.get? (m₁.diff m₂) k = none := by sorry + Const.get? (m₁.diff m₂) k = none := by + revert h + simp_to_model [diff, Const.get?, contains] using List.getValue?_diff_of_containsKey_right /- get -/ theorem get_diff [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} {h_contains : (m₁.diff m₂).contains k} : Const.get (m₁.diff m₂) k h_contains = - Const.get m₁ k ((contains_diff_iff h₁ h₂).1 h_contains).1 := by sorry + Const.get m₁ k ((contains_diff_iff h₁ h₂).1 h_contains).1 := by + simp_to_model [diff, Const.get, contains] using List.getValue_diff /- getD -/ theorem getD_diff [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} {fallback : β} : Const.getD (m₁.diff m₂) k fallback = - if m₂.contains k then fallback else Const.getD m₁ k fallback := by sorry + if m₂.contains k then fallback else Const.getD m₁ k fallback := by + simp_to_model [diff, Const.getD, contains] using List.getValueD_diff -theorem getD_diff_of_contains_left [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) - {k : α} {fallback : β} (h : m₁.contains k) (h' : m₂.contains k = false) : - Const.getD (m₁.diff m₂) k fallback = Const.getD m₁ k fallback := by sorry +theorem getD_diff_of_contains_eq_false_right [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) + {k : α} {fallback : β} (h : m₂.contains k = false) : + Const.getD (m₁.diff m₂) k fallback = Const.getD m₁ k fallback := by + revert h + simp_to_model [diff, contains, Const.getD] using List.getValueD_diff_of_containsKey_eq_false_right theorem getD_diff_of_contains_right [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} {fallback : β} (h : m₂.contains k) : - Const.getD (m₁.diff m₂) k fallback = fallback := by sorry + Const.getD (m₁.diff m₂) k fallback = fallback := by + revert h + simp_to_model [diff, Const.getD, contains] using List.getValueD_diff_of_containsKey_right theorem getD_diff_of_contains_eq_false_left [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} {fallback : β} (h : m₁.contains k = false) : - Const.getD (m₁.diff m₂) k fallback = fallback := by sorry + Const.getD (m₁.diff m₂) k fallback = fallback := by + revert h + simp_to_model [diff, Const.getD, contains] using List.getValueD_diff_of_containsKey_eq_false_left /- get! -/ theorem get!_diff [EquivBEq α] [LawfulHashable α] [Inhabited β] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} : - Const.get! (m₁.diff m₂) k = if m₂.contains k then default else Const.get! m₁ k := by sorry + Const.get! (m₁.diff m₂) k = if m₂.contains k then default else Const.get! m₁ k := by + simp_to_model [diff, Const.get!, contains] using List.getValueD_diff -theorem get!_diff_of_contains_left [EquivBEq α] [LawfulHashable α] [Inhabited β] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) - {k : α} (h : m₁.contains k) (h' : m₂.contains k = false) : - Const.get! (m₁.diff m₂) k = Const.get! m₁ k := by sorry +theorem get!_diff_of_contains_eq_false_right [EquivBEq α] [LawfulHashable α] [Inhabited β] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) + {k : α} (h : m₂.contains k = false) : + Const.get! (m₁.diff m₂) k = Const.get! m₁ k := by + revert h + simp_to_model [diff, contains, Const.get!] using List.getValueD_diff_of_containsKey_eq_false_right theorem get!_diff_of_contains_right [EquivBEq α] [LawfulHashable α] [Inhabited β] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} (h : m₂.contains k) : - Const.get! (m₁.diff m₂) k = default := by sorry + Const.get! (m₁.diff m₂) k = default := by + revert h + simp_to_model [diff, Const.get!, contains] using List.getValueD_diff_of_containsKey_right theorem get!_diff_of_contains_eq_false_left [EquivBEq α] [LawfulHashable α] [Inhabited β] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} (h : m₁.contains k = false) : - Const.get! (m₁.diff m₂) k = default := by sorry + Const.get! (m₁.diff m₂) k = default := by + revert h + simp_to_model [diff, Const.get!, contains] using List.getValueD_diff_of_containsKey_eq_false_left end Const diff --git a/src/Std/Data/Internal/List/Associative.lean b/src/Std/Data/Internal/List/Associative.lean index bf59714a74a8..af35e40e3efb 100644 --- a/src/Std/Data/Internal/List/Associative.lean +++ b/src/Std/Data/Internal/List/Associative.lean @@ -5285,7 +5285,7 @@ theorem getValueCast?_diff_of_containsKey_eq_false_left [BEq α] [LawfulBEq α] theorem List.getValueCast?_diff_of_containsKey_right [BEq α] [LawfulBEq α] {l₁ l₂ : List ((a : α) × β a)} {k : α} (dl₁ : DistinctKeys l₁) : - containsKey k l₂→ getValueCast? k (List.filter (fun p => !containsKey p.fst l₂) l₁) = + containsKey k l₂ → getValueCast? k (List.filter (fun p => !containsKey p.fst l₂) l₁) = none := by intro h₁ rw [List.getValueCast?_diff] @@ -5592,6 +5592,17 @@ theorem List.containsKey_diff_eq_false_iff [BEq α] [EquivBEq α] {l₁ l₂ : L . simp . exact hl₁ +theorem List.getValueCast_diff [BEq α] [LawfulBEq α] + {l₁ l₂ : List ((a : α) × β a)} {k : α} + (dl₁ : DistinctKeys l₁) {h₁ h₂} : + getValueCast k (List.filter (fun p => !containsKey p.fst l₂) l₁) h₁ = getValueCast k l₁ h₂ := by + suffices some (getValueCast k (List.filter (fun p => !containsKey p.fst l₂) l₁) h₁) = some (getValueCast k l₁ h₂) by injections + simp only [← getValueCast?_eq_some_getValueCast] + rw [List.getValueCast?_diff dl₁] + rw [List.containsKey_diff_eq_false_iff] at h₁ + simp [h₁.2] + exact dl₁ + theorem List.getValueCast_filter_containsKey [BEq α] [LawfulBEq α] {l₁ l₂ : List ((a : α) × β a)} {k : α} (dl₁ : DistinctKeys l₁) {h₁ h₂} : From 0b840add73106a1ab39ea9cec28594dade904d80 Mon Sep 17 00:00:00 2001 From: Wojciech Rozowski Date: Fri, 14 Nov 2025 13:33:52 +0100 Subject: [PATCH 57/88] push --- src/Std/Data/Internal/List/Associative.lean | 83 +++++++++++++++++++++ 1 file changed, 83 insertions(+) diff --git a/src/Std/Data/Internal/List/Associative.lean b/src/Std/Data/Internal/List/Associative.lean index af35e40e3efb..09542e5a9552 100644 --- a/src/Std/Data/Internal/List/Associative.lean +++ b/src/Std/Data/Internal/List/Associative.lean @@ -5666,6 +5666,77 @@ theorem List.containsKey_diff_of_containsKey_eq_false_right [BEq α] [EquivBEq . simp [h] . exact hl₁ +theorem List.getValueCastD_diff [BEq α] [LawfulBEq α] + {l₁ l₂ : List ((a : α) × β a)} {k : α} {fallback : β k} + (dl₁ : DistinctKeys l₁) : + getValueCastD k (List.filter (fun p => !containsKey p.fst l₂) l₁) fallback = + if containsKey k l₂ = true then fallback else getValueCastD k l₁ fallback := by + split + case isTrue h => + apply getValueCastD_eq_fallback + apply List.containsKey_diff_of_containsKey_eq_false_right + · exact dl₁ + · exact h + case isFalse h => + rw [getValueCastD_eq_getValueCast?, getValueCastD_eq_getValueCast?] + congr 1 + rw [List.getValueCast?_diff dl₁] + simp [h] + +theorem List.getValueCastD_diff_of_containsKey_eq_false_left [BEq α] [LawfulBEq α] + {l₁ l₂ : List ((a : α) × β a)} {k : α} {fallback} + (dl₁ : DistinctKeys l₁) : + containsKey k l₁ = false → getValueCastD k (List.filter (fun p => !containsKey p.fst l₂) l₁) fallback = fallback := by + intro h + apply getValueCastD_eq_fallback + apply List.containsKey_diff_eq_false_of_containsKey_eq_false_left + · exact dl₁ + · exact h + +theorem List.getKey?_diff [BEq α] [EquivBEq α] + {l₁ l₂ : List ((a : α) × β a)} {k : α} + (dl₁ : DistinctKeys l₁) : + getKey? k (List.filter (fun p => !containsKey p.fst l₂) l₁) = + if containsKey k l₂ = true then none else getKey? k l₁ := by + simp only [getKey?_eq_getEntry?] + split + case isTrue h => + suffices (getEntry? k (List.filter (fun p => !containsKey p.fst l₂) l₁)) = none by + simp [this] + rw [getEntry?_filter] + · simp [Option.filter] + split + case h_1 _ kv heq => + have key_eq := beq_of_getEntry?_eq_some heq + rw [@containsKey_congr _ _ _ _ l₂ _ _ (PartialEquivBEq.symm key_eq)] at h + simp [h] + case h_2 => rfl + · exact dl₁ + case isFalse h => + rw [getEntry?_filter] + · congr 1 + simp [Option.filter] + split + case h_1 _ val heq => + rw [heq] + have key_eq := beq_of_getEntry?_eq_some heq + simp at h + rw [@containsKey_congr _ _ _ _ l₂ _ _ (PartialEquivBEq.symm key_eq)] at h + simp [h] + case h_2 _ heq => + simp only [heq] + · exact dl₁ + +theorem List.getValueCastD_diff_of_containsKey_right [BEq α] [LawfulBEq α] + {l₁ l₂ : List ((a : α) × β a)} {k : α} {fallback} + (dl₁ : DistinctKeys l₁) : + containsKey k l₂ = true → getValueCastD k (List.filter (fun p => !containsKey p.fst l₂) l₁) fallback = fallback := by + intro h + apply getValueCastD_eq_fallback + apply List.containsKey_diff_of_containsKey_eq_false_right + · exact dl₁ + · exact h + theorem getValue?_filter_containsKey_eq_if_containsKey {β : Type v} [BEq α] [EquivBEq α] {l₁ l₂ : List ((_ : α) × β)} {k : α} (dl₁ : DistinctKeys l₁) : @@ -5770,6 +5841,18 @@ theorem List.getValueCastD_filter_containsKey_of_containsKey_eq_false_right [BEq . exact dl₁ . exact h +theorem List.getValueCastD_diff_of_containsKey_eq_false_right [BEq α] [LawfulBEq α] + {l₁ l₂ : List ((a : α) × β a)} {k : α} {fallback} + (dl₁ : DistinctKeys l₁) : + containsKey k l₂ = false → + getValueCastD k (List.filter (fun p => !containsKey p.fst l₂) l₁) fallback = + getValueCastD k l₁ fallback := by + intro h + rw [getValueCastD_eq_getValueCast?, getValueCastD_eq_getValueCast?] + congr 1 + rw [List.getValueCast?_diff dl₁] + simp [h] + theorem List.getKeyD_filter_containsKey [BEq α] [EquivBEq α] {l₁ l₂ : List ((a : α) × β a)} {k : α} {fallback : α} (dl₁ : DistinctKeys l₁) : From d33bb5f561dfd4ff96d593dd7b5bb1cba24af9c3 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Wojciech=20R=C3=B3=C5=BCowski?= Date: Fri, 14 Nov 2025 12:34:13 +0000 Subject: [PATCH 58/88] Update src/Std/Data/Internal/List/Associative.lean Co-authored-by: Markus Himmel --- src/Std/Data/Internal/List/Associative.lean | 2 -- 1 file changed, 2 deletions(-) diff --git a/src/Std/Data/Internal/List/Associative.lean b/src/Std/Data/Internal/List/Associative.lean index d16d183287f9..9e783c82ae7b 100644 --- a/src/Std/Data/Internal/List/Associative.lean +++ b/src/Std/Data/Internal/List/Associative.lean @@ -6517,8 +6517,6 @@ theorem nil_of_containsKey_eq_false [BEq α] [EquivBEq α] {l : List ((a : α) intro hyp k simp [hyp] - - theorem isEmpty_filter_containsKey_left [BEq α] [EquivBEq α] {l₁ l₂ : List ((a : α) × β a)} : l₁.isEmpty → (List.filter (fun p => containsKey p.fst l₂) l₁).isEmpty := by From ae0385b9809c33750d7de640460fb23f75c19bf5 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Wojciech=20R=C3=B3=C5=BCowski?= Date: Fri, 14 Nov 2025 12:34:28 +0000 Subject: [PATCH 59/88] Update src/Std/Data/DHashMap/Basic.lean Co-authored-by: Markus Himmel --- src/Std/Data/DHashMap/Basic.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/src/Std/Data/DHashMap/Basic.lean b/src/Std/Data/DHashMap/Basic.lean index 3be9b0625450..52ccac90dba3 100644 --- a/src/Std/Data/DHashMap/Basic.lean +++ b/src/Std/Data/DHashMap/Basic.lean @@ -339,7 +339,7 @@ This function always merges the smaller map into the larger map, so the expected -/ @[inline] def inter [BEq α] [Hashable α] (m₁ m₂ : DHashMap α β) : DHashMap α β where inner := Raw₀.inter ⟨m₁.1, m₁.2.size_buckets_pos⟩ ⟨m₂.1, m₂.2.size_buckets_pos⟩ - wf := Std.DHashMap.Raw.WF.inter₀ m₁.2 m₂.2 + wf := Std.DHashMap.Raw.WF.inter₀ m₁.2 m₂.2 instance [BEq α] [Hashable α] : Union (DHashMap α β) := ⟨union⟩ instance [BEq α] [Hashable α] : Inter (DHashMap α β) := ⟨inter⟩ From 7aa501b396402f9a68b78f1528ea815f00c8e73a Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Wojciech=20R=C3=B3=C5=BCowski?= Date: Fri, 14 Nov 2025 12:34:52 +0000 Subject: [PATCH 60/88] Update src/Std/Data/DHashMap/Lemmas.lean Co-authored-by: Markus Himmel --- src/Std/Data/DHashMap/Lemmas.lean | 8 ++------ 1 file changed, 2 insertions(+), 6 deletions(-) diff --git a/src/Std/Data/DHashMap/Lemmas.lean b/src/Std/Data/DHashMap/Lemmas.lean index e738c0df2b0a..e3b98e803c86 100644 --- a/src/Std/Data/DHashMap/Lemmas.lean +++ b/src/Std/Data/DHashMap/Lemmas.lean @@ -2237,12 +2237,8 @@ theorem isEmpty_inter_right [EquivBEq α] [LawfulHashable α] (h : m₂.isEmpty) theorem isEmpty_inter_iff [EquivBEq α] [LawfulHashable α] : (m₁ ∩ m₂).isEmpty ↔ ∀ k, k ∈ m₁ → k ∉ m₂ := by - conv => - rhs - ext x - rhs - rw [← contains_eq_false_iff_not_mem] - exact @Raw₀.isEmpty_inter_iff _ _ _ _ ⟨m₁.1, m₁.2.size_buckets_pos⟩ ⟨m₂.1, m₂.2.size_buckets_pos⟩ _ _ m₁.wf m₂.wf + simpa only [mem_iff_contains, Bool.not_eq_true] using + @Raw₀.isEmpty_inter_iff _ _ _ _ ⟨m₁.1, m₁.2.size_buckets_pos⟩ ⟨m₂.1, m₂.2.size_buckets_pos⟩ _ _ m₁.wf m₂.wf end Inter From d3cb4629afbc7559043dd4af7da42868cbbe8d48 Mon Sep 17 00:00:00 2001 From: Wojciech Rozowski Date: Fri, 14 Nov 2025 15:03:00 +0100 Subject: [PATCH 61/88] Fix WF --- src/Std/Data/DHashMap/Internal/RawLemmas.lean | 2 +- src/Std/Data/DHashMap/Internal/WF.lean | 18 ++---------------- 2 files changed, 3 insertions(+), 17 deletions(-) diff --git a/src/Std/Data/DHashMap/Internal/RawLemmas.lean b/src/Std/Data/DHashMap/Internal/RawLemmas.lean index 547848631cda..da6e87efa250 100644 --- a/src/Std/Data/DHashMap/Internal/RawLemmas.lean +++ b/src/Std/Data/DHashMap/Internal/RawLemmas.lean @@ -91,7 +91,7 @@ macro_rules | apply Raw₀.wf_insertMany₀ | apply Raw₀.Const.wf_insertMany₀ | apply Raw₀.Const.wf_insertManyIfNewUnit₀ | apply Raw₀.wf_union₀ | apply Raw.WF.filter₀ | apply Raw₀.wf_map₀ | apply Raw₀.wf_filterMap₀ - | apply Raw.WF.emptyWithCapacity₀ | apply Raw₀.wf_inter₀) <;> wf_trivial) + | apply Raw.WF.emptyWithCapacity₀ | apply Raw.WF.inter₀) <;> wf_trivial) /-- Internal implementation detail of the hash map -/ scoped macro "empty" : tactic => `(tactic| { intros; simp_all [List.isEmpty_iff] } ) diff --git a/src/Std/Data/DHashMap/Internal/WF.lean b/src/Std/Data/DHashMap/Internal/WF.lean index 6c4140cd1576..e43ba71f48b5 100644 --- a/src/Std/Data/DHashMap/Internal/WF.lean +++ b/src/Std/Data/DHashMap/Internal/WF.lean @@ -1149,7 +1149,7 @@ theorem toListModel_unionₘ [BEq α] [Hashable α] [EquivBEq α] [LawfulHashabl · exact toListModel_insertListIfNewₘ ‹_› · exact toListModel_insertListₘ ‹_› -theorem wfimp_inter [BEq α] [EquivBEq α] [Hashable α] [LawfulHashable α] +theorem wfImp_inter [BEq α] [EquivBEq α] [Hashable α] [LawfulHashable α] {m₁ m₂ : Raw α β} {h₁ : 0 < m₁.buckets.size} {h₂ : 0 < m₂.buckets.size} (wh₁ : Raw.WFImp m₁) : Raw.WFImp (Raw₀.inter ⟨m₁, h₁⟩ ⟨m₂, h₂⟩).val := by rw [inter] @@ -1184,7 +1184,7 @@ theorem WF.out [BEq α] [Hashable α] [i₁ : EquivBEq α] [i₂ : LawfulHashabl | constModify₀ _ h => exact Raw₀.Const.wfImp_modify (by apply h) | alter₀ _ h => exact Raw₀.wfImp_alter (by apply h) | constAlter₀ _ h => exact Raw₀.Const.wfImp_alter (by apply h) - | inter₀ _ _ h _ => exact Raw₀.wfimp_inter (by apply h) + | inter₀ _ _ h _ => exact Raw₀.wfImp_inter (by apply h) end Raw @@ -1412,20 +1412,6 @@ theorem toListModel_inter [BEq α] [EquivBEq α] [Hashable α] [LawfulHashable exact Perm.trans (toListModel_interSmallerₘ _ _ hm₁) (List.interSmaller_perm_filter _ _ hm₁.distinct) -theorem wf_inter₀ [BEq α] [Hashable α] [LawfulHashable α] - {m₁ m₂ : Raw α β} {h₁ : 0 < m₁.buckets.size} {h₂ : 0 < m₂.buckets.size} (wh₁ : m₁.WF) : - (Raw₀.inter ⟨m₁, h₁⟩ ⟨m₂, h₂⟩).1.WF := by - rw [inter] - split - · apply Raw.WF.filter₀ wh₁ - · rw [interSmaller] - apply @Raw.fold_induction _ _ _ (fun sofar k x => interSmallerFn ⟨m₁, h₁⟩ sofar k) _ _ (·.val.WF) Raw.WF.emptyWithCapacity₀ - intro acc a b wf - rw [interSmallerFn] - split - · apply Raw.WF.insert₀ wf - · apply wf - /-! # `Const.insertListₘ` -/ theorem Const.toListModel_insertListₘ {β : Type v} [BEq α] [Hashable α] [EquivBEq α] From 3bde3f9e058801f32ea600f22edc39e23e26ce49 Mon Sep 17 00:00:00 2001 From: Wojciech Rozowski Date: Fri, 14 Nov 2025 15:05:33 +0100 Subject: [PATCH 62/88] Fix docstrings --- src/Std/Data/DHashMap/Basic.lean | 2 +- src/Std/Data/HashSet/Basic.lean | 2 +- 2 files changed, 2 insertions(+), 2 deletions(-) diff --git a/src/Std/Data/DHashMap/Basic.lean b/src/Std/Data/DHashMap/Basic.lean index 52ccac90dba3..b8fdbaeaa12a 100644 --- a/src/Std/Data/DHashMap/Basic.lean +++ b/src/Std/Data/DHashMap/Basic.lean @@ -334,7 +334,7 @@ This function always merges the smaller map into the larger map, so the expected /-- Computes the intersection of the given hash maps. The result will only contain entries from the first map. -This function always merges the smaller map into the larger map, so the expected runtime is +This function always iterates through the smaller map, so the expected runtime is `O(min(m₁.size, m₂.size))`. -/ @[inline] def inter [BEq α] [Hashable α] (m₁ m₂ : DHashMap α β) : DHashMap α β where diff --git a/src/Std/Data/HashSet/Basic.lean b/src/Std/Data/HashSet/Basic.lean index c74fa8b0d9c6..91264d6d6a1c 100644 --- a/src/Std/Data/HashSet/Basic.lean +++ b/src/Std/Data/HashSet/Basic.lean @@ -246,7 +246,7 @@ instance [BEq α] [Hashable α] : Union (HashSet α) := ⟨union⟩ /-- Computes the intersection of the given hash sets. The result will only contain entries from the first map. -This function always merges the smaller set into the larger set, so the expected runtime is +This function always iterates through the smaller set, so the expected runtime is `O(min(m₁.size, m₂.size))`. -/ @[inline] def inter [BEq α] [Hashable α] (m₁ m₂ : HashSet α) : HashSet α := From 66304f9ea33caa65f2609b8301bb405bc7d7fc4a Mon Sep 17 00:00:00 2001 From: Wojciech Rozowski Date: Fri, 14 Nov 2025 15:08:05 +0100 Subject: [PATCH 63/88] Fix simp attributes --- src/Std/Data/DHashMap/Internal/RawLemmas.lean | 6 +++--- src/Std/Data/DHashMap/Lemmas.lean | 6 +++--- src/Std/Data/DHashMap/RawLemmas.lean | 6 +++--- src/Std/Data/HashMap/Lemmas.lean | 4 ++-- src/Std/Data/HashMap/RawLemmas.lean | 4 ++-- src/Std/Data/HashSet/Lemmas.lean | 2 +- src/Std/Data/HashSet/RawLemmas.lean | 2 +- 7 files changed, 15 insertions(+), 15 deletions(-) diff --git a/src/Std/Data/DHashMap/Internal/RawLemmas.lean b/src/Std/Data/DHashMap/Internal/RawLemmas.lean index da6e87efa250..4176bf129afa 100644 --- a/src/Std/Data/DHashMap/Internal/RawLemmas.lean +++ b/src/Std/Data/DHashMap/Internal/RawLemmas.lean @@ -2817,7 +2817,7 @@ theorem get?_inter_of_contains_eq_false_right [LawfulBEq α] (h₁ : m₁.val.WF simp_to_model [inter, get?, contains] using List.getValueCast?_filter_containsKey_of_containsKey_eq_false_right /- get -/ -theorem get_inter [LawfulBEq α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) +@[simp] theorem get_inter [LawfulBEq α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} {h_contains : (m₁.inter m₂).contains k} : (m₁.inter m₂).get k h_contains = m₁.get k ((contains_inter_iff h₁ h₂).1 h_contains).1 := by @@ -2899,7 +2899,7 @@ theorem getKey?_inter_of_contains_eq_false_left [EquivBEq α] [LawfulHashable α simp_to_model [contains, getKey?, inter] using List.getKey?_filter_containsKey_of_containsKey_eq_false_left /- getKey -/ -theorem getKey_inter [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) +@[simp] theorem getKey_inter [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} {h_contains : (m₁.inter m₂).contains k} : (m₁.inter m₂).getKey k h_contains = m₁.getKey k (by simp [contains_inter_iff h₁ h₂] at h_contains; exact h_contains.1) := by @@ -3033,7 +3033,7 @@ theorem get?_inter_of_contains_eq_false_right [EquivBEq α] [LawfulHashable α] revert h simp_to_model [inter, Const.get?, contains] using List.getValue?_filter_of_containsKey_eq_false_right -theorem get_inter [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) +@[simp] theorem get_inter [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} {h_contains : (m₁.inter m₂).contains k} : Const.get (m₁.inter m₂) k h_contains = Const.get m₁ k ((contains_inter_iff h₁ h₂).1 h_contains).1 := by diff --git a/src/Std/Data/DHashMap/Lemmas.lean b/src/Std/Data/DHashMap/Lemmas.lean index e3b98e803c86..b33cf30a1dcf 100644 --- a/src/Std/Data/DHashMap/Lemmas.lean +++ b/src/Std/Data/DHashMap/Lemmas.lean @@ -2063,7 +2063,7 @@ theorem get?_inter_of_not_mem_right [LawfulBEq α] exact @Raw₀.get?_inter_of_contains_eq_false_right _ _ _ _ ⟨m₁.1, m₁.2.size_buckets_pos⟩ ⟨m₂.1, m₂.2.size_buckets_pos⟩ _ m₁.2 m₂.2 k not_mem /- get -/ -theorem get_inter [LawfulBEq α] +@[simp] theorem get_inter [LawfulBEq α] {k : α} {h_mem : k ∈ m₁ ∩ m₂} : (m₁ ∩ m₂).get k h_mem = m₁.get k ((mem_inter_iff.1 h_mem).1) := by @@ -2140,7 +2140,7 @@ theorem getKey?_inter_of_not_mem_left [EquivBEq α] [LawfulHashable α] exact @Raw₀.getKey?_inter_of_contains_eq_false_left _ _ _ _ ⟨m₁.1, m₁.2.size_buckets_pos⟩ ⟨m₂.1, m₂.2.size_buckets_pos⟩ _ _ m₁.2 m₂.2 k not_mem /- getKey -/ -theorem getKey_inter [EquivBEq α] [LawfulHashable α] +@[simp] theorem getKey_inter [EquivBEq α] [LawfulHashable α] {k : α} {h_mem : k ∈ m₁ ∩ m₂} : (m₁ ∩ m₂).getKey k h_mem = m₁.getKey k ((mem_inter_iff.1 h_mem).1) := by @@ -2270,7 +2270,7 @@ theorem get?_inter_of_not_mem_right [EquivBEq α] [LawfulHashable α] exact @Raw₀.Const.get?_inter_of_contains_eq_false_right _ _ _ _ ⟨m₁.1, m₁.2.size_buckets_pos⟩ ⟨m₂.1, m₂.2.size_buckets_pos⟩ _ _ m₁.2 m₂.2 k not_mem /- get -/ -theorem get_inter [EquivBEq α] [LawfulHashable α] +@[simp] theorem get_inter [EquivBEq α] [LawfulHashable α] {k : α} {h_mem : k ∈ m₁ ∩ m₂} : Const.get (m₁.inter m₂) k h_mem = Const.get m₁ k ((mem_inter_iff.1 h_mem).1) := by diff --git a/src/Std/Data/DHashMap/RawLemmas.lean b/src/Std/Data/DHashMap/RawLemmas.lean index f91977b5d9fb..5d3aa2fd92be 100644 --- a/src/Std/Data/DHashMap/RawLemmas.lean +++ b/src/Std/Data/DHashMap/RawLemmas.lean @@ -2305,7 +2305,7 @@ theorem get?_inter_of_not_mem_right [LawfulBEq α] (h₁ : m₁.WF) (h₂ : m₂ simp_to_raw using Raw₀.get?_inter_of_contains_eq_false_right /- get -/ -theorem get_inter [LawfulBEq α] (h₁ : m₁.WF) (h₂ : m₂.WF) +@[simp] theorem get_inter [LawfulBEq α] (h₁ : m₁.WF) (h₂ : m₂.WF) {k : α} {h_mem: k ∈ m₁ ∩ m₂} : (m₁ ∩ m₂).get k h_mem = m₁.get k ((mem_inter_iff h₁ h₂).1 h_mem).1 := by @@ -2411,7 +2411,7 @@ theorem getKey?_inter_of_not_mem_left [EquivBEq α] [LawfulHashable α] simp_to_raw using Raw₀.getKey?_inter_of_contains_eq_false_left /- getKey -/ -theorem getKey_inter [EquivBEq α] [LawfulHashable α] (h₁ : m₁.WF) (h₂ : m₂.WF) +@[simp] theorem getKey_inter [EquivBEq α] [LawfulHashable α] (h₁ : m₁.WF) (h₂ : m₂.WF) {k : α} {h_mem : k ∈ m₁ ∩ m₂} : (m₁ ∩ m₂).getKey k h_mem = m₁.getKey k ((mem_inter_iff h₁ h₂).1 h_mem).1 := by @@ -2592,7 +2592,7 @@ theorem get?_inter_of_not_mem_right [EquivBEq α] [LawfulHashable α] (h₁ : m simp_to_raw using Raw₀.Const.get?_inter_of_contains_eq_false_right /- get -/ -theorem get_inter [EquivBEq α] [LawfulHashable α] (h₁ : m₁.WF) (h₂ : m₂.WF) +@[simp] theorem get_inter [EquivBEq α] [LawfulHashable α] (h₁ : m₁.WF) (h₂ : m₂.WF) {k : α} {h_mem : k ∈ m₁ ∩ m₂} : Const.get (m₁ ∩ m₂) k h_mem = Const.get m₁ k ((mem_inter_iff h₁ h₂).1 h_mem).1 := by diff --git a/src/Std/Data/HashMap/Lemmas.lean b/src/Std/Data/HashMap/Lemmas.lean index 103cd25f0455..b8ccd68e7b5b 100644 --- a/src/Std/Data/HashMap/Lemmas.lean +++ b/src/Std/Data/HashMap/Lemmas.lean @@ -1510,7 +1510,7 @@ theorem get?_inter_of_not_mem_right [EquivBEq α] [LawfulHashable α] @DHashMap.Const.get?_inter_of_not_mem_right _ _ _ _ m₁.inner m₂.inner _ _ k not_mem /- get -/ -theorem get_inter [EquivBEq α] [LawfulHashable α] +@[simp] theorem get_inter [EquivBEq α] [LawfulHashable α] {k : α} {h_mem : k ∈ m₁ ∩ m₂} : (m₁ ∩ m₂).get k h_mem = m₁.get k ((mem_inter_iff.1 h_mem).1) := @@ -1580,7 +1580,7 @@ theorem getKey?_inter_of_not_mem_left [EquivBEq α] [LawfulHashable α] @DHashMap.getKey?_inter_of_not_mem_left _ _ _ _ m₁.inner m₂.inner _ _ k not_mem /- getKey -/ -theorem getKey_inter [EquivBEq α] [LawfulHashable α] +@[simp] theorem getKey_inter [EquivBEq α] [LawfulHashable α] {k : α} {h_mem : k ∈ m₁ ∩ m₂} : (m₁ ∩ m₂).getKey k h_mem = m₁.getKey k ((mem_inter_iff.1 h_mem).1) := diff --git a/src/Std/Data/HashMap/RawLemmas.lean b/src/Std/Data/HashMap/RawLemmas.lean index c15e2338315b..3e83d0953089 100644 --- a/src/Std/Data/HashMap/RawLemmas.lean +++ b/src/Std/Data/HashMap/RawLemmas.lean @@ -1486,7 +1486,7 @@ theorem get?_inter_of_not_mem_right [EquivBEq α] [LawfulHashable α] (h₁ : m @DHashMap.Raw.Const.get?_inter_of_not_mem_right _ _ _ _ m₁.inner m₂.inner _ _ h₁.out h₂.out k not_mem /- get -/ -theorem get_inter [EquivBEq α] [LawfulHashable α] (h₁ : m₁.WF) (h₂ : m₂.WF) +@[simp] theorem get_inter [EquivBEq α] [LawfulHashable α] (h₁ : m₁.WF) (h₂ : m₂.WF) {k : α} {h_mem : k ∈ m₁ ∩ m₂} : get (m₁ ∩ m₂) k h_mem = get m₁ k ((mem_inter_iff h₁ h₂).1 h_mem).1 := @@ -1558,7 +1558,7 @@ theorem getKey?_inter_of_not_mem_left [EquivBEq α] [LawfulHashable α] @DHashMap.Raw.getKey?_inter_of_not_mem_left _ _ _ _ m₁.inner m₂.inner _ _ h₁.out h₂.out k not_mem /- getKey -/ -theorem getKey_inter [EquivBEq α] [LawfulHashable α] (h₁ : m₁.WF) (h₂ : m₂.WF) +@[simp] theorem getKey_inter [EquivBEq α] [LawfulHashable α] (h₁ : m₁.WF) (h₂ : m₂.WF) {k : α} {h_mem : k ∈ m₁ ∩ m₂} : (m₁ ∩ m₂).getKey k h_mem = m₁.getKey k ((mem_inter_iff h₁ h₂).1 h_mem).1 := diff --git a/src/Std/Data/HashSet/Lemmas.lean b/src/Std/Data/HashSet/Lemmas.lean index 02a6d11679d1..388fde9a74d2 100644 --- a/src/Std/Data/HashSet/Lemmas.lean +++ b/src/Std/Data/HashSet/Lemmas.lean @@ -901,7 +901,7 @@ theorem get?_inter_of_not_mem_right [EquivBEq α] [LawfulHashable α] @HashMap.getKey?_inter_of_not_mem_right _ _ _ _ m₁.inner m₂.inner _ _ k not_mem /- get -/ -theorem get_inter [EquivBEq α] [LawfulHashable α] +@[simp] theorem get_inter [EquivBEq α] [LawfulHashable α] {k : α} {h_mem : k ∈ m₁ ∩ m₂} : (m₁ ∩ m₂).get k h_mem = m₁.get k ((mem_inter_iff.1 h_mem).1) := diff --git a/src/Std/Data/HashSet/RawLemmas.lean b/src/Std/Data/HashSet/RawLemmas.lean index eaeff1bfe46a..39cfe31d445e 100644 --- a/src/Std/Data/HashSet/RawLemmas.lean +++ b/src/Std/Data/HashSet/RawLemmas.lean @@ -949,7 +949,7 @@ theorem get?_inter_of_not_mem_right [EquivBEq α] [LawfulHashable α] @HashMap.Raw.getKey?_inter_of_not_mem_right _ _ _ _ m₁.inner m₂.inner _ _ h₁.out h₂.out k not_mem /- get -/ -theorem get_inter [EquivBEq α] [LawfulHashable α] (h₁ : m₁.WF) (h₂ : m₂.WF) +@[simp] theorem get_inter [EquivBEq α] [LawfulHashable α] (h₁ : m₁.WF) (h₂ : m₂.WF) {k : α} {h_mem : k ∈ m₁ ∩ m₂} : (m₁ ∩ m₂).get k h_mem = m₁.get k ((mem_inter_iff h₁ h₂).1 h_mem).1 := From 75885a4b71c20b1f3e0c2ae1122b2c5c79cddd1d Mon Sep 17 00:00:00 2001 From: Wojciech Rozowski Date: Fri, 14 Nov 2025 20:28:42 +0100 Subject: [PATCH 64/88] Remove unnecessary List prefix --- src/Std/Data/DHashMap/Internal/RawLemmas.lean | 112 ++++++------ src/Std/Data/DHashMap/Internal/WF.lean | 2 +- src/Std/Data/Internal/List/Associative.lean | 164 +++++++++--------- 3 files changed, 139 insertions(+), 139 deletions(-) diff --git a/src/Std/Data/DHashMap/Internal/RawLemmas.lean b/src/Std/Data/DHashMap/Internal/RawLemmas.lean index 4176bf129afa..87ca92ee2f26 100644 --- a/src/Std/Data/DHashMap/Internal/RawLemmas.lean +++ b/src/Std/Data/DHashMap/Internal/RawLemmas.lean @@ -2482,14 +2482,14 @@ theorem get_union_of_contains_right [LawfulBEq α] (h₁ : m₁.val.WF) (h₂ : revert contains_right simp_to_model [union, get, contains] intro contains_right - apply List.getValueCast_insertList_of_contains_right + apply getValueCast_insertList_of_contains_right all_goals wf_trivial theorem get_union_of_contains_eq_false_left [LawfulBEq α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} (contains_eq_false : m₁.contains k = false) {h'} : (m₁.union m₂).get k h' = m₂.get k (contains_of_contains_union_of_contains_eq_false_left h₁ h₂ h' contains_eq_false) := by revert contains_eq_false - simp_to_model [union, contains, get] using List.getValueCast_insertList_of_contains_eq_false_left + simp_to_model [union, contains, get] using getValueCast_insertList_of_contains_eq_false_left theorem get_union_of_contains_eq_false_right [LawfulBEq α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} (contains_eq_false : m₂.contains k = false) {h'} : @@ -2513,7 +2513,7 @@ theorem getD_union_of_contains_eq_false_left [LawfulBEq α] (h₁ : m₁.val.WF) revert contains_eq_false simp_to_model [union, contains, getD] intro contains_eq_false - apply List.getValueCastD_insertList_of_contains_eq_false_left + apply getValueCastD_insertList_of_contains_eq_false_left all_goals wf_trivial theorem getD_union_of_contains_eq_false_right [LawfulBEq α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) @@ -2540,7 +2540,7 @@ theorem get!_union_of_contains_eq_false_left [LawfulBEq α] (h₁ : m₁.val.WF) revert contains_eq_false simp_to_model [union, contains, get!] intro contains_eq_false - apply List.getValueCastD_insertList_of_contains_eq_false_left + apply getValueCastD_insertList_of_contains_eq_false_left all_goals wf_trivial theorem get!_union_of_contains_eq_false_right [LawfulBEq α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) @@ -2590,7 +2590,7 @@ theorem getKey_union_of_contains_eq_false_left [EquivBEq α] [LawfulHashable α] {k : α} (contains_eq_false : m₁.contains k = false) {h'} : (m₁.union m₂).getKey k h' = m₂.getKey k (contains_of_contains_union_of_contains_eq_false_left h₁ h₂ h' contains_eq_false) := by revert contains_eq_false - simp_to_model [union, contains, getKey] using List.getKey_insertList_of_contains_eq_false_left + simp_to_model [union, contains, getKey] using getKey_insertList_of_contains_eq_false_left theorem getKey_union_of_contains_eq_false_right [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} (contains_eq_false : m₂.contains k = false) {h'} : @@ -2615,7 +2615,7 @@ theorem getKeyD_union_of_contains_eq_false_left [EquivBEq α] [LawfulHashable α revert h' simp_to_model [contains, union, getKeyD] intro h' - apply List.getKeyD_insertList_of_contains_eq_false_left + apply getKeyD_insertList_of_contains_eq_false_left . wf_trivial . wf_trivial . exact h' @@ -2626,7 +2626,7 @@ theorem getKeyD_union_of_contains_eq_false_right [EquivBEq α] [LawfulHashable revert h' simp_to_model [contains, union, getKeyD] intro h' - apply List.getKeyD_insertList_of_contains_eq_false_right h' + apply getKeyD_insertList_of_contains_eq_false_right h' /- getKey! -/ theorem getKey!_union [EquivBEq α] [LawfulHashable α] [Inhabited α] @@ -2640,14 +2640,14 @@ theorem getKey!_union_of_contains_eq_false_left [Inhabited α] (h' : m₁.contains k = false) : (m₁.union m₂).getKey! k = m₂.getKey! k := by revert h' - simp_to_model [getKey!, contains, union] using List.getKeyD_insertList_of_contains_eq_false_left + simp_to_model [getKey!, contains, union] using getKeyD_insertList_of_contains_eq_false_left theorem getKey!_union_of_contains_eq_false_right [Inhabited α] [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} (h' : m₂.contains k = false) : (m₁.union m₂).getKey! k = m₁.getKey! k := by revert h' - simp_to_model [contains, union, getKey!] using List.getKeyD_insertList_of_contains_eq_false_right + simp_to_model [contains, union, getKey!] using getKeyD_insertList_of_contains_eq_false_right /- size -/ theorem size_union_of_not_mem [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) @@ -2683,13 +2683,13 @@ variable {β : Type v} {m₁ m₂ : Raw₀ α (fun _ => β)} /- get? -/ theorem get?_union [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} : Const.get? (m₁.union m₂) k = (Const.get? m₂ k).or (Const.get? m₁ k) := by - simp_to_model [union, Const.get?] using List.getValue?_insertList + simp_to_model [union, Const.get?] using getValue?_insertList theorem get?_union_of_contains_eq_false_left [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} (contains_eq_false : m₁.contains k = false) : Const.get? (m₁.union m₂) k = Const.get? m₂ k := by revert contains_eq_false - simp_to_model [union, contains, Const.get?] using List.getValue?_insertList_of_contains_eq_false_left + simp_to_model [union, contains, Const.get?] using getValue?_insertList_of_contains_eq_false_left theorem get?_union_of_contains_eq_false_right [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} (contains_eq_false : m₂.contains k = false) : @@ -2697,7 +2697,7 @@ theorem get?_union_of_contains_eq_false_right [EquivBEq α] [LawfulHashable α] revert contains_eq_false simp_to_model [union, Const.get?, contains] intro contains_eq_false - apply List.getValue?_insertList_of_contains_eq_false_right contains_eq_false + apply getValue?_insertList_of_contains_eq_false_right contains_eq_false /- get -/ theorem get_union_of_contains_right [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) @@ -2706,14 +2706,14 @@ theorem get_union_of_contains_right [EquivBEq α] [LawfulHashable α] (h₁ : m revert h simp_to_model [union, contains, Const.get] intro h - apply List.getValue_insertList_of_contains_right + apply getValue_insertList_of_contains_right all_goals wf_trivial theorem get_union_of_contains_eq_false_left [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} (contains_eq_false : m₁.contains k = false) {h'} : Const.get (m₁.union m₂) k h' = Const.get m₂ k (contains_of_contains_union_of_contains_eq_false_left h₁ h₂ h' contains_eq_false) := by revert contains_eq_false - simp_to_model [union, contains, Const.get] using List.getValue_insertList_of_contains_eq_false_left + simp_to_model [union, contains, Const.get] using getValue_insertList_of_contains_eq_false_left theorem get_union_of_contains_eq_false_right [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} (contains_eq_false : m₂.contains k = false) {h'} : @@ -2721,7 +2721,7 @@ theorem get_union_of_contains_eq_false_right [EquivBEq α] [LawfulHashable α] ( revert contains_eq_false simp_to_model [union, Const.get, contains] intro contains_eq_false - apply List.getValue_insertList_of_contains_eq_false_right contains_eq_false + apply getValue_insertList_of_contains_eq_false_right contains_eq_false /- getD -/ theorem getD_union [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} {fallback : β} : @@ -2732,7 +2732,7 @@ theorem getD_union_of_contains_eq_false_left [EquivBEq α] [LawfulHashable α] ( {k : α} {fallback : β} (contains_eq_false : m₁.contains k = false) : Const.getD (m₁.union m₂) k fallback = Const.getD m₂ k fallback := by revert contains_eq_false - simp_to_model [union, contains, Const.getD] using List.getValueD_insertList_of_contains_eq_false_left + simp_to_model [union, contains, Const.getD] using getValueD_insertList_of_contains_eq_false_left theorem getD_union_of_contains_eq_false_right [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} {fallback : β} (contains_eq_false : m₂.contains k = false) : @@ -2740,7 +2740,7 @@ theorem getD_union_of_contains_eq_false_right [EquivBEq α] [LawfulHashable α] revert contains_eq_false simp_to_model [union, Const.getD, contains] intro contains_eq_false - apply List.getValueD_insertList_of_contains_eq_false_right contains_eq_false + apply getValueD_insertList_of_contains_eq_false_right contains_eq_false /- get! -/ theorem get!_union [EquivBEq α] [LawfulHashable α] [Inhabited β] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} : @@ -2751,7 +2751,7 @@ theorem get!_union_of_contains_eq_false_left [EquivBEq α] [LawfulHashable α] [ {k : α} (contains_eq_false : m₁.contains k = false) : Const.get! (m₁.union m₂) k = Const.get! m₂ k := by revert contains_eq_false - simp_to_model [union, contains, Const.get!] using List.getValueD_insertList_of_contains_eq_false_left + simp_to_model [union, contains, Const.get!] using getValueD_insertList_of_contains_eq_false_left theorem get!_union_of_contains_eq_false_right [EquivBEq α] [LawfulHashable α] [Inhabited β] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} (contains_eq_false : m₂.contains k = false) : @@ -2759,7 +2759,7 @@ theorem get!_union_of_contains_eq_false_right [EquivBEq α] [LawfulHashable α] revert contains_eq_false simp_to_model [union, Const.get!, contains] intro contains_eq_false - apply List.getValueD_insertList_of_contains_eq_false_right contains_eq_false + apply getValueD_insertList_of_contains_eq_false_right contains_eq_false end Const @@ -2776,184 +2776,184 @@ theorem contains_inter [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) theorem contains_inter_iff [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} : (m₁.inter m₂).contains k ↔ m₁.contains k ∧ m₂.contains k := by - simp_to_model [inter, contains] using List.containsKey_filter_containsKey_iff + simp_to_model [inter, contains] using containsKey_filter_containsKey_iff theorem contains_inter_eq_false_of_contains_eq_false_left [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} (h : m₁.contains k = false) : (m₁.inter m₂).contains k = false := by revert h - simp_to_model [inter, contains] using List.containsKey_filter_containsKey_eq_false_of_containsKey_eq_false_left + simp_to_model [inter, contains] using containsKey_filter_containsKey_eq_false_of_containsKey_eq_false_left theorem contains_inter_eq_false_of_contains_eq_false_right [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} (h : m₂.contains k = false) : (m₁.inter m₂).contains k = false := by revert h - simp_to_model [inter, contains] using List.containsKey_filter_containsKey_eq_false_of_containsKey_eq_false_right + simp_to_model [inter, contains] using containsKey_filter_containsKey_eq_false_of_containsKey_eq_false_right /- get? -/ theorem get?_inter [LawfulBEq α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} : (m₁.inter m₂).get? k = if m₂.contains k then m₁.get? k else none := by - simp_to_model [inter, get?, contains] using List.getValueCast?_filter_containsKey + simp_to_model [inter, get?, contains] using getValueCast?_filter_containsKey theorem get?_inter_of_contains_right [LawfulBEq α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} (h : m₂.contains k) : (m₁.inter m₂).get? k = m₁.get? k := by revert h - simp_to_model [inter, get?, contains] using List.getValueCast?_filter_containsKey_of_containsKey_right + simp_to_model [inter, get?, contains] using getValueCast?_filter_containsKey_of_containsKey_right theorem get?_inter_of_contains_eq_false_left [LawfulBEq α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} (h : m₁.contains k = false) : (m₁.inter m₂).get? k = none := by revert h - simp_to_model [inter, get?, contains] using List.getValueCast?_filter_containsKey_of_containsKey_eq_false_left + simp_to_model [inter, get?, contains] using getValueCast?_filter_containsKey_of_containsKey_eq_false_left theorem get?_inter_of_contains_eq_false_right [LawfulBEq α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} (h : m₂.contains k = false) : (m₁.inter m₂).get? k = none := by revert h - simp_to_model [inter, get?, contains] using List.getValueCast?_filter_containsKey_of_containsKey_eq_false_right + simp_to_model [inter, get?, contains] using getValueCast?_filter_containsKey_of_containsKey_eq_false_right /- get -/ @[simp] theorem get_inter [LawfulBEq α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} {h_contains : (m₁.inter m₂).contains k} : (m₁.inter m₂).get k h_contains = m₁.get k ((contains_inter_iff h₁ h₂).1 h_contains).1 := by - simp_to_model [inter, get, contains] using List.getValueCast_filter_containsKey + simp_to_model [inter, get, contains] using getValueCast_filter_containsKey /- getD -/ theorem getD_inter [LawfulBEq α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} {fallback : β k} : (m₁.inter m₂).getD k fallback = if m₂.contains k then m₁.getD k fallback else fallback := by - simp_to_model [inter, getD, contains] using List.getValueCastD_filter_containsKey + simp_to_model [inter, getD, contains] using getValueCastD_filter_containsKey theorem getD_inter_of_contains_right [LawfulBEq α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} {fallback : β k} (h : m₂.contains k) : (m₁.inter m₂).getD k fallback = m₁.getD k fallback := by revert h - simp_to_model [inter, getD, contains] using List.getValueCastD_filter_containsKey_of_containsKey_right + simp_to_model [inter, getD, contains] using getValueCastD_filter_containsKey_of_containsKey_right theorem getD_inter_of_contains_eq_false_right [LawfulBEq α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} {fallback : β k} (h : m₂.contains k = false) : (m₁.inter m₂).getD k fallback = fallback := by revert h - simp_to_model [inter, getD, contains] using List.getValueCastD_filter_containsKey_of_containsKey_eq_false_right + simp_to_model [inter, getD, contains] using getValueCastD_filter_containsKey_of_containsKey_eq_false_right theorem getD_inter_of_contains_eq_false_left [LawfulBEq α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} {fallback : β k} (h : m₁.contains k = false) : (m₁.inter m₂).getD k fallback = fallback := by revert h - simp_to_model [inter, getD, contains] using List.getValueCastD_filter_containsKey_of_containsKey_eq_false_left + simp_to_model [inter, getD, contains] using getValueCastD_filter_containsKey_of_containsKey_eq_false_left /- get! -/ theorem get!_inter [LawfulBEq α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} [Inhabited (β k)] : (m₁.inter m₂).get! k = if m₂.contains k then m₁.get! k else default := by - simp_to_model [inter, get!, contains] using List.getValueCastD_filter_containsKey + simp_to_model [inter, get!, contains] using getValueCastD_filter_containsKey theorem get!_inter_of_contains_right [LawfulBEq α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} [Inhabited (β k)] (h : m₂.contains k) : (m₁.inter m₂).get! k = m₁.get! k := by revert h - simp_to_model [inter, get!, contains] using List.getValueCastD_filter_containsKey_of_containsKey_right + simp_to_model [inter, get!, contains] using getValueCastD_filter_containsKey_of_containsKey_right theorem get!_inter_of_contains_eq_false_right [LawfulBEq α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} [Inhabited (β k)] (h : m₂.contains k = false) : (m₁.inter m₂).get! k = default := by revert h - simp_to_model [inter, get!, contains] using List.getValueCastD_filter_containsKey_of_containsKey_eq_false_right + simp_to_model [inter, get!, contains] using getValueCastD_filter_containsKey_of_containsKey_eq_false_right theorem get!_inter_of_contains_eq_false_left [LawfulBEq α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} [Inhabited (β k)] (h : m₁.contains k = false) : (m₁.inter m₂).get! k = default := by revert h - simp_to_model [inter, get!, contains] using List.getValueCastD_filter_containsKey_of_containsKey_eq_false_left + simp_to_model [inter, get!, contains] using getValueCastD_filter_containsKey_of_containsKey_eq_false_left /- getKey? -/ theorem getKey?_inter [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} : (m₁.inter m₂).getKey? k = if m₂.contains k then m₁.getKey? k else none := by - simp_to_model [inter, contains, getKey?] using List.getKey?_filter_containsKey + simp_to_model [inter, contains, getKey?] using getKey?_filter_containsKey theorem getKey?_inter_of_contains_right [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} (h : m₂.contains k) : (m₁.inter m₂).getKey? k = m₁.getKey? k := by revert h - simp_to_model [contains, getKey?, inter] using List.getKey?_filter_containsKey_of_containsKey_right + simp_to_model [contains, getKey?, inter] using getKey?_filter_containsKey_of_containsKey_right theorem getKey?_inter_of_contains_eq_false_right [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} (h : m₂.contains k = false) : (m₁.inter m₂).getKey? k = none := by revert h - simp_to_model [contains, getKey?, inter] using List.getKey?_filter_containsKey_of_containsKey_eq_false_right + simp_to_model [contains, getKey?, inter] using getKey?_filter_containsKey_of_containsKey_eq_false_right theorem getKey?_inter_of_contains_eq_false_left [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} (h : m₁.contains k = false) : (m₁.inter m₂).getKey? k = none := by revert h - simp_to_model [contains, getKey?, inter] using List.getKey?_filter_containsKey_of_containsKey_eq_false_left + simp_to_model [contains, getKey?, inter] using getKey?_filter_containsKey_of_containsKey_eq_false_left /- getKey -/ @[simp] theorem getKey_inter [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} {h_contains : (m₁.inter m₂).contains k} : (m₁.inter m₂).getKey k h_contains = m₁.getKey k (by simp [contains_inter_iff h₁ h₂] at h_contains; exact h_contains.1) := by - simp_to_model [inter, contains, getKey] using List.getKey_filter_containsKey + simp_to_model [inter, contains, getKey] using getKey_filter_containsKey /- getKeyD -/ theorem getKeyD_inter [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k fallback : α} : (m₁.inter m₂).getKeyD k fallback = if m₂.contains k then m₁.getKeyD k fallback else fallback := by - simp_to_model [inter, getKeyD, contains] using List.getKeyD_filter_containsKey + simp_to_model [inter, getKeyD, contains] using getKeyD_filter_containsKey theorem getKeyD_inter_of_contains_right [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k fallback : α} (h : m₂.contains k) : (m₁.inter m₂).getKeyD k fallback = m₁.getKeyD k fallback := by revert h - simp_to_model [inter, getKeyD, contains] using List.getKeyD_filter_containsKey_of_containsKey_right + simp_to_model [inter, getKeyD, contains] using getKeyD_filter_containsKey_of_containsKey_right theorem getKeyD_inter_of_contains_eq_false_right [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k fallback : α} (h : m₂.contains k = false) : (m₁.inter m₂).getKeyD k fallback = fallback := by revert h - simp_to_model [inter, getKeyD, contains] using List.getKeyD_filter_containsKey_of_containsKey_eq_false_right + simp_to_model [inter, getKeyD, contains] using getKeyD_filter_containsKey_of_containsKey_eq_false_right theorem getKeyD_inter_of_contains_eq_false_left [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k fallback : α} (h : m₁.contains k = false) : (m₁.inter m₂).getKeyD k fallback = fallback := by revert h - simp_to_model [inter, getKeyD, contains] using List.getKeyD_filter_containsKey_of_containsKey_eq_false_left + simp_to_model [inter, getKeyD, contains] using getKeyD_filter_containsKey_of_containsKey_eq_false_left /- getKey! -/ theorem getKey!_inter [EquivBEq α] [LawfulHashable α] [Inhabited α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} [Inhabited α] : (m₁.inter m₂).getKey! k = if m₂.contains k then m₁.getKey! k else default := by - simp_to_model [inter, getKey!, contains] using List.getKeyD_filter_containsKey + simp_to_model [inter, getKey!, contains] using getKeyD_filter_containsKey theorem getKey!_inter_of_contains_right [EquivBEq α] [LawfulHashable α] [Inhabited α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} (h : m₂.contains k) : (m₁.inter m₂).getKey! k = m₁.getKey! k := by revert h - simp_to_model [inter, getKey!, contains] using List.getKeyD_filter_containsKey_of_containsKey_right + simp_to_model [inter, getKey!, contains] using getKeyD_filter_containsKey_of_containsKey_right theorem getKey!_inter_of_contains_eq_false_right [EquivBEq α] [LawfulHashable α] [Inhabited α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} (h : m₂.contains k = false) : (m₁.inter m₂).getKey! k = default := by revert h - simp_to_model [inter, getKey!, contains] using List.getKeyD_filter_containsKey_of_containsKey_eq_false_right + simp_to_model [inter, getKey!, contains] using getKeyD_filter_containsKey_of_containsKey_eq_false_right theorem getKey!_inter_of_contains_eq_false_left [EquivBEq α] [LawfulHashable α] [Inhabited α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} (h : m₁.contains k = false) : (m₁.inter m₂).getKey! k = default := by revert h - simp_to_model [inter, getKey!, contains] using List.getKeyD_filter_containsKey_of_containsKey_eq_false_left + simp_to_model [inter, getKey!, contains] using getKeyD_filter_containsKey_of_containsKey_eq_false_left /- size -/ theorem size_inter_le_size_left [EquivBEq α] [LawfulHashable α] @@ -2971,19 +2971,19 @@ theorem size_inter_eq_size_left [EquivBEq α] [LawfulHashable α] (h : ∀ (a : α), m₁.contains a → m₂.contains a) : (m₁.inter m₂).1.size = m₁.1.size := by revert h - simp_to_model [inter, size, contains] using List.length_filter_containsKey_eq_length_left + simp_to_model [inter, size, contains] using length_filter_containsKey_eq_length_left theorem size_inter_eq_size_right [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) (h : ∀ (a : α), m₂.contains a → m₁.contains a) : (m₁.inter m₂).1.size = m₂.1.size := by revert h - simp_to_model [inter, size, contains] using List.length_filter_containsKey_of_length_right + simp_to_model [inter, size, contains] using length_filter_containsKey_of_length_right theorem size_add_size_eq_size_union_add_size_inter [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) : m₁.1.size + m₂.1.size = (m₁.union m₂).1.size + (m₁.inter m₂).1.size := by - simp_to_model [union, inter, size] using List.size_add_size_eq_size_insertList_add_size_filter_containsKey + simp_to_model [union, inter, size] using size_add_size_eq_size_insertList_add_size_filter_containsKey /- isEmpty -/ @[simp] @@ -3031,7 +3031,7 @@ theorem get?_inter_of_contains_eq_false_right [EquivBEq α] [LawfulHashable α] {k : α} (h : m₂.contains k = false) : Const.get? (m₁.inter m₂) k = none := by revert h - simp_to_model [inter, Const.get?, contains] using List.getValue?_filter_of_containsKey_eq_false_right + simp_to_model [inter, Const.get?, contains] using getValue?_filter_of_containsKey_eq_false_right @[simp] theorem get_inter [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} {h_contains : (m₁.inter m₂).contains k} : @@ -3056,13 +3056,13 @@ theorem getD_inter_of_contains_eq_false_right [EquivBEq α] [LawfulHashable α] {k : α} {fallback : β} (h : m₂.contains k = false) : Const.getD (m₁.inter m₂) k fallback = fallback := by revert h - simp_to_model [inter, Const.getD, contains] using List.getValueD_filter_of_containsKey_eq_false_right + simp_to_model [inter, Const.getD, contains] using getValueD_filter_of_containsKey_eq_false_right theorem getD_inter_of_contains_eq_false_left [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} {fallback : β} (h : m₁.contains k = false) : Const.getD (m₁.inter m₂) k fallback = fallback := by revert h - simp_to_model [inter, Const.getD, contains] using List.getValueD_filter_of_containsKey_eq_false_left + simp_to_model [inter, Const.getD, contains] using getValueD_filter_of_containsKey_eq_false_left /- get! -/ theorem get!_inter [EquivBEq α] [LawfulHashable α] [Inhabited β] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) @@ -3081,13 +3081,13 @@ theorem get!_inter_of_contains_eq_false_right [EquivBEq α] [LawfulHashable α] {k : α} (h : m₂.contains k = false) : Const.get! (m₁.inter m₂) k = default := by revert h - simp_to_model [inter, Const.get!, contains] using List.getValueD_filter_of_containsKey_eq_false_right + simp_to_model [inter, Const.get!, contains] using getValueD_filter_of_containsKey_eq_false_right theorem get!_inter_of_contains_eq_false_left [EquivBEq α] [LawfulHashable α] [Inhabited β] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} (h : m₁.contains k = false) : Const.get! (m₁.inter m₂) k = default := by revert h - simp_to_model [inter, Const.get!, contains] using List.getValueD_filter_of_containsKey_eq_false_left + simp_to_model [inter, Const.get!, contains] using getValueD_filter_of_containsKey_eq_false_left end Const diff --git a/src/Std/Data/DHashMap/Internal/WF.lean b/src/Std/Data/DHashMap/Internal/WF.lean index e43ba71f48b5..ed535d9e4c38 100644 --- a/src/Std/Data/DHashMap/Internal/WF.lean +++ b/src/Std/Data/DHashMap/Internal/WF.lean @@ -1410,7 +1410,7 @@ theorem toListModel_inter [BEq α] [EquivBEq α] [Hashable α] [LawfulHashable exact toListModel_filterₘ · rw [interSmaller_eq_interSmallerₘ] exact Perm.trans (toListModel_interSmallerₘ _ _ hm₁) - (List.interSmaller_perm_filter _ _ hm₁.distinct) + (interSmaller_perm_filter _ _ hm₁.distinct) /-! # `Const.insertListₘ` -/ diff --git a/src/Std/Data/Internal/List/Associative.lean b/src/Std/Data/Internal/List/Associative.lean index 9e783c82ae7b..37f9ba2d0997 100644 --- a/src/Std/Data/Internal/List/Associative.lean +++ b/src/Std/Data/Internal/List/Associative.lean @@ -3325,7 +3325,7 @@ theorem getKeyD_insertList [BEq α] [EquivBEq α] rw [@getKey?_insertList α β _ _ l toInsert k distinct_l distinct_toInsert] simp only [Option.getD_or] -theorem List.getKeyD_insertList_of_contains_eq_false_right [BEq α] [EquivBEq α] +theorem getKeyD_insertList_of_contains_eq_false_right [BEq α] [EquivBEq α] {l toInsert : List ((a : α) × β a)} {k fallback : α} (not_contains: containsKey k toInsert = false) : List.getKeyD k (insertList l toInsert) fallback = List.getKeyD k l fallback := by @@ -3336,7 +3336,7 @@ theorem List.getKeyD_insertList_of_contains_eq_false_right [BEq α] [EquivBEq α apply getEntry?_insertList_of_contains_eq_false . exact not_contains -theorem List.getKeyD_insertList_of_contains_eq_false_left [BEq α] [EquivBEq α] +theorem getKeyD_insertList_of_contains_eq_false_left [BEq α] [EquivBEq α] {l toInsert : List ((a : α) × β a)} {k fallback : α} (distinct_l : DistinctKeys l) (distinct_toInsert : DistinctKeys toInsert) @@ -3419,7 +3419,7 @@ theorem getValueCastD_of_insertList [BEq α] [LawfulBEq α] rw [getOr] simp only [Option.getD_or, getValueCastD_eq_getValueCast?] -theorem List.getValueCast_insertList_of_contains_eq_false_left [BEq α] [LawfulBEq α] {l toInsert : List ((a : α) × β a)} {k : α} +theorem getValueCast_insertList_of_contains_eq_false_left [BEq α] [LawfulBEq α] {l toInsert : List ((a : α) × β a)} {k : α} (distinct_l : DistinctKeys l) (distinct_toInsert: DistinctKeys toInsert) (contains : containsKey k (insertList l toInsert)) @@ -3430,7 +3430,7 @@ theorem List.getValueCast_insertList_of_contains_eq_false_left [BEq α] [LawfulB rw [← getValueCast?_eq_some_getValueCast, ← getValueCast?_eq_some_getValueCast] exact getValueCast?_insertList_of_contains_eq_false_left distinct_l distinct_toInsert not_contains -theorem List.getKey_insertList_of_contains_eq_false_left [BEq α] [EquivBEq α] {l toInsert : List ((a : α) × β a)} {k : α} +theorem getKey_insertList_of_contains_eq_false_left [BEq α] [EquivBEq α] {l toInsert : List ((a : α) × β a)} {k : α} (distinct_l : DistinctKeys l) (distinct_toInsert: DistinctKeys toInsert) (contains : containsKey k (insertList l toInsert)) @@ -3441,7 +3441,7 @@ theorem List.getKey_insertList_of_contains_eq_false_left [BEq α] [EquivBEq α] simp only [← getKey?_eq_some_getKey] exact getKey?_insertList_of_contains_eq_false_left distinct_l distinct_toInsert not_contains -theorem List.getValueCast_insertList_of_contains_right [BEq α] [LawfulBEq α] {l toInsert : List ((a : α) × β a)} {k : α} +theorem getValueCast_insertList_of_contains_right [BEq α] [LawfulBEq α] {l toInsert : List ((a : α) × β a)} {k : α} (distinct_l : DistinctKeys l) (distinct_toInsert: DistinctKeys toInsert) (contains : containsKey k toInsert) {h} : @@ -3455,7 +3455,7 @@ theorem List.getValueCast_insertList_of_contains_right [BEq α] [LawfulBEq α] { rw [containsKey_eq_isSome_getEntry?] at contains rw [@Option.or_of_isSome _ (getEntry? k toInsert) (getEntry? k l) contains] -theorem List.getValueCastD_insertList_of_contains_eq_false_left [BEq α] [LawfulBEq α] {l toInsert : List ((a : α) × β a)} {k : α} {fallback : β k} +theorem getValueCastD_insertList_of_contains_eq_false_left [BEq α] [LawfulBEq α] {l toInsert : List ((a : α) × β a)} {k : α} {fallback : β k} (distinct_l : DistinctKeys l) (distinct_toInsert: DistinctKeys toInsert) (not_contains: containsKey k l = false) : @@ -3712,14 +3712,14 @@ theorem getValue?_insertListConst_of_contains_eq_false [BEq α] [PartialEquivBEq rw [containsKey_eq_contains_map_fst] simpa using not_contains -theorem List.getValue?_insertList_of_contains_eq_false_right [BEq α] [EquivBEq α] +theorem getValue?_insertList_of_contains_eq_false_right [BEq α] [EquivBEq α] {l toInsert : List ((_ : α) × β)} {k : α} (not_contains : containsKey k toInsert = false) : getValue? k (insertList l toInsert) = getValue? k l := by rw [getValue?_eq_getEntry?, getValue?_eq_getEntry?] simp only [getEntry?_insertList_of_contains_eq_false not_contains] -theorem List.getValueD_insertList_of_contains_eq_false_right [BEq α] [EquivBEq α] +theorem getValueD_insertList_of_contains_eq_false_right [BEq α] [EquivBEq α] {l toInsert : List ((_ : α) × β)} {k : α} {fallback : β} (not_contains : containsKey k toInsert = false) : getValueD k (insertList l toInsert) fallback = getValueD k l fallback := by @@ -3727,7 +3727,7 @@ theorem List.getValueD_insertList_of_contains_eq_false_right [BEq α] [EquivBEq congr 1 exact getValue?_insertList_of_contains_eq_false_right not_contains -theorem List.getValue_insertList_of_contains_eq_false_right [BEq α] [EquivBEq α] +theorem getValue_insertList_of_contains_eq_false_right [BEq α] [EquivBEq α] {l toInsert : List ((_ : α) × β)} {k : α} (not_contains : containsKey k toInsert = false) {p1 : containsKey k (insertList l toInsert) = true} @@ -3736,9 +3736,9 @@ theorem List.getValue_insertList_of_contains_eq_false_right [BEq α] [EquivBEq suffices some (getValue k (insertList l toInsert) p1) = some (getValue k l p2) from by injection this simp only [← getValue?_eq_some_getValue] - apply List.getValue?_insertList_of_contains_eq_false_right not_contains + apply getValue?_insertList_of_contains_eq_false_right not_contains -theorem List.getValue?_insertList_of_contains_eq_false_left [BEq α] [EquivBEq α] {l toInsert : List ((_ : α) × β)} {k : α} +theorem getValue?_insertList_of_contains_eq_false_left [BEq α] [EquivBEq α] {l toInsert : List ((_ : α) × β)} {k : α} (distinct_l : DistinctKeys l) (distinct_toInsert: DistinctKeys toInsert) : (containsKey k l = false) → getValue? k (insertList l toInsert) = getValue? k toInsert := by @@ -3747,7 +3747,7 @@ theorem List.getValue?_insertList_of_contains_eq_false_left [BEq α] [EquivBEq congr 1 exact getEntry?_insertList_of_contains_left_eq_false distinct_l (DistinctKeys_impl_Pairwise_distinct distinct_toInsert) not_contains -theorem List.getValueD_insertList_of_contains_eq_false_left [BEq α] [EquivBEq α] {l toInsert : List ((_ : α) × β)} {k : α} {fallback : β} +theorem getValueD_insertList_of_contains_eq_false_left [BEq α] [EquivBEq α] {l toInsert : List ((_ : α) × β)} {k : α} {fallback : β} (distinct_l : DistinctKeys l) (distinct_toInsert: DistinctKeys toInsert) : (containsKey k l = false) → getValueD k (insertList l toInsert) fallback = getValueD k toInsert fallback := by @@ -3756,7 +3756,7 @@ theorem List.getValueD_insertList_of_contains_eq_false_left [BEq α] [EquivBEq congr 1 exact getValue?_insertList_of_contains_eq_false_left distinct_l distinct_toInsert not_contains -theorem List.getValue?_insertList [BEq α] [EquivBEq α] +theorem getValue?_insertList [BEq α] [EquivBEq α] {l toInsert : List ((_ : α) × β)} {k : α} (distinct_l : DistinctKeys l) (distinct_toInsert: DistinctKeys toInsert) : @@ -3774,9 +3774,9 @@ theorem getValueD_insertList [BEq α] [EquivBEq α] simp only [getValueD_eq_getValue?] simp only [← Option.getD_or] congr 1 - apply List.getValue?_insertList distinct_l distinct_toInsert + apply getValue?_insertList distinct_l distinct_toInsert -theorem List.getValue_insertList_of_contains_eq_false_left [BEq α] [EquivBEq α] {l toInsert : List ((_ : α) × β)} {k : α} +theorem getValue_insertList_of_contains_eq_false_left [BEq α] [EquivBEq α] {l toInsert : List ((_ : α) × β)} {k : α} (distinct_l : DistinctKeys l) (distinct_toInsert: DistinctKeys toInsert) (contains : containsKey k (insertList l toInsert)) @@ -3789,7 +3789,7 @@ theorem List.getValue_insertList_of_contains_eq_false_left [BEq α] [EquivBEq α congr 1 exact getEntry?_insertList_of_contains_left_eq_false distinct_l (DistinctKeys_impl_Pairwise_distinct distinct_toInsert) not_contains -theorem List.getValue_insertList_of_contains_right [BEq α] [EquivBEq α] {l toInsert : List ((_ : α) × β)} {k : α} +theorem getValue_insertList_of_contains_right [BEq α] [EquivBEq α] {l toInsert : List ((_ : α) × β)} {k : α} (distinct_l : DistinctKeys l) (distinct_toInsert: DistinctKeys toInsert) (contains : containsKey k toInsert) {h} : @@ -5201,7 +5201,7 @@ theorem getEntry?_filter [BEq α] [EquivBEq α] intro p simp only [Option.all_guard, BEq.rfl, Bool.or_true] -theorem List.getValueCast?_filter_containsKey [BEq α] [LawfulBEq α] +theorem getValueCast?_filter_containsKey [BEq α] [LawfulBEq α] {l₁ l₂ : List ((a : α) × β a)} {k : α} (dl₁ : DistinctKeys l₁) : getValueCast? k (List.filter (fun p => containsKey p.fst l₂) l₁) = @@ -5233,7 +5233,7 @@ theorem List.getValueCast?_filter_containsKey [BEq α] [LawfulBEq α] case h_2 => rfl exact dl₁ -theorem List.getKey?_filter_containsKey [BEq α] [EquivBEq α] +theorem getKey?_filter_containsKey [BEq α] [EquivBEq α] {l₁ l₂ : List ((a : α) × β a)} {k : α} (dl₁ : DistinctKeys l₁) : getKey? k (List.filter (fun p => containsKey p.fst l₂) l₁) = @@ -5267,67 +5267,67 @@ theorem List.getKey?_filter_containsKey [BEq α] [EquivBEq α] case h_2 => rfl · exact dl₁ -theorem List.getKey?_filter_containsKey_of_containsKey_right [BEq α] [EquivBEq α] +theorem getKey?_filter_containsKey_of_containsKey_right [BEq α] [EquivBEq α] {l₁ l₂ : List ((a : α) × β a)} {k : α} (dl₁ : DistinctKeys l₁) : containsKey k l₂ → getKey? k (List.filter (fun p => containsKey p.fst l₂) l₁) = getKey? k l₁ := by intro h - rw [List.getKey?_filter_containsKey, h] + rw [getKey?_filter_containsKey, h] · simp only [↓reduceIte] · exact dl₁ -theorem List.getKey?_filter_containsKey_of_containsKey_eq_false_left [BEq α] [EquivBEq α] +theorem getKey?_filter_containsKey_of_containsKey_eq_false_left [BEq α] [EquivBEq α] {l₁ l₂ : List ((a : α) × β a)} {k : α} (dl₁ : DistinctKeys l₁) : containsKey k l₁ = false → getKey? k (List.filter (fun p => containsKey p.fst l₂) l₁) = none := by intro h - rw [List.getKey?_filter_containsKey] + rw [getKey?_filter_containsKey] · split · rw [getKey?_eq_none h] · rfl · exact dl₁ -theorem List.getKey?_filter_containsKey_of_containsKey_eq_false_right [BEq α] [EquivBEq α] +theorem getKey?_filter_containsKey_of_containsKey_eq_false_right [BEq α] [EquivBEq α] {l₁ l₂ : List ((a : α) × β a)} {k : α} (dl₁ : DistinctKeys l₁) : containsKey k l₂ = false → getKey? k (List.filter (fun p => containsKey p.fst l₂) l₁) = none := by intro h - rw [List.getKey?_filter_containsKey, h] + rw [getKey?_filter_containsKey, h] · simp only [Bool.false_eq_true, ↓reduceIte] · exact dl₁ -theorem List.getValueCast?_filter_containsKey_of_containsKey_right [BEq α] [LawfulBEq α] +theorem getValueCast?_filter_containsKey_of_containsKey_right [BEq α] [LawfulBEq α] {l₁ l₂ : List ((a : α) × β a)} {k : α} (dl₁ : DistinctKeys l₁) : containsKey k l₂ → getValueCast? k (List.filter (fun p => containsKey p.fst l₂) l₁) = getValueCast? k l₁ := by intro h - rw [List.getValueCast?_filter_containsKey, h] + rw [getValueCast?_filter_containsKey, h] . simp only [↓reduceIte] . exact dl₁ -theorem List.getValueCast?_filter_containsKey_of_containsKey_eq_false_left [BEq α] [LawfulBEq α] +theorem getValueCast?_filter_containsKey_of_containsKey_eq_false_left [BEq α] [LawfulBEq α] {l₁ l₂ : List ((a : α) × β a)} {k : α} (dl₁ : DistinctKeys l₁) : containsKey k l₁ = false → getValueCast? k (List.filter (fun p => containsKey p.fst l₂) l₁) = none := by intro h - rw [List.getValueCast?_filter_containsKey] + rw [getValueCast?_filter_containsKey] . split . rw [getValueCast?_eq_none h] . rfl . exact dl₁ -theorem List.getValueCast?_filter_containsKey_of_containsKey_eq_false_right [BEq α] [LawfulBEq α] +theorem getValueCast?_filter_containsKey_of_containsKey_eq_false_right [BEq α] [LawfulBEq α] {l₁ l₂ : List ((a : α) × β a)} {k : α} (dl₁ : DistinctKeys l₁) : containsKey k l₂ = false → getValueCast? k (List.filter (fun p => containsKey p.fst l₂) l₁) = none := by intro h - rw [List.getValueCast?_filter_containsKey, h] + rw [getValueCast?_filter_containsKey, h] . simp only [Bool.false_eq_true, ↓reduceIte] . exact dl₁ -theorem List.getEntry?_filter_containsKey [BEq α] [EquivBEq α] +theorem getEntry?_filter_containsKey [BEq α] [EquivBEq α] {l₁ l₂ : List ((a : α) × β a)} {k : α} (dl₁ : DistinctKeys l₁) : getEntry? k (List.filter (fun p => containsKey p.fst l₂) l₁) = @@ -5358,34 +5358,34 @@ theorem List.getEntry?_filter_containsKey [BEq α] [EquivBEq α] case h_2 => rfl · exact dl₁ -theorem List.getEntry?_filter_containsKey_of_containsKey_right [BEq α] [EquivBEq α] +theorem getEntry?_filter_containsKey_of_containsKey_right [BEq α] [EquivBEq α] {l₁ l₂ : List ((a : α) × β a)} {k : α} (dl₁ : DistinctKeys l₁) : containsKey k l₂ → getEntry? k (List.filter (fun p => containsKey p.fst l₂) l₁) = getEntry? k l₁ := by intro h - rw [List.getEntry?_filter_containsKey, h] + rw [getEntry?_filter_containsKey, h] · simp only [↓reduceIte] · exact dl₁ -theorem List.getEntry?_filter_containsKey_of_containsKey_eq_false_left [BEq α] [EquivBEq α] +theorem getEntry?_filter_containsKey_of_containsKey_eq_false_left [BEq α] [EquivBEq α] {l₁ l₂ : List ((a : α) × β a)} {k : α} (dl₁ : DistinctKeys l₁) : containsKey k l₁ = false → getEntry? k (List.filter (fun p => containsKey p.fst l₂) l₁) = none := by intro h - rw [List.getEntry?_filter_containsKey] + rw [getEntry?_filter_containsKey] · split · apply (@getEntry?_eq_none _ _ _ l₁ k).2 exact h · rfl · exact dl₁ -theorem List.getEntry?_filter_containsKey_of_containsKey_eq_false_right [BEq α] [EquivBEq α] +theorem getEntry?_filter_containsKey_of_containsKey_eq_false_right [BEq α] [EquivBEq α] {l₁ l₂ : List ((a : α) × β a)} {k : α} (dl₁ : DistinctKeys l₁) : containsKey k l₂ = false → getEntry? k (List.filter (fun p => containsKey p.fst l₂) l₁) = none := by intro h - rw [List.getEntry?_filter_containsKey, h] + rw [getEntry?_filter_containsKey, h] · simp only [Bool.false_eq_true, ↓reduceIte] · exact dl₁ @@ -5422,56 +5422,56 @@ theorem containsKey_filter_containsKey [BEq α] [EquivBEq α] {l₁ l₂ : List apply List.beq_of_getEntry?_eq_some heq . exact hl₁ -theorem List.containsKey_filter_containsKey_iff [BEq α] [EquivBEq α] {l₁ l₂ : List ((a : α) × β a)} {hl₁ : DistinctKeys l₁} {k : α} : +theorem containsKey_filter_containsKey_iff [BEq α] [EquivBEq α] {l₁ l₂ : List ((a : α) × β a)} {hl₁ : DistinctKeys l₁} {k : α} : (containsKey k (List.filter (fun p => containsKey p.fst l₂) l₁)) ↔ (containsKey k l₁ ∧ containsKey k l₂) := by rw [containsKey_filter_containsKey] . simp . exact hl₁ -theorem List.getValueCast_filter_containsKey [BEq α] [LawfulBEq α] +theorem getValueCast_filter_containsKey [BEq α] [LawfulBEq α] {l₁ l₂ : List ((a : α) × β a)} {k : α} (dl₁ : DistinctKeys l₁) {h₁ h₂} : getValueCast k (List.filter (fun p => containsKey p.fst l₂) l₁) h₁ = getValueCast k l₁ h₂ := by suffices some (getValueCast k (List.filter (fun p => containsKey p.fst l₂) l₁) h₁) = some (getValueCast k l₁ h₂) by injections simp only [← getValueCast?_eq_some_getValueCast] - apply List.getValueCast?_filter_containsKey_of_containsKey_right dl₁ - . rw [List.containsKey_filter_containsKey_iff] at h₁ + apply getValueCast?_filter_containsKey_of_containsKey_right dl₁ + . rw [containsKey_filter_containsKey_iff] at h₁ . exact h₁.2 . exact dl₁ -theorem List.getEntry_filter_containsKey [BEq α] [EquivBEq α] +theorem getEntry_filter_containsKey [BEq α] [EquivBEq α] {l₁ l₂ : List ((a : α) × β a)} {k : α} (dl₁ : DistinctKeys l₁) {h₁ h₂} : getEntry k (List.filter (fun p => containsKey p.fst l₂) l₁) h₁ = getEntry k l₁ h₂ := by suffices some (getEntry k (List.filter (fun p => containsKey p.fst l₂) l₁) h₁) = some (getEntry k l₁ h₂) by injections simp only [← getEntry?_eq_some_getEntry] - apply List.getEntry?_filter_containsKey_of_containsKey_right dl₁ - · rw [List.containsKey_filter_containsKey_iff] at h₁ + apply getEntry?_filter_containsKey_of_containsKey_right dl₁ + · rw [containsKey_filter_containsKey_iff] at h₁ · exact h₁.2 · exact dl₁ -theorem List.getKey_filter_containsKey [BEq α] [EquivBEq α] +theorem getKey_filter_containsKey [BEq α] [EquivBEq α] {l₁ l₂ : List ((a : α) × β a)} {k : α} (dl₁ : DistinctKeys l₁) {h₁ h₂} : getKey k (List.filter (fun p => containsKey p.fst l₂) l₁) h₁ = getKey k l₁ h₂ := by suffices some (getKey k (List.filter (fun p => containsKey p.fst l₂) l₁) h₁) = some (getKey k l₁ h₂) by injections simp only [← getKey?_eq_some_getKey] - apply List.getKey?_filter_containsKey_of_containsKey_right dl₁ - · rw [List.containsKey_filter_containsKey_iff] at h₁ + apply getKey?_filter_containsKey_of_containsKey_right dl₁ + · rw [containsKey_filter_containsKey_iff] at h₁ · exact h₁.2 · exact dl₁ -theorem List.containsKey_filter_containsKey_eq_false_of_containsKey_eq_false_left [BEq α] [EquivBEq α] {l₁ l₂ : List ((a : α) × β a)} {hl₁ : DistinctKeys l₁} {k : α} : +theorem containsKey_filter_containsKey_eq_false_of_containsKey_eq_false_left [BEq α] [EquivBEq α] {l₁ l₂ : List ((a : α) × β a)} {hl₁ : DistinctKeys l₁} {k : α} : containsKey k l₁ = false → containsKey k (List.filter (fun p => containsKey p.fst l₂) l₁) = false := by intro h rw [containsKey_filter_containsKey] . simp [h] . exact hl₁ -theorem List.containsKey_filter_containsKey_eq_false_of_containsKey_eq_false_right [BEq α] [EquivBEq α] {l₁ l₂ : List ((a : α) × β a)} {hl₁ : DistinctKeys l₁} {k : α} : +theorem containsKey_filter_containsKey_eq_false_of_containsKey_eq_false_right [BEq α] [EquivBEq α] {l₁ l₂ : List ((a : α) × β a)} {hl₁ : DistinctKeys l₁} {k : α} : containsKey k l₂ = false → containsKey k (List.filter (fun p => containsKey p.fst l₂) l₁) = false := by intro h rw [containsKey_filter_containsKey] @@ -5501,7 +5501,7 @@ theorem getValue?_filter_containsKey_eq_if_containsKey {β : Type v} [BEq α] [E case isFalse h => simp only [Option.map_eq_none_iff, getEntry?_eq_none] simp only [Bool.not_eq_true] at h - apply List.containsKey_filter_containsKey_eq_false_of_containsKey_eq_false_right + apply containsKey_filter_containsKey_eq_false_of_containsKey_eq_false_right . exact dl₁ . exact h @@ -5527,7 +5527,7 @@ theorem getValue?_filter_of_not_contains {β : Type v} [BEq α] [EquivBEq α] · rfl · exact dl₁ -theorem List.getValue?_filter_of_containsKey_eq_false_right {β : Type v} [BEq α] [EquivBEq α] +theorem getValue?_filter_of_containsKey_eq_false_right {β : Type v} [BEq α] [EquivBEq α] {l₁ l₂ : List ((_ : α) × β)} {k : α} (dl₁ : DistinctKeys l₁) : containsKey k l₂ = false → getValue? k (List.filter (fun p => containsKey p.fst l₂) l₁) = none := by @@ -5536,7 +5536,7 @@ theorem List.getValue?_filter_of_containsKey_eq_false_right {β : Type v} [BEq · simp only [Bool.false_eq_true, ↓reduceIte] · exact dl₁ -theorem List.getValueCastD_filter_containsKey [BEq α] [LawfulBEq α] +theorem getValueCastD_filter_containsKey [BEq α] [LawfulBEq α] {l₁ l₂ : List ((a : α) × β a)} {k : α} {fallback : β k} (dl₁ : DistinctKeys l₁) : getValueCastD k (List.filter (fun p => containsKey p.fst l₂) l₁) fallback = @@ -5545,14 +5545,14 @@ theorem List.getValueCastD_filter_containsKey [BEq α] [LawfulBEq α] case isTrue h => rw [getValueCastD_eq_getValueCast?, getValueCastD_eq_getValueCast?] congr 1 - apply List.getValueCast?_filter_containsKey_of_containsKey_right dl₁ h + apply getValueCast?_filter_containsKey_of_containsKey_right dl₁ h case isFalse h => apply getValueCastD_eq_fallback - apply List.containsKey_filter_containsKey_eq_false_of_containsKey_eq_false_right + apply containsKey_filter_containsKey_eq_false_of_containsKey_eq_false_right . exact dl₁ . simp only [h] -theorem List.getValueCastD_filter_containsKey_of_containsKey_right [BEq α] [LawfulBEq α] +theorem getValueCastD_filter_containsKey_of_containsKey_right [BEq α] [LawfulBEq α] {l₁ l₂ : List ((a : α) × β a)} {k : α} {fallback: β k} (dl₁ : DistinctKeys l₁) : containsKey k l₂ → getValueCastD k (List.filter (fun p => containsKey p.fst l₂) l₁) fallback = @@ -5560,29 +5560,29 @@ theorem List.getValueCastD_filter_containsKey_of_containsKey_right [BEq α] [Law intro h rw [getValueCastD_eq_getValueCast?, getValueCastD_eq_getValueCast?] congr 1 - apply List.getValueCast?_filter_containsKey_of_containsKey_right dl₁ h + apply getValueCast?_filter_containsKey_of_containsKey_right dl₁ h -theorem List.getValueCastD_filter_containsKey_of_containsKey_eq_false_left [BEq α] [LawfulBEq α] +theorem getValueCastD_filter_containsKey_of_containsKey_eq_false_left [BEq α] [LawfulBEq α] {l₁ l₂ : List ((a : α) × β a)} {k : α} {fallback} (dl₁ : DistinctKeys l₁) : containsKey k l₁ = false → getValueCastD k (List.filter (fun p => containsKey p.fst l₂) l₁) fallback = fallback := by intro h apply getValueCastD_eq_fallback - apply List.containsKey_filter_containsKey_eq_false_of_containsKey_eq_false_left + apply containsKey_filter_containsKey_eq_false_of_containsKey_eq_false_left . exact dl₁ . exact h -theorem List.getValueCastD_filter_containsKey_of_containsKey_eq_false_right [BEq α] [LawfulBEq α] +theorem getValueCastD_filter_containsKey_of_containsKey_eq_false_right [BEq α] [LawfulBEq α] {l₁ l₂ : List ((a : α) × β a)} {k : α} {fallback} (dl₁ : DistinctKeys l₁) : containsKey k l₂ = false → getValueCastD k (List.filter (fun p => containsKey p.fst l₂) l₁) fallback = fallback := by intro h apply getValueCastD_eq_fallback - apply List.containsKey_filter_containsKey_eq_false_of_containsKey_eq_false_right + apply containsKey_filter_containsKey_eq_false_of_containsKey_eq_false_right . exact dl₁ . exact h -theorem List.getKeyD_filter_containsKey [BEq α] [EquivBEq α] +theorem getKeyD_filter_containsKey [BEq α] [EquivBEq α] {l₁ l₂ : List ((a : α) × β a)} {k : α} {fallback : α} (dl₁ : DistinctKeys l₁) : getKeyD k (List.filter (fun p => containsKey p.fst l₂) l₁) fallback = @@ -5591,14 +5591,14 @@ theorem List.getKeyD_filter_containsKey [BEq α] [EquivBEq α] case isTrue h => rw [getKeyD_eq_getKey?, getKeyD_eq_getKey?] congr 1 - apply List.getKey?_filter_containsKey_of_containsKey_right dl₁ h + apply getKey?_filter_containsKey_of_containsKey_right dl₁ h case isFalse h => apply getKeyD_eq_fallback - apply List.containsKey_filter_containsKey_eq_false_of_containsKey_eq_false_right + apply containsKey_filter_containsKey_eq_false_of_containsKey_eq_false_right · exact dl₁ · simp only [h] -theorem List.getKeyD_filter_containsKey_of_containsKey_right [BEq α] [EquivBEq α] +theorem getKeyD_filter_containsKey_of_containsKey_right [BEq α] [EquivBEq α] {l₁ l₂ : List ((a : α) × β a)} {k : α} {fallback : α} (dl₁ : DistinctKeys l₁) : containsKey k l₂ → getKeyD k (List.filter (fun p => containsKey p.fst l₂) l₁) fallback = @@ -5606,25 +5606,25 @@ theorem List.getKeyD_filter_containsKey_of_containsKey_right [BEq α] [EquivBEq intro h rw [getKeyD_eq_getKey?, getKeyD_eq_getKey?] congr 1 - apply List.getKey?_filter_containsKey_of_containsKey_right dl₁ h + apply getKey?_filter_containsKey_of_containsKey_right dl₁ h -theorem List.getKeyD_filter_containsKey_of_containsKey_eq_false_left [BEq α] [EquivBEq α] +theorem getKeyD_filter_containsKey_of_containsKey_eq_false_left [BEq α] [EquivBEq α] {l₁ l₂ : List ((a : α) × β a)} {k : α} {fallback : α} (dl₁ : DistinctKeys l₁) : containsKey k l₁ = false → getKeyD k (List.filter (fun p => containsKey p.fst l₂) l₁) fallback = fallback := by intro h apply getKeyD_eq_fallback - apply List.containsKey_filter_containsKey_eq_false_of_containsKey_eq_false_left + apply containsKey_filter_containsKey_eq_false_of_containsKey_eq_false_left · exact dl₁ · exact h -theorem List.getKeyD_filter_containsKey_of_containsKey_eq_false_right [BEq α] [EquivBEq α] +theorem getKeyD_filter_containsKey_of_containsKey_eq_false_right [BEq α] [EquivBEq α] {l₁ l₂ : List ((a : α) × β a)} {k : α} {fallback : α} (dl₁ : DistinctKeys l₁) : containsKey k l₂ = false → getKeyD k (List.filter (fun p => containsKey p.fst l₂) l₁) fallback = fallback := by intro h apply getKeyD_eq_fallback - apply List.containsKey_filter_containsKey_eq_false_of_containsKey_eq_false_right + apply containsKey_filter_containsKey_eq_false_of_containsKey_eq_false_right · exact dl₁ · exact h @@ -5636,7 +5636,7 @@ theorem getValue_filter_containsKey_eq_containsKey {β : Type v} [BEq α] [Equiv injections simp only [← getValue?_eq_some_getValue] apply getValue?_filter_containsKey_eq_containsKey dl₁ - · rw [List.containsKey_filter_containsKey_iff] at h₁ + · rw [containsKey_filter_containsKey_iff] at h₁ · exact h₁.2 · exact dl₁ @@ -5652,7 +5652,7 @@ theorem getValueD_filter_containsKey_eq_if_containsKey {β : Type v} [BEq α] [E apply getValue?_filter_containsKey_eq_containsKey dl₁ h case isFalse h => apply getValueD_eq_fallback - apply List.containsKey_filter_containsKey_eq_false_of_containsKey_eq_false_right + apply containsKey_filter_containsKey_eq_false_of_containsKey_eq_false_right · exact dl₁ · simp only [h] @@ -5666,23 +5666,23 @@ theorem getValueD_filter_containsKey_eq_containsKey {β : Type v} [BEq α] [Equi congr 1 apply getValue?_filter_containsKey_eq_containsKey dl₁ h -theorem List.getValueD_filter_of_containsKey_eq_false_left {β : Type v} [BEq α] [EquivBEq α] +theorem getValueD_filter_of_containsKey_eq_false_left {β : Type v} [BEq α] [EquivBEq α] {l₁ l₂ : List ((_ : α) × β)} {k : α} {fallback : β} (dl₁ : DistinctKeys l₁) : containsKey k l₁ = false → getValueD k (List.filter (fun p => containsKey p.fst l₂) l₁) fallback = fallback := by intro h apply getValueD_eq_fallback - apply List.containsKey_filter_containsKey_eq_false_of_containsKey_eq_false_left + apply containsKey_filter_containsKey_eq_false_of_containsKey_eq_false_left · exact dl₁ · exact h -theorem List.getValueD_filter_of_containsKey_eq_false_right {β : Type v} [BEq α] [EquivBEq α] +theorem getValueD_filter_of_containsKey_eq_false_right {β : Type v} [BEq α] [EquivBEq α] {l₁ l₂ : List ((_ : α) × β)} {k : α} {fallback : β} (dl₁ : DistinctKeys l₁) : containsKey k l₂ = false → getValueD k (List.filter (fun p => containsKey p.fst l₂) l₁) fallback = fallback := by intro h apply getValueD_eq_fallback - apply List.containsKey_filter_containsKey_eq_false_of_containsKey_eq_false_right + apply containsKey_filter_containsKey_eq_false_of_containsKey_eq_false_right · exact dl₁ · exact h @@ -6193,7 +6193,7 @@ theorem length_filter_key_eq_length_iff [BEq α] [EquivBEq α] {f : α → Bool} simp only [getKey, getKey?_eq_getEntry?, this] at h exact h -theorem List.length_filter_containsKey_eq_length_left [BEq α] [EquivBEq α] +theorem length_filter_containsKey_eq_length_left [BEq α] [EquivBEq α] {l₁ l₂ : List ((a : α) × β a)} (hl₁ : DistinctKeys l₁) (w : ∀ (a : α), containsKey a l₁ → containsKey a l₂) : (l₁.filter fun p => containsKey p.fst l₂).length = l₁.length := by @@ -6214,7 +6214,7 @@ theorem List.length_filter_containsKey_eq_length_left [BEq α] [EquivBEq α] . apply w rw [containsKey_cons, BEq.rfl, Bool.true_or] -theorem List.length_filter_containsKey_of_length_right [BEq α] [EquivBEq α] +theorem length_filter_containsKey_of_length_right [BEq α] [EquivBEq α] {l₁ l₂ : List ((a : α) × β a)} (hl₁ : DistinctKeys l₁) (hl₂ : DistinctKeys l₂) (w : ∀ (a : α), containsKey a l₂ → containsKey a l₁) : (l₁.filter fun p => containsKey p.fst l₂).length = l₂.length := by @@ -6326,7 +6326,7 @@ theorem List.length_filter_containsKey_of_length_right [BEq α] [EquivBEq α] case w.inr => exact inr -theorem List.size_add_size_eq_size_insertList_add_size_filter_containsKey [BEq α] [EquivBEq α] +theorem size_add_size_eq_size_insertList_add_size_filter_containsKey [BEq α] [EquivBEq α] {l₁ l₂ : List ((a : α) × β a)} (hl₁ : DistinctKeys l₁) (hl₂ : DistinctKeys l₂) : l₁.length + l₂.length = (insertList l₁ l₂).length + @@ -9054,7 +9054,7 @@ def interSmallerFn [BEq α] (l sofar : List ((a : α) × β a)) (k : α) : List | some kv' => List.insertEntry kv'.1 kv'.2 sofar | none => sofar -theorem List.distinctKeys_interSmallerFn [BEq α] [PartialEquivBEq α] +theorem distinctKeys_interSmallerFn [BEq α] [PartialEquivBEq α] (l sofar : List ((a : α) × β a)) (k : α) (hs : DistinctKeys sofar) : DistinctKeys (List.interSmallerFn l sofar k) := by rw [List.interSmallerFn] @@ -9119,7 +9119,7 @@ theorem getEntry?_interSmaller [BEq α] [PartialEquivBEq α] (l₁ l₂ : List ( simp only [List.containsKey_cons, Bool.or_comm (k == _), Option.filter_or, Option.or_assoc, List.getEntry?_interSmallerFn] -theorem List.distinctKeys_interSmaller [BEq α] [PartialEquivBEq α] {l₁ l₂ : List ((a : α) × β a)} : +theorem distinctKeys_interSmaller [BEq α] [PartialEquivBEq α] {l₁ l₂ : List ((a : α) × β a)} : DistinctKeys (interSmaller l₁ l₂) := by rw [interSmaller] suffices ∀ l, DistinctKeys l → DistinctKeys (l₂.foldl (fun sofar kv => List.interSmallerFn l₁ sofar kv.1) l) by @@ -9132,10 +9132,10 @@ theorem List.distinctKeys_interSmaller [BEq α] [PartialEquivBEq α] {l₁ l₂ apply ih exact distinctKeys_interSmallerFn _ _ _ hl -theorem List.interSmaller_perm_filter [BEq α] [EquivBEq α] (l₁ l₂ : List ((a : α) × β a)) (h₁ : DistinctKeys l₁) : +theorem interSmaller_perm_filter [BEq α] [EquivBEq α] (l₁ l₂ : List ((a : α) × β a)) (h₁ : DistinctKeys l₁) : List.Perm (List.interSmaller l₁ l₂) (l₁.filter (fun kv => containsKey kv.1 l₂)) := by apply List.getEntry?_ext - · exact List.distinctKeys_interSmaller + · exact distinctKeys_interSmaller · exact h₁.filter (f := fun k v => containsKey k l₂) · intro k' rw [List.getEntry?_filter h₁, List.getEntry?_interSmaller] From dac309f780f9b7fb8ee561f2316e84facbd3705e Mon Sep 17 00:00:00 2001 From: Wojciech Rozowski Date: Sat, 15 Nov 2025 23:28:02 +0100 Subject: [PATCH 65/88] fixes --- src/Std/Data/Internal/List/Associative.lean | 18 ++++++++---------- 1 file changed, 8 insertions(+), 10 deletions(-) diff --git a/src/Std/Data/Internal/List/Associative.lean b/src/Std/Data/Internal/List/Associative.lean index 157f9a871734..bb18901ade08 100644 --- a/src/Std/Data/Internal/List/Associative.lean +++ b/src/Std/Data/Internal/List/Associative.lean @@ -5295,7 +5295,7 @@ theorem List.getValueCast?_diff_of_containsKey_right [BEq α] [LawfulBEq α] contradiction · exact dl₁ -theorem List.getKey?_filter_containsKey [BEq α] [EquivBEq α] +theorem getKey?_filter_containsKey [BEq α] [EquivBEq α] {l₁ l₂ : List ((a : α) × β a)} {k : α} (dl₁ : DistinctKeys l₁) : getKey? k (List.filter (fun p => containsKey p.fst l₂) l₁) = @@ -5339,8 +5339,6 @@ theorem getKey?_filter_containsKey_of_containsKey_right [BEq α] [EquivBEq α] · simp only [↓reduceIte] · exact dl₁ - - /-- Internal implementation detail of the hash map -/ def eraseList [BEq α] (l toErase : List ((a : α) × β a)) : List ((a : α) × β a) := match toErase with @@ -5418,7 +5416,7 @@ theorem eraseList_perm_of_perm_first [BEq α] [EquivBEq α] {l1 l2 toErase : Lis simp only [eraseList] apply ih (eraseKey_of_perm distinct h) (DistinctKeys.eraseKey distinct) -theorem List.getKey?_filter_containsKey_of_containsKey_eq_false_left [BEq α] [EquivBEq α] +theorem getKey?_filter_containsKey_of_containsKey_eq_false_left [BEq α] [EquivBEq α] {l₁ l₂ : List ((a : α) × β a)} {k : α} (dl₁ : DistinctKeys l₁) : containsKey k l₁ = false → getKey? k (List.filter (fun p => containsKey p.fst l₂) l₁) = none := by @@ -5645,21 +5643,21 @@ theorem containsKey_filter_containsKey_eq_false_of_containsKey_eq_false_left [BE . simp [h] . exact hl₁ -theorem List.containsKey_diff_eq_false_of_containsKey_eq_false_left [BEq α] [EquivBEq α] {l₁ l₂ : List ((a : α) × β a)} {hl₁ : DistinctKeys l₁} {k : α} : +theorem containsKey_diff_eq_false_of_containsKey_eq_false_left [BEq α] [EquivBEq α] {l₁ l₂ : List ((a : α) × β a)} {hl₁ : DistinctKeys l₁} {k : α} : containsKey k l₁ = false → containsKey k (List.filter (fun p => !containsKey p.fst l₂) l₁) = false := by intro h rw [List.containsKey_diff_eq_false] . simp [h] . exact hl₁ -theorem List.containsKey_filter_containsKey_eq_false_of_containsKey_eq_false_right [BEq α] [EquivBEq α] {l₁ l₂ : List ((a : α) × β a)} {hl₁ : DistinctKeys l₁} {k : α} : +theorem containsKey_filter_containsKey_eq_false_of_containsKey_eq_false_right [BEq α] [EquivBEq α] {l₁ l₂ : List ((a : α) × β a)} {hl₁ : DistinctKeys l₁} {k : α} : containsKey k l₂ = false → containsKey k (List.filter (fun p => containsKey p.fst l₂) l₁) = false := by intro h rw [containsKey_filter_containsKey] . simp [h] . exact hl₁ -theorem List.containsKey_diff_of_containsKey_eq_false_right [BEq α] [EquivBEq α] {l₁ l₂ : List ((a : α) × β a)} {hl₁ : DistinctKeys l₁} {k : α} : +theorem containsKey_diff_of_containsKey_eq_false_right [BEq α] [EquivBEq α] {l₁ l₂ : List ((a : α) × β a)} {hl₁ : DistinctKeys l₁} {k : α} : containsKey k l₂ = true → containsKey k (List.filter (fun p => !containsKey p.fst l₂) l₁) = false := by intro h rw [List.containsKey_diff_eq_false] @@ -5674,7 +5672,7 @@ theorem List.getValueCastD_diff [BEq α] [LawfulBEq α] split case isTrue h => apply getValueCastD_eq_fallback - apply List.containsKey_diff_of_containsKey_eq_false_right + apply containsKey_diff_of_containsKey_eq_false_right · exact dl₁ · exact h case isFalse h => @@ -5689,7 +5687,7 @@ theorem List.getValueCastD_diff_of_containsKey_eq_false_left [BEq α] [LawfulBEq containsKey k l₁ = false → getValueCastD k (List.filter (fun p => !containsKey p.fst l₂) l₁) fallback = fallback := by intro h apply getValueCastD_eq_fallback - apply List.containsKey_diff_eq_false_of_containsKey_eq_false_left + apply containsKey_diff_eq_false_of_containsKey_eq_false_left · exact dl₁ · exact h @@ -5733,7 +5731,7 @@ theorem List.getValueCastD_diff_of_containsKey_right [BEq α] [LawfulBEq α] containsKey k l₂ = true → getValueCastD k (List.filter (fun p => !containsKey p.fst l₂) l₁) fallback = fallback := by intro h apply getValueCastD_eq_fallback - apply List.containsKey_diff_of_containsKey_eq_false_right + apply containsKey_diff_of_containsKey_eq_false_right · exact dl₁ · exact h From b56b1bb151ec5f8fb2a6b20ae47da5c9a4b24fac Mon Sep 17 00:00:00 2001 From: Wojciech Rozowski Date: Sat, 15 Nov 2025 23:42:46 +0100 Subject: [PATCH 66/88] fix compilation --- src/Std/Data/DHashMap/Internal/RawLemmas.lean | 372 +++++++++--------- src/Std/Data/Internal/List/Associative.lean | 4 +- 2 files changed, 188 insertions(+), 188 deletions(-) diff --git a/src/Std/Data/DHashMap/Internal/RawLemmas.lean b/src/Std/Data/DHashMap/Internal/RawLemmas.lean index 0a168fd6df1f..f17911f64bc3 100644 --- a/src/Std/Data/DHashMap/Internal/RawLemmas.lean +++ b/src/Std/Data/DHashMap/Internal/RawLemmas.lean @@ -3212,115 +3212,115 @@ theorem getKey?_diff [EquivBEq α] [LawfulHashable α] (m₁.diff m₂).getKey? k = if m₂.contains k then none else m₁.getKey? k := by simp_to_model [diff, contains, getKey?] using List.getKey?_diff -theorem getKey?_diff_of_contains_eq_false_right [EquivBEq α] [LawfulHashable α] - (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} (h : m₂.contains k = false) : - (m₁.diff m₂).getKey? k = m₁.getKey? k := by - revert h - simp_to_model [contains, getKey?, diff] using List.getKey?_diff_of_containsKey_eq_false_right - -theorem getKey?_diff_of_contains_eq_false_left [EquivBEq α] [LawfulHashable α] - (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} (h : m₁.contains k = false) : - (m₁.diff m₂).getKey? k = none := by - revert h - simp_to_model [contains, getKey?, diff] using List.getKey?_diff_of_containsKey_eq_false_left - -theorem getKey?_diff_of_contains_right [EquivBEq α] [LawfulHashable α] - (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} (h : m₂.contains k) : - (m₁.diff m₂).getKey? k = none := by - revert h - simp_to_model [contains, getKey?, diff] using List.getKey?_diff_of_containsKey_right - -/- getKey -/ -theorem getKey_diff [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) - {k : α} {h_contains : (m₁.diff m₂).contains k} : - (m₁.diff m₂).getKey k h_contains = - m₁.getKey k ((contains_diff_iff h₁ h₂).1 h_contains).1 := by - simp_to_model [diff, contains, getKey] using List.getKey_diff - -/- getKeyD -/ -theorem getKeyD_diff [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) - (h₂ : m₂.val.WF) {k fallback : α} : - (m₁.diff m₂).getKeyD k fallback = - if m₂.contains k then fallback else m₁.getKeyD k fallback := by - simp_to_model [diff, getKeyD, contains] using List.getKeyD_diff - -theorem getKeyD_diff_of_contains_eq_false_right [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) - (h₂ : m₂.val.WF) {k fallback : α} (h : m₂.contains k = false) : - (m₁.diff m₂).getKeyD k fallback = m₁.getKeyD k fallback := by - revert h - simp_to_model [contains, diff, getKeyD] using List.getKeyD_diff_of_containsKey_eq_false_right - -theorem getKeyD_diff_of_contains_right [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) - (h₂ : m₂.val.WF) {k fallback : α} (h : m₂.contains k) : - (m₁.diff m₂).getKeyD k fallback = fallback := by - revert h - simp_to_model [diff, getKeyD, contains] using List.getKeyD_diff_of_containsKey_right - -theorem getKeyD_diff_of_contains_eq_false_left [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) - (h₂ : m₂.val.WF) {k fallback : α} (h : m₁.contains k = false) : - (m₁.diff m₂).getKeyD k fallback = fallback := by - revert h - simp_to_model [diff, getKeyD, contains] using List.getKeyD_diff_of_containsKey_eq_false_left - -/- getKey! -/ -theorem getKey!_diff [EquivBEq α] [LawfulHashable α] [Inhabited α] (h₁ : m₁.val.WF) - (h₂ : m₂.val.WF) {k : α} : - (m₁.diff m₂).getKey! k = if m₂.contains k then default else m₁.getKey! k := by - simp_to_model [diff, getKey!, contains] using List.getKeyD_diff - -theorem getKey!_diff_of_contains_eq_false_right [EquivBEq α] [LawfulHashable α] [Inhabited α] (h₁ : m₁.val.WF) - (h₂ : m₂.val.WF) {k : α} (h : m₂.contains k = false) : - (m₁.diff m₂).getKey! k = m₁.getKey! k := by - revert h - simp_to_model [diff, getKey!, contains] using List.getKeyD_diff_of_containsKey_eq_false_right - -theorem getKey!_diff_of_contains_right [EquivBEq α] [LawfulHashable α] [Inhabited α] (h₁ : m₁.val.WF) - (h₂ : m₂.val.WF) {k : α} (h : m₂.contains k) : - (m₁.diff m₂).getKey! k = default := by - revert h - simp_to_model [diff, getKey!, contains] using List.getKeyD_diff_of_containsKey_right - -theorem getKey!_diff_of_contains_eq_false_left [EquivBEq α] [LawfulHashable α] [Inhabited α] (h₁ : m₁.val.WF) - (h₂ : m₂.val.WF) {k : α} (h : m₁.contains k = false) : - (m₁.diff m₂).getKey! k = default := by - revert h - simp_to_model [diff, getKey!, contains] using List.getKeyD_diff_of_containsKey_eq_false_left - -/- size -/ -theorem size_diff_le_size_left [EquivBEq α] [LawfulHashable α] - (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) : - (m₁.diff m₂).1.size ≤ m₁.1.size := by - simp_to_model [diff, size] using List.length_diff_le - -theorem size_diff_eq_size_left_of_disjoint [EquivBEq α] [LawfulHashable α] - (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) - (h : ∀ (a : α), m₁.contains a → m₂.contains a = false) : - (m₁.diff m₂).1.size = m₁.1.size := by - revert h - simp_to_model [diff, size, contains] using List.length_diff_eq_length_left - -theorem size_diff_add_size_inter_eq_size_left [EquivBEq α] [LawfulHashable α] - (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) : - (m₁.diff m₂).1.size + (m₁.inter m₂).1.size = m₁.1.size := by - simp_to_model [diff, inter, size] using List.size_diff_add_size_inter_eq_size_left - -/- isEmpty -/ -@[simp] -theorem isEmpty_diff_of_isEmpty_left [EquivBEq α] [LawfulHashable α] - (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) (h : m₁.1.isEmpty) : - (m₁.diff m₂).1.isEmpty = true := by - revert h - simp_to_model [isEmpty, diff] using List.isEmpty_diff_left - -theorem isEmpty_diff_iff [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) : - (m₁.diff m₂).1.isEmpty ↔ ∀ k, m₁.contains k → m₂.contains k := by - simp_to_model [diff, contains, isEmpty] using List.isEmpty_diff_iff - -theorem isEmpty_diff_of_subset [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) - (h : ∀ k, m₁.contains k → m₂.contains k) : - (m₁.diff m₂).1.isEmpty = true := by - revert h - simp_to_model [diff, contains, isEmpty] using List.isEmpty_diff_of_subset +-- theorem getKey?_diff_of_contains_eq_false_right [EquivBEq α] [LawfulHashable α] +-- (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} (h : m₂.contains k = false) : +-- (m₁.diff m₂).getKey? k = m₁.getKey? k := by +-- revert h +-- simp_to_model [contains, getKey?, diff] using List.getKey?_diff_of_containsKey_eq_false_right + +-- theorem getKey?_diff_of_contains_eq_false_left [EquivBEq α] [LawfulHashable α] +-- (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} (h : m₁.contains k = false) : +-- (m₁.diff m₂).getKey? k = none := by +-- revert h +-- simp_to_model [contains, getKey?, diff] using List.getKey?_diff_of_containsKey_eq_false_left + +-- theorem getKey?_diff_of_contains_right [EquivBEq α] [LawfulHashable α] +-- (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} (h : m₂.contains k) : +-- (m₁.diff m₂).getKey? k = none := by +-- revert h +-- simp_to_model [contains, getKey?, diff] using List.getKey?_diff_of_containsKey_right + +-- /- getKey -/ +-- theorem getKey_diff [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) +-- {k : α} {h_contains : (m₁.diff m₂).contains k} : +-- (m₁.diff m₂).getKey k h_contains = +-- m₁.getKey k ((contains_diff_iff h₁ h₂).1 h_contains).1 := by +-- simp_to_model [diff, contains, getKey] using List.getKey_diff + +-- /- getKeyD -/ +-- theorem getKeyD_diff [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) +-- (h₂ : m₂.val.WF) {k fallback : α} : +-- (m₁.diff m₂).getKeyD k fallback = +-- if m₂.contains k then fallback else m₁.getKeyD k fallback := by +-- simp_to_model [diff, getKeyD, contains] using List.getKeyD_diff + +-- theorem getKeyD_diff_of_contains_eq_false_right [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) +-- (h₂ : m₂.val.WF) {k fallback : α} (h : m₂.contains k = false) : +-- (m₁.diff m₂).getKeyD k fallback = m₁.getKeyD k fallback := by +-- revert h +-- simp_to_model [contains, diff, getKeyD] using List.getKeyD_diff_of_containsKey_eq_false_right + +-- theorem getKeyD_diff_of_contains_right [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) +-- (h₂ : m₂.val.WF) {k fallback : α} (h : m₂.contains k) : +-- (m₁.diff m₂).getKeyD k fallback = fallback := by +-- revert h +-- simp_to_model [diff, getKeyD, contains] using List.getKeyD_diff_of_containsKey_right + +-- theorem getKeyD_diff_of_contains_eq_false_left [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) +-- (h₂ : m₂.val.WF) {k fallback : α} (h : m₁.contains k = false) : +-- (m₁.diff m₂).getKeyD k fallback = fallback := by +-- revert h +-- simp_to_model [diff, getKeyD, contains] using List.getKeyD_diff_of_containsKey_eq_false_left + +-- /- getKey! -/ +-- theorem getKey!_diff [EquivBEq α] [LawfulHashable α] [Inhabited α] (h₁ : m₁.val.WF) +-- (h₂ : m₂.val.WF) {k : α} : +-- (m₁.diff m₂).getKey! k = if m₂.contains k then default else m₁.getKey! k := by +-- simp_to_model [diff, getKey!, contains] using List.getKeyD_diff + +-- theorem getKey!_diff_of_contains_eq_false_right [EquivBEq α] [LawfulHashable α] [Inhabited α] (h₁ : m₁.val.WF) +-- (h₂ : m₂.val.WF) {k : α} (h : m₂.contains k = false) : +-- (m₁.diff m₂).getKey! k = m₁.getKey! k := by +-- revert h +-- simp_to_model [diff, getKey!, contains] using List.getKeyD_diff_of_containsKey_eq_false_right + +-- theorem getKey!_diff_of_contains_right [EquivBEq α] [LawfulHashable α] [Inhabited α] (h₁ : m₁.val.WF) +-- (h₂ : m₂.val.WF) {k : α} (h : m₂.contains k) : +-- (m₁.diff m₂).getKey! k = default := by +-- revert h +-- simp_to_model [diff, getKey!, contains] using List.getKeyD_diff_of_containsKey_right + +-- theorem getKey!_diff_of_contains_eq_false_left [EquivBEq α] [LawfulHashable α] [Inhabited α] (h₁ : m₁.val.WF) +-- (h₂ : m₂.val.WF) {k : α} (h : m₁.contains k = false) : +-- (m₁.diff m₂).getKey! k = default := by +-- revert h +-- simp_to_model [diff, getKey!, contains] using List.getKeyD_diff_of_containsKey_eq_false_left + +-- /- size -/ +-- theorem size_diff_le_size_left [EquivBEq α] [LawfulHashable α] +-- (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) : +-- (m₁.diff m₂).1.size ≤ m₁.1.size := by +-- simp_to_model [diff, size] using List.length_diff_le + +-- theorem size_diff_eq_size_left_of_disjoint [EquivBEq α] [LawfulHashable α] +-- (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) +-- (h : ∀ (a : α), m₁.contains a → m₂.contains a = false) : +-- (m₁.diff m₂).1.size = m₁.1.size := by +-- revert h +-- simp_to_model [diff, size, contains] using List.length_diff_eq_length_left + +-- theorem size_diff_add_size_inter_eq_size_left [EquivBEq α] [LawfulHashable α] +-- (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) : +-- (m₁.diff m₂).1.size + (m₁.inter m₂).1.size = m₁.1.size := by +-- simp_to_model [diff, inter, size] using List.size_diff_add_size_inter_eq_size_left + +-- /- isEmpty -/ +-- @[simp] +-- theorem isEmpty_diff_of_isEmpty_left [EquivBEq α] [LawfulHashable α] +-- (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) (h : m₁.1.isEmpty) : +-- (m₁.diff m₂).1.isEmpty = true := by +-- revert h +-- simp_to_model [isEmpty, diff] using List.isEmpty_diff_left + +-- theorem isEmpty_diff_iff [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) : +-- (m₁.diff m₂).1.isEmpty ↔ ∀ k, m₁.contains k → m₂.contains k := by +-- simp_to_model [diff, contains, isEmpty] using List.isEmpty_diff_iff + +-- theorem isEmpty_diff_of_subset [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) +-- (h : ∀ k, m₁.contains k → m₂.contains k) : +-- (m₁.diff m₂).1.isEmpty = true := by +-- revert h +-- simp_to_model [diff, contains, isEmpty] using List.isEmpty_diff_of_subset end Diff @@ -3328,83 +3328,83 @@ namespace Const variable {β : Type v} {m₁ m₂ : Raw₀ α (fun _ => β)} -/- get? -/ -theorem get?_diff [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} : - Const.get? (m₁.diff m₂) k = if m₂.contains k then none else Const.get? m₁ k := by - simp_to_model [diff, Const.get?, contains] using List.getValue?_diff - -theorem get?_diff_of_contains_eq_false_right [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) - {k : α} (h : m₂.contains k = false) : - Const.get? (m₁.diff m₂) k = Const.get? m₁ k := by - revert h - simp_to_model [diff, contains, Const.get?] using List.getValue?_diff_of_containsKey_eq_false_right - -theorem get?_diff_of_contains_eq_false_left [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) - {k : α} (h : m₁.contains k = false) : - Const.get? (m₁.diff m₂) k = none := by - revert h - simp_to_model [diff, Const.get?, contains] using List.getValue?_diff_of_containsKey_eq_false_left - -theorem get?_diff_of_contains_right [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) - {k : α} (h : m₂.contains k) : - Const.get? (m₁.diff m₂) k = none := by - revert h - simp_to_model [diff, Const.get?, contains] using List.getValue?_diff_of_containsKey_right - -/- get -/ -theorem get_diff [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) - {k : α} {h_contains : (m₁.diff m₂).contains k} : - Const.get (m₁.diff m₂) k h_contains = - Const.get m₁ k ((contains_diff_iff h₁ h₂).1 h_contains).1 := by - simp_to_model [diff, Const.get, contains] using List.getValue_diff - -/- getD -/ -theorem getD_diff [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) - {k : α} {fallback : β} : - Const.getD (m₁.diff m₂) k fallback = - if m₂.contains k then fallback else Const.getD m₁ k fallback := by - simp_to_model [diff, Const.getD, contains] using List.getValueD_diff - -theorem getD_diff_of_contains_eq_false_right [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) - {k : α} {fallback : β} (h : m₂.contains k = false) : - Const.getD (m₁.diff m₂) k fallback = Const.getD m₁ k fallback := by - revert h - simp_to_model [diff, contains, Const.getD] using List.getValueD_diff_of_containsKey_eq_false_right - -theorem getD_diff_of_contains_right [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) - {k : α} {fallback : β} (h : m₂.contains k) : - Const.getD (m₁.diff m₂) k fallback = fallback := by - revert h - simp_to_model [diff, Const.getD, contains] using List.getValueD_diff_of_containsKey_right - -theorem getD_diff_of_contains_eq_false_left [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) - {k : α} {fallback : β} (h : m₁.contains k = false) : - Const.getD (m₁.diff m₂) k fallback = fallback := by - revert h - simp_to_model [diff, Const.getD, contains] using List.getValueD_diff_of_containsKey_eq_false_left - -/- get! -/ -theorem get!_diff [EquivBEq α] [LawfulHashable α] [Inhabited β] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} : - Const.get! (m₁.diff m₂) k = if m₂.contains k then default else Const.get! m₁ k := by - simp_to_model [diff, Const.get!, contains] using List.getValueD_diff - -theorem get!_diff_of_contains_eq_false_right [EquivBEq α] [LawfulHashable α] [Inhabited β] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) - {k : α} (h : m₂.contains k = false) : - Const.get! (m₁.diff m₂) k = Const.get! m₁ k := by - revert h - simp_to_model [diff, contains, Const.get!] using List.getValueD_diff_of_containsKey_eq_false_right - -theorem get!_diff_of_contains_right [EquivBEq α] [LawfulHashable α] [Inhabited β] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) - {k : α} (h : m₂.contains k) : - Const.get! (m₁.diff m₂) k = default := by - revert h - simp_to_model [diff, Const.get!, contains] using List.getValueD_diff_of_containsKey_right - -theorem get!_diff_of_contains_eq_false_left [EquivBEq α] [LawfulHashable α] [Inhabited β] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) - {k : α} (h : m₁.contains k = false) : - Const.get! (m₁.diff m₂) k = default := by - revert h - simp_to_model [diff, Const.get!, contains] using List.getValueD_diff_of_containsKey_eq_false_left +-- /- get? -/ +-- theorem get?_diff [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} : +-- Const.get? (m₁.diff m₂) k = if m₂.contains k then none else Const.get? m₁ k := by +-- simp_to_model [diff, Const.get?, contains] using List.getValue?_diff + +-- theorem get?_diff_of_contains_eq_false_right [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) +-- {k : α} (h : m₂.contains k = false) : +-- Const.get? (m₁.diff m₂) k = Const.get? m₁ k := by +-- revert h +-- simp_to_model [diff, contains, Const.get?] using List.getValue?_diff_of_containsKey_eq_false_right + +-- theorem get?_diff_of_contains_eq_false_left [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) +-- {k : α} (h : m₁.contains k = false) : +-- Const.get? (m₁.diff m₂) k = none := by +-- revert h +-- simp_to_model [diff, Const.get?, contains] using List.getValue?_diff_of_containsKey_eq_false_left + +-- theorem get?_diff_of_contains_right [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) +-- {k : α} (h : m₂.contains k) : +-- Const.get? (m₁.diff m₂) k = none := by +-- revert h +-- simp_to_model [diff, Const.get?, contains] using List.getValue?_diff_of_containsKey_right + +-- /- get -/ +-- theorem get_diff [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) +-- {k : α} {h_contains : (m₁.diff m₂).contains k} : +-- Const.get (m₁.diff m₂) k h_contains = +-- Const.get m₁ k ((contains_diff_iff h₁ h₂).1 h_contains).1 := by +-- simp_to_model [diff, Const.get, contains] using List.getValue_diff + +-- /- getD -/ +-- theorem getD_diff [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) +-- {k : α} {fallback : β} : +-- Const.getD (m₁.diff m₂) k fallback = +-- if m₂.contains k then fallback else Const.getD m₁ k fallback := by +-- simp_to_model [diff, Const.getD, contains] using List.getValueD_diff + +-- theorem getD_diff_of_contains_eq_false_right [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) +-- {k : α} {fallback : β} (h : m₂.contains k = false) : +-- Const.getD (m₁.diff m₂) k fallback = Const.getD m₁ k fallback := by +-- revert h +-- simp_to_model [diff, contains, Const.getD] using List.getValueD_diff_of_containsKey_eq_false_right + +-- theorem getD_diff_of_contains_right [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) +-- {k : α} {fallback : β} (h : m₂.contains k) : +-- Const.getD (m₁.diff m₂) k fallback = fallback := by +-- revert h +-- simp_to_model [diff, Const.getD, contains] using List.getValueD_diff_of_containsKey_right + +-- theorem getD_diff_of_contains_eq_false_left [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) +-- {k : α} {fallback : β} (h : m₁.contains k = false) : +-- Const.getD (m₁.diff m₂) k fallback = fallback := by +-- revert h +-- simp_to_model [diff, Const.getD, contains] using List.getValueD_diff_of_containsKey_eq_false_left + +-- /- get! -/ +-- theorem get!_diff [EquivBEq α] [LawfulHashable α] [Inhabited β] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} : +-- Const.get! (m₁.diff m₂) k = if m₂.contains k then default else Const.get! m₁ k := by +-- simp_to_model [diff, Const.get!, contains] using List.getValueD_diff + +-- theorem get!_diff_of_contains_eq_false_right [EquivBEq α] [LawfulHashable α] [Inhabited β] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) +-- {k : α} (h : m₂.contains k = false) : +-- Const.get! (m₁.diff m₂) k = Const.get! m₁ k := by +-- revert h +-- simp_to_model [diff, contains, Const.get!] using List.getValueD_diff_of_containsKey_eq_false_right + +-- theorem get!_diff_of_contains_right [EquivBEq α] [LawfulHashable α] [Inhabited β] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) +-- {k : α} (h : m₂.contains k) : +-- Const.get! (m₁.diff m₂) k = default := by +-- revert h +-- simp_to_model [diff, Const.get!, contains] using List.getValueD_diff_of_containsKey_right + +-- theorem get!_diff_of_contains_eq_false_left [EquivBEq α] [LawfulHashable α] [Inhabited β] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) +-- {k : α} (h : m₁.contains k = false) : +-- Const.get! (m₁.diff m₂) k = default := by +-- revert h +-- simp_to_model [diff, Const.get!, contains] using List.getValueD_diff_of_containsKey_eq_false_left end Const diff --git a/src/Std/Data/Internal/List/Associative.lean b/src/Std/Data/Internal/List/Associative.lean index bb18901ade08..f245451d53b6 100644 --- a/src/Std/Data/Internal/List/Associative.lean +++ b/src/Std/Data/Internal/List/Associative.lean @@ -5601,7 +5601,7 @@ theorem List.getValueCast_diff [BEq α] [LawfulBEq α] simp [h₁.2] exact dl₁ -theorem List.getValueCast_filter_containsKey [BEq α] [LawfulBEq α] +theorem getValueCast_filter_containsKey [BEq α] [LawfulBEq α] {l₁ l₂ : List ((a : α) × β a)} {k : α} (dl₁ : DistinctKeys l₁) {h₁ h₂} : getValueCast k (List.filter (fun p => containsKey p.fst l₂) l₁) h₁ = getValueCast k l₁ h₂ := by @@ -5851,7 +5851,7 @@ theorem List.getValueCastD_diff_of_containsKey_eq_false_right [BEq α] [LawfulBE rw [List.getValueCast?_diff dl₁] simp [h] -theorem List.getKeyD_filter_containsKey [BEq α] [EquivBEq α] +theorem getKeyD_filter_containsKey [BEq α] [EquivBEq α] {l₁ l₂ : List ((a : α) × β a)} {k : α} {fallback : α} (dl₁ : DistinctKeys l₁) : getKeyD k (List.filter (fun p => containsKey p.fst l₂) l₁) fallback = From 6303bdc14ab54a643d34b5d647241416e7f4d375 Mon Sep 17 00:00:00 2001 From: Wojciech Rozowski Date: Sun, 16 Nov 2025 00:02:33 +0100 Subject: [PATCH 67/88] chore: remove untrue lemma --- src/Std/Data/DHashMap/Internal/RawLemmas.lean | 6 ------ 1 file changed, 6 deletions(-) diff --git a/src/Std/Data/DHashMap/Internal/RawLemmas.lean b/src/Std/Data/DHashMap/Internal/RawLemmas.lean index f17911f64bc3..c099ab2da72c 100644 --- a/src/Std/Data/DHashMap/Internal/RawLemmas.lean +++ b/src/Std/Data/DHashMap/Internal/RawLemmas.lean @@ -3121,12 +3121,6 @@ theorem contains_diff_eq_false_of_contains_right [EquivBEq α] [LawfulHashable revert h simp_to_model [diff, contains] using List.containsKey_diff_of_containsKey_eq_false_right -/- Equiv -/ -theorem diff_erase_right_equiv_erase_diff [EquivBEq α] [LawfulHashable α] {k : α} - (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) : - (m₁.diff (m₂.erase k)).1.Equiv ((m₁.diff m₂).erase k).1 := by - sorry - /- get? -/ theorem get?_diff [LawfulBEq α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} : (m₁.diff m₂).get? k = if m₂.contains k then none else m₁.get? k := by From fa5196993d6ba05c91cc2bcf5f3ea785b0a9035f Mon Sep 17 00:00:00 2001 From: Wojciech Rozowski Date: Mon, 17 Nov 2025 12:49:22 +0100 Subject: [PATCH 68/88] More set difference lemmas --- src/Std/Data/DHashMap/Internal/RawLemmas.lean | 356 +++++++++--------- src/Std/Data/Internal/List/Associative.lean | 315 +++++++++++++++- 2 files changed, 470 insertions(+), 201 deletions(-) diff --git a/src/Std/Data/DHashMap/Internal/RawLemmas.lean b/src/Std/Data/DHashMap/Internal/RawLemmas.lean index c099ab2da72c..2fe0b6ce661f 100644 --- a/src/Std/Data/DHashMap/Internal/RawLemmas.lean +++ b/src/Std/Data/DHashMap/Internal/RawLemmas.lean @@ -3124,7 +3124,7 @@ theorem contains_diff_eq_false_of_contains_right [EquivBEq α] [LawfulHashable /- get? -/ theorem get?_diff [LawfulBEq α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} : (m₁.diff m₂).get? k = if m₂.contains k then none else m₁.get? k := by - simp_to_model [diff, get?, contains] using List.getValueCast?_diff + simp_to_model [diff, get?, contains] using getValueCast?_diff theorem get?_diff_of_contains_eq_false_right [LawfulBEq α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} (h : m₂.contains k = false) : @@ -3162,7 +3162,7 @@ theorem getD_diff_of_contains_eq_false_right [LawfulBEq α] (h₁ : m₁.val.WF) {k : α} {fallback : β k} (h : m₂.contains k = false) : (m₁.diff m₂).getD k fallback = m₁.getD k fallback := by revert h - simp_to_model [diff, contains, getD] using List.getValueCastD_diff_of_containsKey_eq_false_right + simp_to_model [diff, contains, getD] using getValueCastD_diff_of_containsKey_eq_false_right theorem getD_diff_of_contains_right [LawfulBEq α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} {fallback : β k} (h : m₂.contains k) : @@ -3186,7 +3186,7 @@ theorem get!_diff_of_contains_eq_false_right [LawfulBEq α] (h₁ : m₁.val.WF) {k : α} [Inhabited (β k)] (h : m₂.contains k = false) : (m₁.diff m₂).get! k = m₁.get! k := by revert h - simp_to_model [diff, contains, get!] using List.getValueCastD_diff_of_containsKey_eq_false_right + simp_to_model [diff, contains, get!] using getValueCastD_diff_of_containsKey_eq_false_right theorem get!_diff_of_contains_right [LawfulBEq α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} [Inhabited (β k)] (h : m₂.contains k) : @@ -3206,115 +3206,109 @@ theorem getKey?_diff [EquivBEq α] [LawfulHashable α] (m₁.diff m₂).getKey? k = if m₂.contains k then none else m₁.getKey? k := by simp_to_model [diff, contains, getKey?] using List.getKey?_diff --- theorem getKey?_diff_of_contains_eq_false_right [EquivBEq α] [LawfulHashable α] --- (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} (h : m₂.contains k = false) : --- (m₁.diff m₂).getKey? k = m₁.getKey? k := by --- revert h --- simp_to_model [contains, getKey?, diff] using List.getKey?_diff_of_containsKey_eq_false_right - --- theorem getKey?_diff_of_contains_eq_false_left [EquivBEq α] [LawfulHashable α] --- (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} (h : m₁.contains k = false) : --- (m₁.diff m₂).getKey? k = none := by --- revert h --- simp_to_model [contains, getKey?, diff] using List.getKey?_diff_of_containsKey_eq_false_left - --- theorem getKey?_diff_of_contains_right [EquivBEq α] [LawfulHashable α] --- (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} (h : m₂.contains k) : --- (m₁.diff m₂).getKey? k = none := by --- revert h --- simp_to_model [contains, getKey?, diff] using List.getKey?_diff_of_containsKey_right - --- /- getKey -/ --- theorem getKey_diff [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) --- {k : α} {h_contains : (m₁.diff m₂).contains k} : --- (m₁.diff m₂).getKey k h_contains = --- m₁.getKey k ((contains_diff_iff h₁ h₂).1 h_contains).1 := by --- simp_to_model [diff, contains, getKey] using List.getKey_diff - --- /- getKeyD -/ --- theorem getKeyD_diff [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) --- (h₂ : m₂.val.WF) {k fallback : α} : --- (m₁.diff m₂).getKeyD k fallback = --- if m₂.contains k then fallback else m₁.getKeyD k fallback := by --- simp_to_model [diff, getKeyD, contains] using List.getKeyD_diff - --- theorem getKeyD_diff_of_contains_eq_false_right [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) --- (h₂ : m₂.val.WF) {k fallback : α} (h : m₂.contains k = false) : --- (m₁.diff m₂).getKeyD k fallback = m₁.getKeyD k fallback := by --- revert h --- simp_to_model [contains, diff, getKeyD] using List.getKeyD_diff_of_containsKey_eq_false_right - --- theorem getKeyD_diff_of_contains_right [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) --- (h₂ : m₂.val.WF) {k fallback : α} (h : m₂.contains k) : --- (m₁.diff m₂).getKeyD k fallback = fallback := by --- revert h --- simp_to_model [diff, getKeyD, contains] using List.getKeyD_diff_of_containsKey_right - --- theorem getKeyD_diff_of_contains_eq_false_left [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) --- (h₂ : m₂.val.WF) {k fallback : α} (h : m₁.contains k = false) : --- (m₁.diff m₂).getKeyD k fallback = fallback := by --- revert h --- simp_to_model [diff, getKeyD, contains] using List.getKeyD_diff_of_containsKey_eq_false_left - --- /- getKey! -/ --- theorem getKey!_diff [EquivBEq α] [LawfulHashable α] [Inhabited α] (h₁ : m₁.val.WF) --- (h₂ : m₂.val.WF) {k : α} : --- (m₁.diff m₂).getKey! k = if m₂.contains k then default else m₁.getKey! k := by --- simp_to_model [diff, getKey!, contains] using List.getKeyD_diff - --- theorem getKey!_diff_of_contains_eq_false_right [EquivBEq α] [LawfulHashable α] [Inhabited α] (h₁ : m₁.val.WF) --- (h₂ : m₂.val.WF) {k : α} (h : m₂.contains k = false) : --- (m₁.diff m₂).getKey! k = m₁.getKey! k := by --- revert h --- simp_to_model [diff, getKey!, contains] using List.getKeyD_diff_of_containsKey_eq_false_right - --- theorem getKey!_diff_of_contains_right [EquivBEq α] [LawfulHashable α] [Inhabited α] (h₁ : m₁.val.WF) --- (h₂ : m₂.val.WF) {k : α} (h : m₂.contains k) : --- (m₁.diff m₂).getKey! k = default := by --- revert h --- simp_to_model [diff, getKey!, contains] using List.getKeyD_diff_of_containsKey_right - --- theorem getKey!_diff_of_contains_eq_false_left [EquivBEq α] [LawfulHashable α] [Inhabited α] (h₁ : m₁.val.WF) --- (h₂ : m₂.val.WF) {k : α} (h : m₁.contains k = false) : --- (m₁.diff m₂).getKey! k = default := by --- revert h --- simp_to_model [diff, getKey!, contains] using List.getKeyD_diff_of_containsKey_eq_false_left - --- /- size -/ --- theorem size_diff_le_size_left [EquivBEq α] [LawfulHashable α] --- (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) : --- (m₁.diff m₂).1.size ≤ m₁.1.size := by --- simp_to_model [diff, size] using List.length_diff_le +theorem getKey?_diff_of_contains_eq_false_right [EquivBEq α] [LawfulHashable α] + (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} (h : m₂.contains k = false) : + (m₁.diff m₂).getKey? k = m₁.getKey? k := by + revert h + simp_to_model [contains, getKey?, diff] using List.getKey?_diff_of_containsKey_eq_false_right + +theorem getKey?_diff_of_contains_eq_false_left [EquivBEq α] [LawfulHashable α] + (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} (h : m₁.contains k = false) : + (m₁.diff m₂).getKey? k = none := by + revert h + simp_to_model [contains, getKey?, diff] using List.getKey?_diff_of_containsKey_eq_false_left + +theorem getKey?_diff_of_contains_right [EquivBEq α] [LawfulHashable α] + (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} (h : m₂.contains k) : + (m₁.diff m₂).getKey? k = none := by + revert h + simp_to_model [contains, getKey?, diff] using List.getKey?_diff_of_containsKey_right + +/- getKey -/ +theorem getKey_diff [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) + {k : α} {h_contains : (m₁.diff m₂).contains k} : + (m₁.diff m₂).getKey k h_contains = + m₁.getKey k ((contains_diff_iff h₁ h₂).1 h_contains).1 := by + simp_to_model [diff, contains, getKey] using List.getKey_diff + +/- getKeyD -/ +theorem getKeyD_diff [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) + (h₂ : m₂.val.WF) {k fallback : α} : + (m₁.diff m₂).getKeyD k fallback = + if m₂.contains k then fallback else m₁.getKeyD k fallback := by + simp_to_model [diff, getKeyD, contains] using List.getKeyD_diff + +theorem getKeyD_diff_of_contains_eq_false_right [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) + (h₂ : m₂.val.WF) {k fallback : α} (h : m₂.contains k = false) : + (m₁.diff m₂).getKeyD k fallback = m₁.getKeyD k fallback := by + revert h + simp_to_model [contains, diff, getKeyD] using List.getKeyD_diff_of_containsKey_eq_false_right + +theorem getKeyD_diff_of_contains_right [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) + (h₂ : m₂.val.WF) {k fallback : α} (h : m₂.contains k) : + (m₁.diff m₂).getKeyD k fallback = fallback := by + revert h + simp_to_model [diff, getKeyD, contains] using List.getKeyD_diff_of_containsKey_right --- theorem size_diff_eq_size_left_of_disjoint [EquivBEq α] [LawfulHashable α] --- (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) --- (h : ∀ (a : α), m₁.contains a → m₂.contains a = false) : --- (m₁.diff m₂).1.size = m₁.1.size := by --- revert h --- simp_to_model [diff, size, contains] using List.length_diff_eq_length_left +theorem getKeyD_diff_of_contains_eq_false_left [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) + (h₂ : m₂.val.WF) {k fallback : α} (h : m₁.contains k = false) : + (m₁.diff m₂).getKeyD k fallback = fallback := by + revert h + simp_to_model [diff, getKeyD, contains] using List.getKeyD_diff_of_containsKey_eq_false_left + +/- getKey! -/ +theorem getKey!_diff [EquivBEq α] [LawfulHashable α] [Inhabited α] (h₁ : m₁.val.WF) + (h₂ : m₂.val.WF) {k : α} : + (m₁.diff m₂).getKey! k = if m₂.contains k then default else m₁.getKey! k := by + simp_to_model [diff, getKey!, contains] using List.getKeyD_diff + +theorem getKey!_diff_of_contains_eq_false_right [EquivBEq α] [LawfulHashable α] [Inhabited α] (h₁ : m₁.val.WF) + (h₂ : m₂.val.WF) {k : α} (h : m₂.contains k = false) : + (m₁.diff m₂).getKey! k = m₁.getKey! k := by + revert h + simp_to_model [diff, getKey!, contains] using List.getKeyD_diff_of_containsKey_eq_false_right + +theorem getKey!_diff_of_contains_right [EquivBEq α] [LawfulHashable α] [Inhabited α] (h₁ : m₁.val.WF) + (h₂ : m₂.val.WF) {k : α} (h : m₂.contains k) : + (m₁.diff m₂).getKey! k = default := by + revert h + simp_to_model [diff, getKey!, contains] using List.getKeyD_diff_of_containsKey_right + +theorem getKey!_diff_of_contains_eq_false_left [EquivBEq α] [LawfulHashable α] [Inhabited α] (h₁ : m₁.val.WF) + (h₂ : m₂.val.WF) {k : α} (h : m₁.contains k = false) : + (m₁.diff m₂).getKey! k = default := by + revert h + simp_to_model [diff, getKey!, contains] using List.getKeyD_diff_of_containsKey_eq_false_left + +/- size -/ +theorem size_diff_le_size_left [EquivBEq α] [LawfulHashable α] + (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) : + (m₁.diff m₂).1.size ≤ m₁.1.size := by + simp_to_model [diff, size] using List.length_filter_le + +theorem size_diff_eq_size_left [EquivBEq α] [LawfulHashable α] + (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) + (h : ∀ (a : α), m₁.contains a → m₂.contains a = false) : + (m₁.diff m₂).1.size = m₁.1.size := by + revert h + simp_to_model [diff, size, contains] using List.length_diff_eq_length_left -- theorem size_diff_add_size_inter_eq_size_left [EquivBEq α] [LawfulHashable α] -- (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) : -- (m₁.diff m₂).1.size + (m₁.inter m₂).1.size = m₁.1.size := by -- simp_to_model [diff, inter, size] using List.size_diff_add_size_inter_eq_size_left --- /- isEmpty -/ --- @[simp] --- theorem isEmpty_diff_of_isEmpty_left [EquivBEq α] [LawfulHashable α] --- (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) (h : m₁.1.isEmpty) : --- (m₁.diff m₂).1.isEmpty = true := by --- revert h --- simp_to_model [isEmpty, diff] using List.isEmpty_diff_left - --- theorem isEmpty_diff_iff [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) : --- (m₁.diff m₂).1.isEmpty ↔ ∀ k, m₁.contains k → m₂.contains k := by --- simp_to_model [diff, contains, isEmpty] using List.isEmpty_diff_iff +/- isEmpty -/ +@[simp] +theorem isEmpty_diff_left [EquivBEq α] [LawfulHashable α] + (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) (h : m₁.1.isEmpty) : + (m₁.diff m₂).1.isEmpty = true := by + revert h + simp_to_model [isEmpty, diff] using List.isEmpty_diff_left --- theorem isEmpty_diff_of_subset [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) --- (h : ∀ k, m₁.contains k → m₂.contains k) : --- (m₁.diff m₂).1.isEmpty = true := by --- revert h --- simp_to_model [diff, contains, isEmpty] using List.isEmpty_diff_of_subset +theorem isEmpty_diff_iff [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) : + (m₁.diff m₂).1.isEmpty ↔ ∀ k, m₁.contains k → m₂.contains k := by + simp_to_model [diff, contains, isEmpty] using List.isEmpty_diff_iff end Diff @@ -3322,83 +3316,83 @@ namespace Const variable {β : Type v} {m₁ m₂ : Raw₀ α (fun _ => β)} --- /- get? -/ --- theorem get?_diff [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} : --- Const.get? (m₁.diff m₂) k = if m₂.contains k then none else Const.get? m₁ k := by --- simp_to_model [diff, Const.get?, contains] using List.getValue?_diff - --- theorem get?_diff_of_contains_eq_false_right [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) --- {k : α} (h : m₂.contains k = false) : --- Const.get? (m₁.diff m₂) k = Const.get? m₁ k := by --- revert h --- simp_to_model [diff, contains, Const.get?] using List.getValue?_diff_of_containsKey_eq_false_right - --- theorem get?_diff_of_contains_eq_false_left [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) --- {k : α} (h : m₁.contains k = false) : --- Const.get? (m₁.diff m₂) k = none := by --- revert h --- simp_to_model [diff, Const.get?, contains] using List.getValue?_diff_of_containsKey_eq_false_left - --- theorem get?_diff_of_contains_right [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) --- {k : α} (h : m₂.contains k) : --- Const.get? (m₁.diff m₂) k = none := by --- revert h --- simp_to_model [diff, Const.get?, contains] using List.getValue?_diff_of_containsKey_right - --- /- get -/ --- theorem get_diff [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) --- {k : α} {h_contains : (m₁.diff m₂).contains k} : --- Const.get (m₁.diff m₂) k h_contains = --- Const.get m₁ k ((contains_diff_iff h₁ h₂).1 h_contains).1 := by --- simp_to_model [diff, Const.get, contains] using List.getValue_diff - --- /- getD -/ --- theorem getD_diff [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) --- {k : α} {fallback : β} : --- Const.getD (m₁.diff m₂) k fallback = --- if m₂.contains k then fallback else Const.getD m₁ k fallback := by --- simp_to_model [diff, Const.getD, contains] using List.getValueD_diff - --- theorem getD_diff_of_contains_eq_false_right [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) --- {k : α} {fallback : β} (h : m₂.contains k = false) : --- Const.getD (m₁.diff m₂) k fallback = Const.getD m₁ k fallback := by --- revert h --- simp_to_model [diff, contains, Const.getD] using List.getValueD_diff_of_containsKey_eq_false_right - --- theorem getD_diff_of_contains_right [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) --- {k : α} {fallback : β} (h : m₂.contains k) : --- Const.getD (m₁.diff m₂) k fallback = fallback := by --- revert h --- simp_to_model [diff, Const.getD, contains] using List.getValueD_diff_of_containsKey_right - --- theorem getD_diff_of_contains_eq_false_left [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) --- {k : α} {fallback : β} (h : m₁.contains k = false) : --- Const.getD (m₁.diff m₂) k fallback = fallback := by --- revert h --- simp_to_model [diff, Const.getD, contains] using List.getValueD_diff_of_containsKey_eq_false_left - --- /- get! -/ --- theorem get!_diff [EquivBEq α] [LawfulHashable α] [Inhabited β] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} : --- Const.get! (m₁.diff m₂) k = if m₂.contains k then default else Const.get! m₁ k := by --- simp_to_model [diff, Const.get!, contains] using List.getValueD_diff - --- theorem get!_diff_of_contains_eq_false_right [EquivBEq α] [LawfulHashable α] [Inhabited β] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) --- {k : α} (h : m₂.contains k = false) : --- Const.get! (m₁.diff m₂) k = Const.get! m₁ k := by --- revert h --- simp_to_model [diff, contains, Const.get!] using List.getValueD_diff_of_containsKey_eq_false_right - --- theorem get!_diff_of_contains_right [EquivBEq α] [LawfulHashable α] [Inhabited β] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) --- {k : α} (h : m₂.contains k) : --- Const.get! (m₁.diff m₂) k = default := by --- revert h --- simp_to_model [diff, Const.get!, contains] using List.getValueD_diff_of_containsKey_right - --- theorem get!_diff_of_contains_eq_false_left [EquivBEq α] [LawfulHashable α] [Inhabited β] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) --- {k : α} (h : m₁.contains k = false) : --- Const.get! (m₁.diff m₂) k = default := by --- revert h --- simp_to_model [diff, Const.get!, contains] using List.getValueD_diff_of_containsKey_eq_false_left +/- get? -/ +theorem get?_diff [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} : + Const.get? (m₁.diff m₂) k = if m₂.contains k then none else Const.get? m₁ k := by + simp_to_model [diff, Const.get?, contains] using List.getValue?_diff + +theorem get?_diff_of_contains_eq_false_right [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) + {k : α} (h : m₂.contains k = false) : + Const.get? (m₁.diff m₂) k = Const.get? m₁ k := by + revert h + simp_to_model [diff, contains, Const.get?] using List.getValue?_diff_of_containsKey_eq_false_right + +theorem get?_diff_of_contains_eq_false_left [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) + {k : α} (h : m₁.contains k = false) : + Const.get? (m₁.diff m₂) k = none := by + revert h + simp_to_model [diff, Const.get?, contains] using List.getValue?_diff_of_containsKey_eq_false_left + +theorem get?_diff_of_contains_right [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) + {k : α} (h : m₂.contains k) : + Const.get? (m₁.diff m₂) k = none := by + revert h + simp_to_model [diff, Const.get?, contains] using List.getValue?_diff_of_containsKey_right + +/- get -/ +theorem get_diff [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) + {k : α} {h_contains : (m₁.diff m₂).contains k} : + Const.get (m₁.diff m₂) k h_contains = + Const.get m₁ k ((contains_diff_iff h₁ h₂).1 h_contains).1 := by + simp_to_model [diff, Const.get, contains] using List.getValue_diff + +/- getD -/ +theorem getD_diff [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) + {k : α} {fallback : β} : + Const.getD (m₁.diff m₂) k fallback = + if m₂.contains k then fallback else Const.getD m₁ k fallback := by + simp_to_model [diff, Const.getD, contains] using List.getValueD_diff + +theorem getD_diff_of_contains_eq_false_right [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) + {k : α} {fallback : β} (h : m₂.contains k = false) : + Const.getD (m₁.diff m₂) k fallback = Const.getD m₁ k fallback := by + revert h + simp_to_model [diff, contains, Const.getD] using List.getValueD_diff_of_containsKey_eq_false_right + +theorem getD_diff_of_contains_right [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) + {k : α} {fallback : β} (h : m₂.contains k) : + Const.getD (m₁.diff m₂) k fallback = fallback := by + revert h + simp_to_model [diff, Const.getD, contains] using List.getValueD_diff_of_containsKey_right + +theorem getD_diff_of_contains_eq_false_left [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) + {k : α} {fallback : β} (h : m₁.contains k = false) : + Const.getD (m₁.diff m₂) k fallback = fallback := by + revert h + simp_to_model [diff, Const.getD, contains] using List.getValueD_diff_of_containsKey_eq_false_left + +/- get! -/ +theorem get!_diff [EquivBEq α] [LawfulHashable α] [Inhabited β] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} : + Const.get! (m₁.diff m₂) k = if m₂.contains k then default else Const.get! m₁ k := by + simp_to_model [diff, Const.get!, contains] using List.getValueD_diff + +theorem get!_diff_of_contains_eq_false_right [EquivBEq α] [LawfulHashable α] [Inhabited β] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) + {k : α} (h : m₂.contains k = false) : + Const.get! (m₁.diff m₂) k = Const.get! m₁ k := by + revert h + simp_to_model [diff, contains, Const.get!] using List.getValueD_diff_of_containsKey_eq_false_right + +theorem get!_diff_of_contains_right [EquivBEq α] [LawfulHashable α] [Inhabited β] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) + {k : α} (h : m₂.contains k) : + Const.get! (m₁.diff m₂) k = default := by + revert h + simp_to_model [diff, Const.get!, contains] using List.getValueD_diff_of_containsKey_right + +theorem get!_diff_of_contains_eq_false_left [EquivBEq α] [LawfulHashable α] [Inhabited β] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) + {k : α} (h : m₁.contains k = false) : + Const.get! (m₁.diff m₂) k = default := by + revert h + simp_to_model [diff, Const.get!, contains] using List.getValueD_diff_of_containsKey_eq_false_left end Const diff --git a/src/Std/Data/Internal/List/Associative.lean b/src/Std/Data/Internal/List/Associative.lean index f245451d53b6..3de3ba53b314 100644 --- a/src/Std/Data/Internal/List/Associative.lean +++ b/src/Std/Data/Internal/List/Associative.lean @@ -5233,7 +5233,7 @@ theorem getValueCast?_filter_containsKey [BEq α] [LawfulBEq α] case h_2 => rfl exact dl₁ -theorem List.getValueCast?_diff [BEq α] [LawfulBEq α] +theorem getValueCast?_diff [BEq α] [LawfulBEq α] {l₁ l₂ : List ((a : α) × β a)} {k : α} (dl₁ : DistinctKeys l₁) : getValueCast? k (List.filter (fun p => !containsKey p.fst l₂) l₁) = @@ -5269,7 +5269,7 @@ theorem getValueCast?_diff_of_containsKey_eq_false_right [BEq α] [LawfulBEq α] containsKey k l₂ = false → getValueCast? k (List.filter (fun p => !containsKey p.fst l₂) l₁) = getValueCast? k l₁ := by intro h₁ - rw [@List.getValueCast?_diff, h₁] + rw [getValueCast?_diff, h₁] all_goals simp [dl₁] theorem getValueCast?_diff_of_containsKey_eq_false_left [BEq α] [LawfulBEq α] @@ -5278,17 +5278,17 @@ theorem getValueCast?_diff_of_containsKey_eq_false_left [BEq α] [LawfulBEq α] containsKey k l₁ = false → getValueCast? k (List.filter (fun p => !containsKey p.fst l₂) l₁) = none := by intro h₁ - rw [List.getValueCast?_diff] + rw [getValueCast?_diff] · simp [getValueCast?_eq_none h₁] · exact dl₁ -theorem List.getValueCast?_diff_of_containsKey_right [BEq α] [LawfulBEq α] +theorem getValueCast?_diff_of_containsKey_right [BEq α] [LawfulBEq α] {l₁ l₂ : List ((a : α) × β a)} {k : α} (dl₁ : DistinctKeys l₁) : containsKey k l₂ → getValueCast? k (List.filter (fun p => !containsKey p.fst l₂) l₁) = none := by intro h₁ - rw [List.getValueCast?_diff] + rw [getValueCast?_diff] · simp only [ite_eq_left_iff, Bool.not_eq_true] intro h₂ rw [h₂] at h₁ @@ -5568,7 +5568,7 @@ theorem containsKey_filter_containsKey_iff [BEq α] [EquivBEq α] {l₁ l₂ : L . simp . exact hl₁ -theorem List.containsKey_diff_eq_false [BEq α] [EquivBEq α] {l₁ l₂ : List ((a : α) × β a)} {hl₁ : DistinctKeys l₁} {k : α} : +theorem containsKey_diff_eq_false [BEq α] [EquivBEq α] {l₁ l₂ : List ((a : α) × β a)} {hl₁ : DistinctKeys l₁} {k : α} : (containsKey k (List.filter (fun p => !containsKey p.fst (l₂)) l₁)) = (containsKey k l₁ && !containsKey k l₂) := by rw [containsKey_eq_isSome_getEntry?, containsKey_eq_isSome_getEntry?] @@ -5583,21 +5583,21 @@ theorem List.containsKey_diff_eq_false [BEq α] [EquivBEq α] {l₁ l₂ : List apply List.beq_of_getEntry?_eq_some heq . exact hl₁ -theorem List.containsKey_diff_eq_false_iff [BEq α] [EquivBEq α] {l₁ l₂ : List ((a : α) × β a)} {hl₁ : DistinctKeys l₁} {k : α} : +theorem containsKey_diff_eq_false_iff [BEq α] [EquivBEq α] {l₁ l₂ : List ((a : α) × β a)} {hl₁ : DistinctKeys l₁} {k : α} : (containsKey k (List.filter (fun p => !containsKey p.fst l₂) l₁)) ↔ (containsKey k l₁ ∧ ¬ containsKey k l₂) := by - rw [List.containsKey_diff_eq_false] + rw [containsKey_diff_eq_false] . simp . exact hl₁ -theorem List.getValueCast_diff [BEq α] [LawfulBEq α] +theorem getValueCast_diff [BEq α] [LawfulBEq α] {l₁ l₂ : List ((a : α) × β a)} {k : α} (dl₁ : DistinctKeys l₁) {h₁ h₂} : getValueCast k (List.filter (fun p => !containsKey p.fst l₂) l₁) h₁ = getValueCast k l₁ h₂ := by suffices some (getValueCast k (List.filter (fun p => !containsKey p.fst l₂) l₁) h₁) = some (getValueCast k l₁ h₂) by injections simp only [← getValueCast?_eq_some_getValueCast] - rw [List.getValueCast?_diff dl₁] - rw [List.containsKey_diff_eq_false_iff] at h₁ + rw [getValueCast?_diff dl₁] + rw [containsKey_diff_eq_false_iff] at h₁ simp [h₁.2] exact dl₁ @@ -5646,7 +5646,7 @@ theorem containsKey_filter_containsKey_eq_false_of_containsKey_eq_false_left [BE theorem containsKey_diff_eq_false_of_containsKey_eq_false_left [BEq α] [EquivBEq α] {l₁ l₂ : List ((a : α) × β a)} {hl₁ : DistinctKeys l₁} {k : α} : containsKey k l₁ = false → containsKey k (List.filter (fun p => !containsKey p.fst l₂) l₁) = false := by intro h - rw [List.containsKey_diff_eq_false] + rw [containsKey_diff_eq_false] . simp [h] . exact hl₁ @@ -5660,11 +5660,11 @@ theorem containsKey_filter_containsKey_eq_false_of_containsKey_eq_false_right [B theorem containsKey_diff_of_containsKey_eq_false_right [BEq α] [EquivBEq α] {l₁ l₂ : List ((a : α) × β a)} {hl₁ : DistinctKeys l₁} {k : α} : containsKey k l₂ = true → containsKey k (List.filter (fun p => !containsKey p.fst l₂) l₁) = false := by intro h - rw [List.containsKey_diff_eq_false] + rw [containsKey_diff_eq_false] . simp [h] . exact hl₁ -theorem List.getValueCastD_diff [BEq α] [LawfulBEq α] +theorem getValueCastD_diff [BEq α] [LawfulBEq α] {l₁ l₂ : List ((a : α) × β a)} {k : α} {fallback : β k} (dl₁ : DistinctKeys l₁) : getValueCastD k (List.filter (fun p => !containsKey p.fst l₂) l₁) fallback = @@ -5678,10 +5678,10 @@ theorem List.getValueCastD_diff [BEq α] [LawfulBEq α] case isFalse h => rw [getValueCastD_eq_getValueCast?, getValueCastD_eq_getValueCast?] congr 1 - rw [List.getValueCast?_diff dl₁] + rw [getValueCast?_diff dl₁] simp [h] -theorem List.getValueCastD_diff_of_containsKey_eq_false_left [BEq α] [LawfulBEq α] +theorem getValueCastD_diff_of_containsKey_eq_false_left [BEq α] [LawfulBEq α] {l₁ l₂ : List ((a : α) × β a)} {k : α} {fallback} (dl₁ : DistinctKeys l₁) : containsKey k l₁ = false → getValueCastD k (List.filter (fun p => !containsKey p.fst l₂) l₁) fallback = fallback := by @@ -5691,7 +5691,7 @@ theorem List.getValueCastD_diff_of_containsKey_eq_false_left [BEq α] [LawfulBEq · exact dl₁ · exact h -theorem List.getKey?_diff [BEq α] [EquivBEq α] +theorem getKey?_diff [BEq α] [EquivBEq α] {l₁ l₂ : List ((a : α) × β a)} {k : α} (dl₁ : DistinctKeys l₁) : getKey? k (List.filter (fun p => !containsKey p.fst l₂) l₁) = @@ -5725,7 +5725,282 @@ theorem List.getKey?_diff [BEq α] [EquivBEq α] simp only [heq] · exact dl₁ -theorem List.getValueCastD_diff_of_containsKey_right [BEq α] [LawfulBEq α] +theorem getKey?_diff_of_containsKey_eq_false_right [BEq α] [EquivBEq α] + {l₁ l₂ : List ((a : α) × β a)} {k : α} + (dl₁ : DistinctKeys l₁) : + containsKey k l₂ = false → getKey? k (List.filter (fun p => !containsKey p.fst l₂) l₁) = + getKey? k l₁ := by + intro h + rw [getKey?_diff, h] + · simp only [Bool.false_eq_true, ↓reduceIte] + · exact dl₁ + +theorem getKey?_diff_of_containsKey_eq_false_left [BEq α] [EquivBEq α] + {l₁ l₂ : List ((a : α) × β a)} {k : α} + (dl₁ : DistinctKeys l₁) : + containsKey k l₁ = false → getKey? k (List.filter (fun p => !containsKey p.fst l₂) l₁) = none := by + intro h + rw [getKey?_diff] + · split + · rfl + · rw [getKey?_eq_none h] + · exact dl₁ + +theorem getKey?_diff_of_containsKey_right [BEq α] [EquivBEq α] + {l₁ l₂ : List ((a : α) × β a)} {k : α} + (dl₁ : DistinctKeys l₁) : + containsKey k l₂ = true → getKey? k (List.filter (fun p => !containsKey p.fst l₂) l₁) = none := by + intro h + rw [getKey?_diff, h] + · simp only [↓reduceIte] + · exact dl₁ + +theorem getKey_diff [BEq α] [EquivBEq α] + {l₁ l₂ : List ((a : α) × β a)} {k : α} + (dl₁ : DistinctKeys l₁) {h₁ h₂} : + getKey k (List.filter (fun p => !containsKey p.fst l₂) l₁) h₁ = getKey k l₁ h₂ := by + suffices some (getKey k (List.filter (fun p => !containsKey p.fst l₂) l₁) h₁) = some (getKey k l₁ h₂) by + injections + simp only [← getKey?_eq_some_getKey] + apply getKey?_diff_of_containsKey_eq_false_right dl₁ + · rw [containsKey_diff_eq_false_iff] at h₁ + · simp at h₁ + exact h₁.2 + · exact dl₁ + +theorem getKeyD_diff [BEq α] [EquivBEq α] + {l₁ l₂ : List ((a : α) × β a)} {k : α} {fallback : α} + (dl₁ : DistinctKeys l₁) : + getKeyD k (List.filter (fun p => !containsKey p.fst l₂) l₁) fallback = + if containsKey k l₂ = true then fallback else getKeyD k l₁ fallback := by + split + case isTrue h => + apply getKeyD_eq_fallback + apply containsKey_diff_of_containsKey_eq_false_right + · exact dl₁ + · exact h + case isFalse h => + rw [getKeyD_eq_getKey?, getKeyD_eq_getKey?] + congr 1 + apply getKey?_diff_of_containsKey_eq_false_right dl₁ + simp only [h] + +theorem getKeyD_diff_of_containsKey_eq_false_right [BEq α] [EquivBEq α] + {l₁ l₂ : List ((a : α) × β a)} {k : α} {fallback : α} + (dl₁ : DistinctKeys l₁) : + containsKey k l₂ = false → getKeyD k (List.filter (fun p => !containsKey p.fst l₂) l₁) fallback = + getKeyD k l₁ fallback := by + intro h + rw [getKeyD_eq_getKey?, getKeyD_eq_getKey?] + congr 1 + apply getKey?_diff_of_containsKey_eq_false_right dl₁ h + +theorem getKeyD_diff_of_containsKey_right [BEq α] [EquivBEq α] + {l₁ l₂ : List ((a : α) × β a)} {k : α} {fallback : α} + (dl₁ : DistinctKeys l₁) : + containsKey k l₂ = true → getKeyD k (List.filter (fun p => !containsKey p.fst l₂) l₁) fallback = fallback := by + intro h + apply getKeyD_eq_fallback + apply containsKey_diff_of_containsKey_eq_false_right + · exact dl₁ + · exact h + +theorem getKeyD_diff_of_containsKey_eq_false_left [BEq α] [EquivBEq α] + {l₁ l₂ : List ((a : α) × β a)} {k : α} {fallback : α} + (dl₁ : DistinctKeys l₁) : + containsKey k l₁ = false → getKeyD k (List.filter (fun p => !containsKey p.fst l₂) l₁) fallback = fallback := by + intro h + apply getKeyD_eq_fallback + apply containsKey_diff_eq_false_of_containsKey_eq_false_left + · exact dl₁ + · exact h + +theorem length_diff_eq_length_left [BEq α] [EquivBEq α] + {l₁ l₂ : List ((a : α) × β a)} (hl₁ : DistinctKeys l₁) + (w : ∀ (a : α), containsKey a l₁ → containsKey a l₂ = false) : + (l₁.filter fun p => !containsKey p.fst l₂).length = l₁.length := by + rw [← List.countP_eq_length_filter] + induction l₁ + case nil => simp + case cons h t ih => + simp + rw [List.countP_cons_of_pos] + · rw [List.distinctKeys_cons_iff] at hl₁ + specialize ih hl₁.1 + simp only [Nat.add_right_cancel_iff] + apply ih + intro a hc + apply w + rw [containsKey_cons] + simp [hc] + · simp only [Bool.not_eq_true'] + apply w + rw [containsKey_cons, BEq.rfl, Bool.true_or] + +theorem isEmpty_diff_left [BEq α] [EquivBEq α] + {l₁ l₂ : List ((a : α) × β a)} : + l₁.isEmpty → (List.filter (fun p => !containsKey p.fst l₂) l₁).isEmpty := by + intro hyp + simp at hyp + simp [hyp] + +theorem isEmpty_diff_iff [BEq α] [EquivBEq α] + {l₁ l₂ : List ((a : α) × β a)} (dl₁ : DistinctKeys l₁) : + (List.filter (fun p => !containsKey p.fst l₂) l₁).isEmpty = true ↔ + ∀ (k : α), containsKey k l₁ = true → containsKey k l₂ = true := by + constructor + case mpr => + intro hyp + rw [List.isEmpty_iff] + simp only [List.filter_eq_nil_iff, Bool.not_eq_true] + intro ⟨k, v⟩ mem + simp only [Bool.not_eq_eq_eq_not, Bool.not_false] + apply hyp + apply List.containsKey_of_mem mem + case mp => + intro hyp k mem + rw [List.isEmpty_iff] at hyp + induction l₁ with + | nil => simp at mem + | cons h t ih => + rw [containsKey_cons] at mem + simp only [Bool.or_eq_true] at mem + rw [List.filter_cons] at hyp + split at hyp + case cons.isTrue _ => + simp at hyp + case cons.isFalse heq => + simp only [Bool.not_eq_eq_eq_not, Bool.not_true] at heq + cases mem + case inl heq2 => + rw [containsKey_congr (PartialEquivBEq.symm heq2)] + simp [heq] + case inr heq2 => + apply ih + · rw [List.distinctKeys_cons_iff] at dl₁ + exact dl₁.1 + · exact hyp + · exact heq2 + +theorem getValue?_diff {β : Type v} [BEq α] [EquivBEq α] + {l₁ l₂ : List ((_ : α) × β)} {k : α} + (dl₁ : DistinctKeys l₁) : + getValue? k (List.filter (fun p => !containsKey p.fst l₂) l₁) = + if containsKey k l₂ = true then none else getValue? k l₁ := by + simp only [getValue?_eq_getEntry?] + split + case isTrue h => + simp only [Option.map_eq_none_iff, getEntry?_eq_none] + apply containsKey_diff_of_containsKey_eq_false_right + · exact dl₁ + · exact h + case isFalse h => + congr 1 + rw [getEntry?_filter] + · simp only [Option.filter] + split + case h_1 _ val heq => + simp [heq] + have := PartialEquivBEq.symm <| @beq_of_getEntry?_eq_some α (fun _ => β) _ l₁ k val heq + rw [← containsKey_congr this] + simp at h + exact h + case h_2 _ heq => + simp [heq] + · exact dl₁ + +theorem getValue?_diff_of_containsKey_eq_false_right {β : Type v} [BEq α] [EquivBEq α] + {l₁ l₂ : List ((_ : α) × β)} {k : α} + (dl₁ : DistinctKeys l₁) : + containsKey k l₂ = false → getValue? k (List.filter (fun p => !containsKey p.fst l₂) l₁) = + getValue? k l₁ := by + intro h + rw [getValue?_diff, h] + · simp only [Bool.false_eq_true, ↓reduceIte] + · exact dl₁ + +theorem getValue?_diff_of_containsKey_eq_false_left {β : Type v} [BEq α] [EquivBEq α] + {l₁ l₂ : List ((_ : α) × β)} {k : α} + (dl₁ : DistinctKeys l₁) : + containsKey k l₁ = false → getValue? k (List.filter (fun p => !containsKey p.fst l₂) l₁) = none := by + intro h + rw [getValue?_diff] + · split + · rfl + · rw [@getValue?_eq_none α β _ l₁ k] + exact h + · exact dl₁ + +theorem getValue?_diff_of_containsKey_right {β : Type v} [BEq α] [EquivBEq α] + {l₁ l₂ : List ((_ : α) × β)} {k : α} + (dl₁ : DistinctKeys l₁) : + containsKey k l₂ = true → getValue? k (List.filter (fun p => !containsKey p.fst l₂) l₁) = none := by + intro h + rw [getValue?_diff, h] + · simp only [↓reduceIte] + · exact dl₁ + +theorem getValue_diff {β : Type v} [BEq α] [EquivBEq α] + {l₁ l₂ : List ((_ : α) × β)} {k : α} + (dl₁ : DistinctKeys l₁) {h₁ h₂} : + getValue k (List.filter (fun p => !containsKey p.fst l₂) l₁) h₁ = getValue k l₁ h₂ := by + suffices some (getValue k (List.filter (fun p => !containsKey p.fst l₂) l₁) h₁) = some (getValue k l₁ h₂) by + injections + simp only [← getValue?_eq_some_getValue] + apply getValue?_diff_of_containsKey_eq_false_right dl₁ + · rw [containsKey_diff_eq_false_iff] at h₁ + · simp at h₁ + exact h₁.2 + · exact dl₁ + +theorem getValueD_diff {β : Type v} [BEq α] [EquivBEq α] + {l₁ l₂ : List ((_ : α) × β)} {k : α} {fallback : β} + (dl₁ : DistinctKeys l₁) : + getValueD k (List.filter (fun p => !containsKey p.fst l₂) l₁) fallback = + if containsKey k l₂ = true then fallback else getValueD k l₁ fallback := by + split + case isTrue h => + apply getValueD_eq_fallback + apply containsKey_diff_of_containsKey_eq_false_right + · exact dl₁ + · exact h + case isFalse h => + rw [getValueD_eq_getValue?, getValueD_eq_getValue?] + congr 1 + apply getValue?_diff_of_containsKey_eq_false_right dl₁ + simp only [h] + +theorem getValueD_diff_of_containsKey_eq_false_right {β : Type v} [BEq α] [EquivBEq α] + {l₁ l₂ : List ((_ : α) × β)} {k : α} {fallback : β} + (dl₁ : DistinctKeys l₁) : + containsKey k l₂ = false → getValueD k (List.filter (fun p => !containsKey p.fst l₂) l₁) fallback = + getValueD k l₁ fallback := by + intro h + rw [getValueD_eq_getValue?, getValueD_eq_getValue?] + congr 1 + apply getValue?_diff_of_containsKey_eq_false_right dl₁ h + +theorem getValueD_diff_of_containsKey_right {β : Type v} [BEq α] [EquivBEq α] + {l₁ l₂ : List ((_ : α) × β)} {k : α} {fallback : β} + (dl₁ : DistinctKeys l₁) : + containsKey k l₂ = true → getValueD k (List.filter (fun p => !containsKey p.fst l₂) l₁) fallback = fallback := by + intro h + apply getValueD_eq_fallback + apply containsKey_diff_of_containsKey_eq_false_right + · exact dl₁ + · exact h + +theorem getValueD_diff_of_containsKey_eq_false_left {β : Type v} [BEq α] [EquivBEq α] + {l₁ l₂ : List ((_ : α) × β)} {k : α} {fallback : β} + (dl₁ : DistinctKeys l₁) : + containsKey k l₁ = false → getValueD k (List.filter (fun p => !containsKey p.fst l₂) l₁) fallback = fallback := by + intro h + apply getValueD_eq_fallback + apply containsKey_diff_eq_false_of_containsKey_eq_false_left + · exact dl₁ + · exact h + +theorem getValueCastD_diff_of_containsKey_right [BEq α] [LawfulBEq α] {l₁ l₂ : List ((a : α) × β a)} {k : α} {fallback} (dl₁ : DistinctKeys l₁) : containsKey k l₂ = true → getValueCastD k (List.filter (fun p => !containsKey p.fst l₂) l₁) fallback = fallback := by @@ -5839,7 +6114,7 @@ theorem getValueCastD_filter_containsKey_of_containsKey_eq_false_right [BEq α] . exact dl₁ . exact h -theorem List.getValueCastD_diff_of_containsKey_eq_false_right [BEq α] [LawfulBEq α] +theorem getValueCastD_diff_of_containsKey_eq_false_right [BEq α] [LawfulBEq α] {l₁ l₂ : List ((a : α) × β a)} {k : α} {fallback} (dl₁ : DistinctKeys l₁) : containsKey k l₂ = false → @@ -5848,7 +6123,7 @@ theorem List.getValueCastD_diff_of_containsKey_eq_false_right [BEq α] [LawfulBE intro h rw [getValueCastD_eq_getValueCast?, getValueCastD_eq_getValueCast?] congr 1 - rw [List.getValueCast?_diff dl₁] + rw [getValueCast?_diff dl₁] simp [h] theorem getKeyD_filter_containsKey [BEq α] [EquivBEq α] From 07b32d68556a57933de38885d7f6475d4ffdbb30 Mon Sep 17 00:00:00 2001 From: Wojciech Rozowski Date: Mon, 17 Nov 2025 12:53:31 +0100 Subject: [PATCH 69/88] Finish remaining intersection lemma --- src/Std/Data/DHashMap/Internal/RawLemmas.lean | 8 ++++---- src/Std/Data/Internal/List/Associative.lean | 20 +++++++++++++++++++ 2 files changed, 24 insertions(+), 4 deletions(-) diff --git a/src/Std/Data/DHashMap/Internal/RawLemmas.lean b/src/Std/Data/DHashMap/Internal/RawLemmas.lean index 2fe0b6ce661f..dcfb5afd1549 100644 --- a/src/Std/Data/DHashMap/Internal/RawLemmas.lean +++ b/src/Std/Data/DHashMap/Internal/RawLemmas.lean @@ -3293,10 +3293,10 @@ theorem size_diff_eq_size_left [EquivBEq α] [LawfulHashable α] revert h simp_to_model [diff, size, contains] using List.length_diff_eq_length_left --- theorem size_diff_add_size_inter_eq_size_left [EquivBEq α] [LawfulHashable α] --- (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) : --- (m₁.diff m₂).1.size + (m₁.inter m₂).1.size = m₁.1.size := by --- simp_to_model [diff, inter, size] using List.size_diff_add_size_inter_eq_size_left +theorem size_diff_add_size_inter_eq_size_left [EquivBEq α] [LawfulHashable α] + (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) : + (m₁.diff m₂).1.size + (m₁.inter m₂).1.size = m₁.1.size := by + simp_to_model [diff, inter, size] using List.size_diff_add_size_inter_eq_size_left /- isEmpty -/ @[simp] diff --git a/src/Std/Data/Internal/List/Associative.lean b/src/Std/Data/Internal/List/Associative.lean index 3de3ba53b314..6ee909f06042 100644 --- a/src/Std/Data/Internal/List/Associative.lean +++ b/src/Std/Data/Internal/List/Associative.lean @@ -5837,6 +5837,26 @@ theorem length_diff_eq_length_left [BEq α] [EquivBEq α] apply w rw [containsKey_cons, BEq.rfl, Bool.true_or] +theorem size_diff_add_size_inter_eq_size_left [BEq α] [EquivBEq α] + {l₁ l₂ : List ((a : α) × β a)} (dl₁ : DistinctKeys l₁) : + (List.filter (fun p => !containsKey p.fst l₂) l₁).length + + (List.filter (fun p => containsKey p.fst l₂) l₁).length = l₁.length := by + induction l₁ with + | nil => simp + | cons h t ih => + rw [List.distinctKeys_cons_iff] at dl₁ + specialize ih dl₁.1 + simp only [List.filter_cons] + split + case isTrue hc => + simp only [Bool.not_eq_eq_eq_not, Bool.not_true] at hc + simp only [List.length_cons, hc, Bool.false_eq_true, ↓reduceIte] + omega + case isFalse hc => + simp only [Bool.not_eq_eq_eq_not, Bool.not_true, Bool.not_eq_false] at hc + simp only [hc, ↓reduceIte, List.length_cons] + omega + theorem isEmpty_diff_left [BEq α] [EquivBEq α] {l₁ l₂ : List ((a : α) × β a)} : l₁.isEmpty → (List.filter (fun p => !containsKey p.fst l₂) l₁).isEmpty := by From 3e6ed8ddc487894b859d35977e0a306b29dd6f7a Mon Sep 17 00:00:00 2001 From: Wojciech Rozowski Date: Mon, 17 Nov 2025 13:20:37 +0100 Subject: [PATCH 70/88] Rename sublemmas --- src/Std/Data/DHashMap/Internal/RawLemmas.lean | 27 ++++++++--------- src/Std/Data/HashSet/Raw.lean | 2 -- src/Std/Data/Internal/List/Associative.lean | 30 +++++++++---------- 3 files changed, 28 insertions(+), 31 deletions(-) diff --git a/src/Std/Data/DHashMap/Internal/RawLemmas.lean b/src/Std/Data/DHashMap/Internal/RawLemmas.lean index 61c3a32df20a..f19758d6b0d7 100644 --- a/src/Std/Data/DHashMap/Internal/RawLemmas.lean +++ b/src/Std/Data/DHashMap/Internal/RawLemmas.lean @@ -3215,7 +3215,6 @@ theorem isEmpty_inter_iff [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) simp_to_model [inter, contains, isEmpty] using List.isEmpty_filter_containsKey_iff end Inter - namespace Const variable {β : Type v} {m₁ m₂ : Raw₀ α (fun _ => β)} @@ -3224,81 +3223,81 @@ variable {β : Type v} {m₁ m₂ : Raw₀ α (fun _ => β)} theorem get?_inter [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} : Const.get? (m₁.inter m₂) k = if m₂.contains k then Const.get? m₁ k else none := by - simp_to_model [inter, Const.get?, contains] using List.getValue?_filter_containsKey_eq_if_containsKey + simp_to_model [inter, Const.get?, contains] using List.getValue?_filter_containsKey theorem get?_inter_of_contains_right [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} (h : m₂.contains k) : Const.get? (m₁.inter m₂) k = Const.get? m₁ k := by revert h - simp_to_model [inter, Const.get?, contains] using List.getValue?_filter_containsKey_eq_containsKey + simp_to_model [inter, Const.get?, contains] using List.getValue?_filter_containsKey_of_containsKey_right theorem get?_inter_of_contains_eq_false_left [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} (h : m₁.contains k = false) : Const.get? (m₁.inter m₂) k = none := by revert h - simp_to_model [inter, Const.get?, contains] using List.getValue?_filter_of_not_contains + simp_to_model [inter, Const.get?, contains] using List.getValue?_filter_containsKey_of_containsKey_eq_false_left theorem get?_inter_of_contains_eq_false_right [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} (h : m₂.contains k = false) : Const.get? (m₁.inter m₂) k = none := by revert h - simp_to_model [inter, Const.get?, contains] using getValue?_filter_of_containsKey_eq_false_right + simp_to_model [inter, Const.get?, contains] using getValue?_filter_containsKey_of_containsKey_eq_false_right @[simp] theorem get_inter [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} {h_contains : (m₁.inter m₂).contains k} : Const.get (m₁.inter m₂) k h_contains = Const.get m₁ k ((contains_inter_iff h₁ h₂).1 h_contains).1 := by - simp_to_model [inter, Const.get, contains] using List.getValue_filter_containsKey_eq_containsKey + simp_to_model [inter, Const.get, contains] using List.getValue_filter_containsKey /- getD -/ theorem getD_inter [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} {fallback : β} : Const.getD (m₁.inter m₂) k fallback = if m₂.contains k then Const.getD m₁ k fallback else fallback := by - simp_to_model [inter, Const.getD, contains] using List.getValueD_filter_containsKey_eq_if_containsKey + simp_to_model [inter, Const.getD, contains] using List.getValueD_filter_containsKey theorem getD_inter_of_contains_right [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} {fallback : β} (h : m₂.contains k) : Const.getD (m₁.inter m₂) k fallback = Const.getD m₁ k fallback := by revert h - simp_to_model [inter, Const.getD, contains] using List.getValueD_filter_containsKey_eq_containsKey + simp_to_model [inter, Const.getD, contains] using List.getValueD_filter_containsKey_of_containsKey_right theorem getD_inter_of_contains_eq_false_right [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} {fallback : β} (h : m₂.contains k = false) : Const.getD (m₁.inter m₂) k fallback = fallback := by revert h - simp_to_model [inter, Const.getD, contains] using getValueD_filter_of_containsKey_eq_false_right + simp_to_model [inter, Const.getD, contains] using getValueD_filter_containsKey_of_containsKey_eq_false_right theorem getD_inter_of_contains_eq_false_left [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} {fallback : β} (h : m₁.contains k = false) : Const.getD (m₁.inter m₂) k fallback = fallback := by revert h - simp_to_model [inter, Const.getD, contains] using getValueD_filter_of_containsKey_eq_false_left + simp_to_model [inter, Const.getD, contains] using getValueD_filter_containsKey_of_containsKey_eq_false_left /- get! -/ theorem get!_inter [EquivBEq α] [LawfulHashable α] [Inhabited β] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} : Const.get! (m₁.inter m₂) k = if m₂.contains k then Const.get! m₁ k else default := by - simp_to_model [inter, Const.get!, contains] using List.getValueD_filter_containsKey_eq_if_containsKey + simp_to_model [inter, Const.get!, contains] using List.getValueD_filter_containsKey theorem get!_inter_of_contains_right [EquivBEq α] [LawfulHashable α] [Inhabited β] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} (h : m₂.contains k) : Const.get! (m₁.inter m₂) k = Const.get! m₁ k := by revert h - simp_to_model [inter, Const.get!, contains] using List.getValueD_filter_containsKey_eq_containsKey + simp_to_model [inter, Const.get!, contains] using List.getValueD_filter_containsKey_of_containsKey_right theorem get!_inter_of_contains_eq_false_right [EquivBEq α] [LawfulHashable α] [Inhabited β] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} (h : m₂.contains k = false) : Const.get! (m₁.inter m₂) k = default := by revert h - simp_to_model [inter, Const.get!, contains] using getValueD_filter_of_containsKey_eq_false_right + simp_to_model [inter, Const.get!, contains] using getValueD_filter_containsKey_of_containsKey_eq_false_right theorem get!_inter_of_contains_eq_false_left [EquivBEq α] [LawfulHashable α] [Inhabited β] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} (h : m₁.contains k = false) : Const.get! (m₁.inter m₂) k = default := by revert h - simp_to_model [inter, Const.get!, contains] using getValueD_filter_of_containsKey_eq_false_left + simp_to_model [inter, Const.get!, contains] using getValueD_filter_containsKey_of_containsKey_eq_false_left end Const diff --git a/src/Std/Data/HashSet/Raw.lean b/src/Std/Data/HashSet/Raw.lean index 836c09eb2c4e..49824e3e7fb4 100644 --- a/src/Std/Data/HashSet/Raw.lean +++ b/src/Std/Data/HashSet/Raw.lean @@ -253,8 +253,6 @@ section Unverified /-- Check if any element satisfies the predicate, short-circuiting if a predicate succeeds. -/ @[inline] def any (m : Raw α) (p : α → Bool) : Bool := m.inner.any (fun x _ => p x) -section Unverified - /-! We currently do not provide lemmas for the functions below. -/ /-- diff --git a/src/Std/Data/Internal/List/Associative.lean b/src/Std/Data/Internal/List/Associative.lean index a474550b7f9a..3c00f21ee143 100644 --- a/src/Std/Data/Internal/List/Associative.lean +++ b/src/Std/Data/Internal/List/Associative.lean @@ -5480,7 +5480,7 @@ theorem containsKey_filter_containsKey_eq_false_of_containsKey_eq_false_right [B . simp [h] . exact hl₁ -theorem getValue?_filter_containsKey_eq_if_containsKey {β : Type v} [BEq α] [EquivBEq α] +theorem List.getValue?_filter_containsKey {β : Type v} [BEq α] [EquivBEq α] {l₁ l₂ : List ((_ : α) × β)} {k : α} (dl₁ : DistinctKeys l₁) : getValue? k (List.filter (fun p => containsKey p.fst l₂) l₁) = @@ -5507,34 +5507,34 @@ theorem getValue?_filter_containsKey_eq_if_containsKey {β : Type v} [BEq α] [E . exact dl₁ . exact h -theorem getValue?_filter_containsKey_eq_containsKey {β : Type v} [BEq α] [EquivBEq α] +theorem List.getValue?_filter_containsKey_of_containsKey_right {β : Type v} [BEq α] [EquivBEq α] {l₁ l₂ : List ((_ : α) × β)} {k : α} (dl₁ : DistinctKeys l₁) : containsKey k l₂ → getValue? k (List.filter (fun p => containsKey p.fst l₂) l₁) = getValue? k l₁ := by intro h - rw [getValue?_filter_containsKey_eq_if_containsKey, h] + rw [List.getValue?_filter_containsKey, h] · simp only [↓reduceIte] · exact dl₁ -theorem getValue?_filter_of_not_contains {β : Type v} [BEq α] [EquivBEq α] +theorem List.getValue?_filter_containsKey_of_containsKey_eq_false_left {β : Type v} [BEq α] [EquivBEq α] {l₁ l₂ : List ((_ : α) × β)} {k : α} (dl₁ : DistinctKeys l₁) : containsKey k l₁ = false → getValue? k (List.filter (fun p => containsKey p.fst l₂) l₁) = none := by intro h - rw [getValue?_filter_containsKey_eq_if_containsKey] + rw [List.getValue?_filter_containsKey] · split · rw [@getValue?_eq_none α β _ l₁ k] exact h · rfl · exact dl₁ -theorem getValue?_filter_of_containsKey_eq_false_right {β : Type v} [BEq α] [EquivBEq α] +theorem getValue?_filter_containsKey_of_containsKey_eq_false_right {β : Type v} [BEq α] [EquivBEq α] {l₁ l₂ : List ((_ : α) × β)} {k : α} (dl₁ : DistinctKeys l₁) : containsKey k l₂ = false → getValue? k (List.filter (fun p => containsKey p.fst l₂) l₁) = none := by intro h - rw [getValue?_filter_containsKey_eq_if_containsKey, h] + rw [List.getValue?_filter_containsKey, h] · simp only [Bool.false_eq_true, ↓reduceIte] · exact dl₁ @@ -5630,19 +5630,19 @@ theorem getKeyD_filter_containsKey_of_containsKey_eq_false_right [BEq α] [Equiv · exact dl₁ · exact h -theorem getValue_filter_containsKey_eq_containsKey {β : Type v} [BEq α] [EquivBEq α] +theorem List.getValue_filter_containsKey {β : Type v} [BEq α] [EquivBEq α] {l₁ l₂ : List ((_ : α) × β)} {k : α} (dl₁ : DistinctKeys l₁) {h₁ h₂} : getValue k (List.filter (fun p => containsKey p.fst l₂) l₁) h₁ = getValue k l₁ h₂ := by suffices some (getValue k (List.filter (fun p => containsKey p.fst l₂) l₁) h₁) = some (getValue k l₁ h₂) by injections simp only [← getValue?_eq_some_getValue] - apply getValue?_filter_containsKey_eq_containsKey dl₁ + apply List.getValue?_filter_containsKey_of_containsKey_right dl₁ · rw [containsKey_filter_containsKey_iff] at h₁ · exact h₁.2 · exact dl₁ -theorem getValueD_filter_containsKey_eq_if_containsKey {β : Type v} [BEq α] [EquivBEq α] +theorem List.getValueD_filter_containsKey {β : Type v} [BEq α] [EquivBEq α] {l₁ l₂ : List ((_ : α) × β)} {k : α} {fallback : β} (dl₁ : DistinctKeys l₁) : getValueD k (List.filter (fun p => containsKey p.fst l₂) l₁) fallback = @@ -5651,14 +5651,14 @@ theorem getValueD_filter_containsKey_eq_if_containsKey {β : Type v} [BEq α] [E case isTrue h => rw [getValueD_eq_getValue?, getValueD_eq_getValue?] congr 1 - apply getValue?_filter_containsKey_eq_containsKey dl₁ h + apply List.getValue?_filter_containsKey_of_containsKey_right dl₁ h case isFalse h => apply getValueD_eq_fallback apply containsKey_filter_containsKey_eq_false_of_containsKey_eq_false_right · exact dl₁ · simp only [h] -theorem getValueD_filter_containsKey_eq_containsKey {β : Type v} [BEq α] [EquivBEq α] +theorem List.getValueD_filter_containsKey_of_containsKey_right {β : Type v} [BEq α] [EquivBEq α] {l₁ l₂ : List ((_ : α) × β)} {k : α} {fallback : β} (dl₁ : DistinctKeys l₁) : containsKey k l₂ → getValueD k (List.filter (fun p => containsKey p.fst l₂) l₁) fallback = @@ -5666,9 +5666,9 @@ theorem getValueD_filter_containsKey_eq_containsKey {β : Type v} [BEq α] [Equi intro h rw [getValueD_eq_getValue?, getValueD_eq_getValue?] congr 1 - apply getValue?_filter_containsKey_eq_containsKey dl₁ h + apply List.getValue?_filter_containsKey_of_containsKey_right dl₁ h -theorem getValueD_filter_of_containsKey_eq_false_left {β : Type v} [BEq α] [EquivBEq α] +theorem getValueD_filter_containsKey_of_containsKey_eq_false_left {β : Type v} [BEq α] [EquivBEq α] {l₁ l₂ : List ((_ : α) × β)} {k : α} {fallback : β} (dl₁ : DistinctKeys l₁) : containsKey k l₁ = false → getValueD k (List.filter (fun p => containsKey p.fst l₂) l₁) fallback = fallback := by @@ -5678,7 +5678,7 @@ theorem getValueD_filter_of_containsKey_eq_false_left {β : Type v} [BEq α] [Eq · exact dl₁ · exact h -theorem getValueD_filter_of_containsKey_eq_false_right {β : Type v} [BEq α] [EquivBEq α] +theorem getValueD_filter_containsKey_of_containsKey_eq_false_right {β : Type v} [BEq α] [EquivBEq α] {l₁ l₂ : List ((_ : α) × β)} {k : α} {fallback : β} (dl₁ : DistinctKeys l₁) : containsKey k l₂ = false → getValueD k (List.filter (fun p => containsKey p.fst l₂) l₁) fallback = fallback := by From 1d458402c0864ff3030aaf9879142e2e6a8ce2d6 Mon Sep 17 00:00:00 2001 From: Wojciech Rozowski Date: Mon, 17 Nov 2025 16:19:32 +0100 Subject: [PATCH 71/88] Rename `eraseMany` to `eraseManyEntries` --- src/Std/Data/DHashMap/Internal/Defs.lean | 4 ++-- src/Std/Data/DHashMap/Internal/Model.lean | 8 ++++---- src/Std/Data/DHashMap/Internal/WF.lean | 20 ++++++++++---------- src/Std/Data/DHashMap/Raw.lean | 4 ++-- 4 files changed, 18 insertions(+), 18 deletions(-) diff --git a/src/Std/Data/DHashMap/Internal/Defs.lean b/src/Std/Data/DHashMap/Internal/Defs.lean index b027ec0e6407..0255c6340ab5 100644 --- a/src/Std/Data/DHashMap/Internal/Defs.lean +++ b/src/Std/Data/DHashMap/Internal/Defs.lean @@ -452,7 +452,7 @@ def insertMany {ρ : Type w} [ForIn Id ρ ((a : α) × β a)] [BEq α] [Hashable return r /-- Internal implementation detail of the hash map -/ -def eraseMany {ρ : Type w} [ForIn Id ρ ((a : α) × β a)] [BEq α] [Hashable α] +def eraseManyEntries {ρ : Type w} [ForIn Id ρ ((a : α) × β a)] [BEq α] [Hashable α] (m : Raw₀ α β) (l : ρ) : { m' : Raw₀ α β // ∀ (P : Raw₀ α β → Prop), (∀ {m'' a}, P m'' → P (m''.erase a)) → P m → P m' } := Id.run do let mut r : { m' : Raw₀ α β // ∀ (P : Raw₀ α β → Prop), @@ -492,7 +492,7 @@ def inter [BEq α] [Hashable α] (m₁ m₂ : Raw₀ α β) : Raw₀ α β := /-- Internal implementation detail of the hash map -/ @[inline] def diff [BEq α] [Hashable α] (m₁ m₂ : Raw₀ α β) : Raw₀ α β := - if m₁.1.size ≤ m₂.1.size then m₁.filter (fun k _ => !m₂.contains k) else (eraseMany m₁ m₂.1).1 + if m₁.1.size ≤ m₂.1.size then m₁.filter (fun k _ => !m₂.contains k) else (eraseManyEntries m₁ m₂.1).1 section diff --git a/src/Std/Data/DHashMap/Internal/Model.lean b/src/Std/Data/DHashMap/Internal/Model.lean index d4ecf9175410..c91efd2a18b3 100644 --- a/src/Std/Data/DHashMap/Internal/Model.lean +++ b/src/Std/Data/DHashMap/Internal/Model.lean @@ -418,7 +418,7 @@ def diffₘ [BEq α] [Hashable α] (m₁ m₂ : Raw₀ α β) : Raw₀ α β := if m₁.1.size ≤ m₂.1.size then filterₘ m₁ (fun k _ => !containsₘ m₂ k) else - eraseMany m₁ (toListModel m₂.1.buckets) + eraseManyEntries m₁ (toListModel m₂.1.buckets) /-- Internal implementation detail of the hash map -/ def insertListIfNewₘ [BEq α] [Hashable α] (m : Raw₀ α β) (l : List ((a : α) × β a)) : Raw₀ α β := @@ -669,9 +669,9 @@ theorem insertMany_eq_insertListₘ [BEq α] [Hashable α] (m : Raw₀ α β) (l simp only [List.foldl_cons, insertListₘ] apply ih -theorem eraseMany_eq_eraseListₘ [BEq α] [Hashable α] (m : Raw₀ α β) (l : List ((a : α) × β a)) : - eraseMany m l = eraseListₘ m l := by - simp only [eraseMany, Id.run_pure, pure_bind, List.forIn_pure_yield_eq_foldl] +theorem eraseManyEntries_eq_eraseListₘ [BEq α] [Hashable α] (m : Raw₀ α β) (l : List ((a : α) × β a)) : + eraseManyEntries m l = eraseListₘ m l := by + simp only [eraseManyEntries, Id.run_pure, pure_bind, List.forIn_pure_yield_eq_foldl] suffices ∀ (t : { m' // ∀ (P : Raw₀ α β → Prop), (∀ {m'' : Raw₀ α β} {a : α}, P m'' → P (m''.erase a)) → P m → P m' }), (List.foldl (fun m' p => ⟨m'.val.erase p.1, fun P h₁ h₂ => h₁ (m'.2 _ h₁ h₂)⟩) t l).val = diff --git a/src/Std/Data/DHashMap/Internal/WF.lean b/src/Std/Data/DHashMap/Internal/WF.lean index 9a0d0481ae1d..c38edb9a19d0 100644 --- a/src/Std/Data/DHashMap/Internal/WF.lean +++ b/src/Std/Data/DHashMap/Internal/WF.lean @@ -1312,10 +1312,10 @@ theorem toListModel_insertMany_list [BEq α] [Hashable α] [EquivBEq α] [Lawful /-! # `eraseMany` -/ -theorem WF.eraseMany [BEq α] [Hashable α] {ρ : Type w} [ForIn Id ρ ((a : α) × β a)] {m : Raw α β} - {l : ρ} (h : m.WF) : (m.eraseMany l).WF := by - simpa [Raw.eraseMany, h.size_buckets_pos] using - (Raw₀.eraseMany ⟨m, h.size_buckets_pos⟩ l).2 _ Raw.WF.erase₀ h +theorem WF.eraseManyEntries [BEq α] [Hashable α] {ρ : Type w} [ForIn Id ρ ((a : α) × β a)] {m : Raw α β} + {l : ρ} (h : m.WF) : (m.eraseManyEntries l).WF := by + simpa [Raw.eraseManyEntries, h.size_buckets_pos] using + (Raw₀.eraseManyEntries ⟨m, h.size_buckets_pos⟩ l).2 _ Raw.WF.erase₀ h theorem wf_insertMany₀ [BEq α] [Hashable α] [EquivBEq α] [LawfulHashable α] {ρ : Type w} [ForIn Id ρ ((a : α) × β a)] {m : Raw α β} {h : 0 < m.buckets.size} {l : ρ} (h' : m.WF) : @@ -1324,12 +1324,12 @@ theorem wf_insertMany₀ [BEq α] [Hashable α] [EquivBEq α] [LawfulHashable α theorem wf_eraseMany₀ [BEq α] [Hashable α] [EquivBEq α] [LawfulHashable α] {ρ : Type w} [ForIn Id ρ ((a : α) × β a)] {m : Raw α β} {h : 0 < m.buckets.size} {l : ρ} (h' : m.WF) : - (Raw₀.eraseMany ⟨m, h⟩ l).1.1.WF := - (Raw₀.eraseMany ⟨m, h⟩ l).2 _ Raw.WF.erase₀ h' + (eraseManyEntries ⟨m, h⟩ l).1.1.WF := + (eraseManyEntries ⟨m, h⟩ l).2 _ Raw.WF.erase₀ h' theorem eraseMany_eq_eraseListₘ_toListModel [BEq α] [Hashable α] (m m₂ : Raw₀ α β) : - eraseMany m m₂.1 = eraseListₘ m (toListModel m₂.1.buckets) := by - simp only [eraseMany, bind_pure_comp, map_pure, bind_pure] + eraseManyEntries m m₂.1 = eraseListₘ m (toListModel m₂.1.buckets) := by + simp only [eraseManyEntries, bind_pure_comp, map_pure, bind_pure] simp only [ForIn.forIn] simp only [Raw.forIn_eq_forIn_toListModel, forIn_pure_yield_eq_foldl, Id.run_pure] generalize toListModel m₂.val.buckets = l @@ -1358,7 +1358,7 @@ theorem toListModel_diffₘ [BEq α] [Hashable α] [EquivBEq α] [LawfulHashable ext p congr rw [containsₘ_eq_containsKey h₂] - · rw [eraseMany_eq_eraseListₘ] + · rw [eraseManyEntries_eq_eraseListₘ] apply Perm.trans (toListModel_eraseListₘ h₁) · apply eraseList_perm_filter_containsKey · exact h₁.distinct @@ -1370,7 +1370,7 @@ theorem diff_eq_diffₘ [BEq α] [Hashable α] (m₁ m₂ : Raw₀ α β) : · rw [filter_eq_filterₘ] congr · rw [eraseMany_eq_eraseListₘ_toListModel] - rw [eraseMany_eq_eraseListₘ] + rw [eraseManyEntries_eq_eraseListₘ] theorem toListModel_diff [BEq α] [Hashable α] [EquivBEq α] [LawfulHashable α] {m₁ m₂ : Raw₀ α β} (h₁ : Raw.WFImp m₁.1) (h₂ : Raw.WFImp m₂.1) : diff --git a/src/Std/Data/DHashMap/Raw.lean b/src/Std/Data/DHashMap/Raw.lean index 6f57ccd436f7..77c250495722 100644 --- a/src/Std/Data/DHashMap/Raw.lean +++ b/src/Std/Data/DHashMap/Raw.lean @@ -555,10 +555,10 @@ Erases multiple keys from the hash map by iterating over the given collection an If the same key appears multiple times in the collection, subsequent erasures have no effect after the first one removes the key. -/ -@[inline] def eraseMany [BEq α] [Hashable α] {ρ : Type w} [ForIn Id ρ ((a : α) × β a)] +@[inline] def eraseManyEntries [BEq α] [Hashable α] {ρ : Type w} [ForIn Id ρ ((a : α) × β a)] (m : Raw α β) (l : ρ) : Raw α β := if h : 0 < m.buckets.size then - (Raw₀.eraseMany ⟨m, h⟩ l).1 + (Raw₀.eraseManyEntries ⟨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 α] From 6b5dba15a2fbbaa2dc7120512da0b1900adcedd9 Mon Sep 17 00:00:00 2001 From: Wojciech Rozowski Date: Tue, 18 Nov 2025 13:43:31 +0100 Subject: [PATCH 72/88] Apply Markus' suggestion --- src/Std/Data/DHashMap/Internal/Model.lean | 10 +++++----- src/Std/Data/DHashMap/Internal/WF.lean | 8 ++++---- src/Std/Data/Internal/List/Associative.lean | 10 +++++----- src/lake/tests/clone/hello | 1 + src/lake/tests/depTree/a | 1 + src/lake/tests/depTree/b | 1 + src/lake/tests/depTree/c | 1 + src/lake/tests/depTree/d | 1 + src/lake/tests/manifest/bar | 1 + src/lake/tests/packageOverrides/bar1 | 1 + 10 files changed, 21 insertions(+), 14 deletions(-) create mode 160000 src/lake/tests/clone/hello create mode 160000 src/lake/tests/depTree/a create mode 160000 src/lake/tests/depTree/b create mode 160000 src/lake/tests/depTree/c create mode 160000 src/lake/tests/depTree/d create mode 160000 src/lake/tests/manifest/bar create mode 160000 src/lake/tests/packageOverrides/bar1 diff --git a/src/Std/Data/DHashMap/Internal/Model.lean b/src/Std/Data/DHashMap/Internal/Model.lean index c91efd2a18b3..7ebc9dc29b2b 100644 --- a/src/Std/Data/DHashMap/Internal/Model.lean +++ b/src/Std/Data/DHashMap/Internal/Model.lean @@ -408,10 +408,10 @@ def insertListₘ [BEq α] [Hashable α] (m : Raw₀ α β) (l : List ((a : α) | .cons hd tl => insertListₘ (m.insert hd.1 hd.2) tl /-- Internal implementation detail of the hash map -/ -def eraseListₘ [BEq α] [Hashable α] (m : Raw₀ α β) (l : List ((a : α) × β a)) : Raw₀ α β := +def eraseListₘ [BEq α] [Hashable α] (m : Raw₀ α β) (l : List α) : Raw₀ α β := match l with | .nil => m - | .cons hd tl => eraseListₘ (m.erase hd.1) tl + | .cons hd tl => eraseListₘ (m.erase hd) tl /-- Internal implementation detail of the hash map -/ def diffₘ [BEq α] [Hashable α] (m₁ m₂ : Raw₀ α β) : Raw₀ α β := @@ -670,17 +670,17 @@ theorem insertMany_eq_insertListₘ [BEq α] [Hashable α] (m : Raw₀ α β) (l apply ih theorem eraseManyEntries_eq_eraseListₘ [BEq α] [Hashable α] (m : Raw₀ α β) (l : List ((a : α) × β a)) : - eraseManyEntries m l = eraseListₘ m l := by + eraseManyEntries m l = eraseListₘ m (l.map (·.1)) := by simp only [eraseManyEntries, Id.run_pure, pure_bind, List.forIn_pure_yield_eq_foldl] suffices ∀ (t : { m' // ∀ (P : Raw₀ α β → Prop), (∀ {m'' : Raw₀ α β} {a : α}, P m'' → P (m''.erase a)) → P m → P m' }), (List.foldl (fun m' p => ⟨m'.val.erase p.1, fun P h₁ h₂ => h₁ (m'.2 _ h₁ h₂)⟩) t l).val = - t.val.eraseListₘ l from this _ + t.val.eraseListₘ (l.map (·.1)) from this _ intro t induction l generalizing m with | nil => simp [eraseListₘ] | cons hd tl ih => - simp only [List.foldl_cons, eraseListₘ] + simp only [List.foldl_cons] apply ih theorem insertManyIfNew_eq_insertListIfNewₘ [BEq α] [Hashable α] (m : Raw₀ α β) (l : List ((a : α) × β a)) : diff --git a/src/Std/Data/DHashMap/Internal/WF.lean b/src/Std/Data/DHashMap/Internal/WF.lean index a028ef911a1a..465f433d8d53 100644 --- a/src/Std/Data/DHashMap/Internal/WF.lean +++ b/src/Std/Data/DHashMap/Internal/WF.lean @@ -1092,7 +1092,7 @@ theorem toListModel_insertListₘ [BEq α] [Hashable α] [EquivBEq α] [LawfulHa apply List.insertList_perm_of_perm_first (toListModel_insert h) (wfImp_insert h).distinct theorem toListModel_eraseListₘ [BEq α] [Hashable α] [EquivBEq α] [LawfulHashable α] - {m : Raw₀ α β} {l : List ((a : α) × β a)} (h : Raw.WFImp m.1) : + {m : Raw₀ α β} {l : List α} (h : Raw.WFImp m.1) : Perm (toListModel (eraseListₘ m l).1.buckets) (List.eraseList (toListModel m.1.buckets) l) := by induction l generalizing m with @@ -1337,7 +1337,7 @@ theorem wf_eraseMany₀ [BEq α] [Hashable α] [EquivBEq α] [LawfulHashable α] (eraseManyEntries ⟨m, h⟩ l).2 _ Raw.WF.erase₀ h' theorem eraseMany_eq_eraseListₘ_toListModel [BEq α] [Hashable α] (m m₂ : Raw₀ α β) : - eraseManyEntries m m₂.1 = eraseListₘ m (toListModel m₂.1.buckets) := by + eraseManyEntries m m₂.1 = eraseListₘ m ((toListModel m₂.1.buckets).map (·.1)) := by simp only [eraseManyEntries, bind_pure_comp, map_pure, bind_pure] simp only [ForIn.forIn] simp only [Raw.forIn_eq_forIn_toListModel, forIn_pure_yield_eq_foldl, Id.run_pure] @@ -1345,12 +1345,12 @@ theorem eraseMany_eq_eraseListₘ_toListModel [BEq α] [Hashable α] (m m₂ : R suffices ∀ (t : { m' // ∀ (P : Raw₀ α β → Prop), (∀ {m'' : Raw₀ α β} {a : α}, P m'' → P (m''.erase a)) → P m → P m' }), (List.foldl (fun m' p => ⟨m'.val.erase p.1, fun P h₁ h₂ => h₁ (m'.2 _ h₁ h₂)⟩) t l).val = - t.val.eraseListₘ l from this _ + t.val.eraseListₘ (l.map (·.1)) from this _ intro t induction l generalizing m with | nil => simp [eraseListₘ] | cons hd tl ih => - simp only [List.foldl_cons, eraseListₘ] + simp only [List.foldl_cons] apply ih theorem toListModel_diffₘ [BEq α] [Hashable α] [EquivBEq α] [LawfulHashable α] diff --git a/src/Std/Data/Internal/List/Associative.lean b/src/Std/Data/Internal/List/Associative.lean index 1b5e6fe4b996..388cdd6570a9 100644 --- a/src/Std/Data/Internal/List/Associative.lean +++ b/src/Std/Data/Internal/List/Associative.lean @@ -5349,13 +5349,13 @@ theorem getKey?_filter_containsKey_of_containsKey_right [BEq α] [EquivBEq α] · exact dl₁ /-- Internal implementation detail of the hash map -/ -def eraseList [BEq α] (l toErase : List ((a : α) × β a)) : List ((a : α) × β a) := +def eraseList [BEq α] (l : List ((a : α) × β a)) (toErase : List α): List ((a : α) × β a) := match toErase with | .nil => l - | .cons ⟨k, _⟩ toErase => eraseList (eraseKey k l) toErase + | .cons k toErase => eraseList (eraseKey k l) toErase theorem eraseList_perm_filter_containsKey [BEq α] [EquivBEq α] (l toErase : List ((a : α) × β a)) (hl : DistinctKeys l): - List.Perm (eraseList l toErase) (l.filter (fun p => !containsKey p.fst toErase)) := by + List.Perm (eraseList l (toErase.map (·.1))) (l.filter (fun p => !containsKey p.fst toErase)) := by induction toErase generalizing l with | nil => simp [eraseList, containsKey] @@ -5366,7 +5366,7 @@ theorem eraseList_perm_filter_containsKey [BEq α] [EquivBEq α] (l toErase : Li simp only [List.filter_cons_of_pos, List.perm_cons] apply t_ih | cons kv tl ih => - simp only [eraseList] + simp only [List.map_cons, eraseList] apply Perm.trans · apply ih · apply DistinctKeys.eraseKey hl @@ -5416,7 +5416,7 @@ theorem eraseList_perm_filter_containsKey [BEq α] [EquivBEq α] (l toErase : Li · exact hl · apply DistinctKeys.eraseKey hl -theorem eraseList_perm_of_perm_first [BEq α] [EquivBEq α] {l1 l2 toErase : List ((a : α) × β a)} +theorem eraseList_perm_of_perm_first [BEq α] [EquivBEq α] {l1 l2 : List ((a : α) × β a)} {toErase : List α} (h : Perm l1 l2) (distinct : DistinctKeys l1) : Perm (eraseList l1 toErase) (eraseList l2 toErase) := by induction toErase generalizing l1 l2 with diff --git a/src/lake/tests/clone/hello b/src/lake/tests/clone/hello new file mode 160000 index 000000000000..3a924f941e40 --- /dev/null +++ b/src/lake/tests/clone/hello @@ -0,0 +1 @@ +Subproject commit 3a924f941e40846cc134b2c09c19ae71f4dbf9cd diff --git a/src/lake/tests/depTree/a b/src/lake/tests/depTree/a new file mode 160000 index 000000000000..5776f391d0eb --- /dev/null +++ b/src/lake/tests/depTree/a @@ -0,0 +1 @@ +Subproject commit 5776f391d0eb2e841041ebadc3f847d16bc1843a diff --git a/src/lake/tests/depTree/b b/src/lake/tests/depTree/b new file mode 160000 index 000000000000..f335505571b9 --- /dev/null +++ b/src/lake/tests/depTree/b @@ -0,0 +1 @@ +Subproject commit f335505571b9c41572d3b56511f454ceb73772ae diff --git a/src/lake/tests/depTree/c b/src/lake/tests/depTree/c new file mode 160000 index 000000000000..9f7bbb33a1c1 --- /dev/null +++ b/src/lake/tests/depTree/c @@ -0,0 +1 @@ +Subproject commit 9f7bbb33a1c18f4dc4733a3ba0ac338e867fb88d diff --git a/src/lake/tests/depTree/d b/src/lake/tests/depTree/d new file mode 160000 index 000000000000..32bf98fb6b71 --- /dev/null +++ b/src/lake/tests/depTree/d @@ -0,0 +1 @@ +Subproject commit 32bf98fb6b712f4096ecadd1cde3981897d5fd30 diff --git a/src/lake/tests/manifest/bar b/src/lake/tests/manifest/bar new file mode 160000 index 000000000000..977bed86d8e2 --- /dev/null +++ b/src/lake/tests/manifest/bar @@ -0,0 +1 @@ +Subproject commit 977bed86d8e2b00c59af7e8c7d0552969d0fcba9 diff --git a/src/lake/tests/packageOverrides/bar1 b/src/lake/tests/packageOverrides/bar1 new file mode 160000 index 000000000000..cc741410bf5f --- /dev/null +++ b/src/lake/tests/packageOverrides/bar1 @@ -0,0 +1 @@ +Subproject commit cc741410bf5f604dfbf2cded3cf70580e64c802a From ac3049f5c2fffd3ef69086ff32832a7c1ef7dd9b Mon Sep 17 00:00:00 2001 From: Wojciech Rozowski Date: Tue, 18 Nov 2025 13:44:27 +0100 Subject: [PATCH 73/88] Untrack files added by accident --- src/lake/tests/clone/hello | 1 - src/lake/tests/depTree/a | 1 - src/lake/tests/depTree/b | 1 - src/lake/tests/depTree/c | 1 - src/lake/tests/depTree/d | 1 - src/lake/tests/manifest/bar | 1 - src/lake/tests/packageOverrides/bar1 | 1 - 7 files changed, 7 deletions(-) delete mode 160000 src/lake/tests/clone/hello delete mode 160000 src/lake/tests/depTree/a delete mode 160000 src/lake/tests/depTree/b delete mode 160000 src/lake/tests/depTree/c delete mode 160000 src/lake/tests/depTree/d delete mode 160000 src/lake/tests/manifest/bar delete mode 160000 src/lake/tests/packageOverrides/bar1 diff --git a/src/lake/tests/clone/hello b/src/lake/tests/clone/hello deleted file mode 160000 index 3a924f941e40..000000000000 --- a/src/lake/tests/clone/hello +++ /dev/null @@ -1 +0,0 @@ -Subproject commit 3a924f941e40846cc134b2c09c19ae71f4dbf9cd diff --git a/src/lake/tests/depTree/a b/src/lake/tests/depTree/a deleted file mode 160000 index 5776f391d0eb..000000000000 --- a/src/lake/tests/depTree/a +++ /dev/null @@ -1 +0,0 @@ -Subproject commit 5776f391d0eb2e841041ebadc3f847d16bc1843a diff --git a/src/lake/tests/depTree/b b/src/lake/tests/depTree/b deleted file mode 160000 index f335505571b9..000000000000 --- a/src/lake/tests/depTree/b +++ /dev/null @@ -1 +0,0 @@ -Subproject commit f335505571b9c41572d3b56511f454ceb73772ae diff --git a/src/lake/tests/depTree/c b/src/lake/tests/depTree/c deleted file mode 160000 index 9f7bbb33a1c1..000000000000 --- a/src/lake/tests/depTree/c +++ /dev/null @@ -1 +0,0 @@ -Subproject commit 9f7bbb33a1c18f4dc4733a3ba0ac338e867fb88d diff --git a/src/lake/tests/depTree/d b/src/lake/tests/depTree/d deleted file mode 160000 index 32bf98fb6b71..000000000000 --- a/src/lake/tests/depTree/d +++ /dev/null @@ -1 +0,0 @@ -Subproject commit 32bf98fb6b712f4096ecadd1cde3981897d5fd30 diff --git a/src/lake/tests/manifest/bar b/src/lake/tests/manifest/bar deleted file mode 160000 index 977bed86d8e2..000000000000 --- a/src/lake/tests/manifest/bar +++ /dev/null @@ -1 +0,0 @@ -Subproject commit 977bed86d8e2b00c59af7e8c7d0552969d0fcba9 diff --git a/src/lake/tests/packageOverrides/bar1 b/src/lake/tests/packageOverrides/bar1 deleted file mode 160000 index cc741410bf5f..000000000000 --- a/src/lake/tests/packageOverrides/bar1 +++ /dev/null @@ -1 +0,0 @@ -Subproject commit cc741410bf5f604dfbf2cded3cf70580e64c802a From c19e2b7d15aa7c9668c334fe18225c296d044a56 Mon Sep 17 00:00:00 2001 From: Wojciech Rozowski Date: Tue, 18 Nov 2025 15:13:03 +0100 Subject: [PATCH 74/88] Add lemmas --- src/Std/Data/DHashMap/Basic.lean | 11 + src/Std/Data/DHashMap/Internal/Raw.lean | 4 + src/Std/Data/DHashMap/Lemmas.lean | 302 ++++++++++++++++++ src/Std/Data/DHashMap/Raw.lean | 21 ++ src/Std/Data/DHashMap/RawLemmas.lean | 389 +++++++++++++++++++++++- src/Std/Data/HashMap/Raw.lean | 3 + 6 files changed, 729 insertions(+), 1 deletion(-) diff --git a/src/Std/Data/DHashMap/Basic.lean b/src/Std/Data/DHashMap/Basic.lean index 88abac99b4f6..f580b576f491 100644 --- a/src/Std/Data/DHashMap/Basic.lean +++ b/src/Std/Data/DHashMap/Basic.lean @@ -346,8 +346,19 @@ This function always iterates through the smaller map, so the expected runtime i inner := Raw₀.inter ⟨m₁.1, m₁.2.size_buckets_pos⟩ ⟨m₂.1, m₂.2.size_buckets_pos⟩ wf := Std.DHashMap.Raw.WF.inter₀ m₁.2 m₂.2 +/-- +Computes the difference of the given hash maps. + +This function always iterates through the smaller map, so the expected runtime is +`O(min(m₁.size, m₂.size))`. +-/ +@[inline] def diff [BEq α] [Hashable α] (m₁ m₂ : DHashMap α β) : DHashMap α β where + inner := Raw₀.diff ⟨m₁.1, m₁.2.size_buckets_pos⟩ ⟨m₂.1, m₂.2.size_buckets_pos⟩ + wf := Std.DHashMap.Raw.WF.diff₀ m₁.2 m₂.2 + instance [BEq α] [Hashable α] : Union (DHashMap α β) := ⟨union⟩ instance [BEq α] [Hashable α] : Inter (DHashMap α β) := ⟨inter⟩ +instance [BEq α] [Hashable α] : SDiff (DHashMap α β) := ⟨diff⟩ section Unverified diff --git a/src/Std/Data/DHashMap/Internal/Raw.lean b/src/Std/Data/DHashMap/Internal/Raw.lean index b421e0f6caf2..52010d8636b6 100644 --- a/src/Std/Data/DHashMap/Internal/Raw.lean +++ b/src/Std/Data/DHashMap/Internal/Raw.lean @@ -149,6 +149,10 @@ theorem inter_eq [BEq α] [Hashable α] {m₁ m₂ : Raw α β} (h₁ : m₁.WF) m₁.inter m₂ = Raw₀.inter ⟨m₁, h₁.size_buckets_pos⟩ ⟨m₂, h₂.size_buckets_pos⟩ := by simp [Raw.inter, h₁.size_buckets_pos, h₂.size_buckets_pos] +theorem diff_eq [BEq α] [Hashable α] {m₁ m₂ : Raw α β} (h₁ : m₁.WF) (h₂ : m₂.WF) : + m₁.diff m₂ = Raw₀.diff ⟨m₁, h₁.size_buckets_pos⟩ ⟨m₂, h₂.size_buckets_pos⟩ := by + simp [Raw.diff, h₁.size_buckets_pos, h₂.size_buckets_pos] + section variable {β : Type v} diff --git a/src/Std/Data/DHashMap/Lemmas.lean b/src/Std/Data/DHashMap/Lemmas.lean index 7e63f8a5397c..70bf124200ab 100644 --- a/src/Std/Data/DHashMap/Lemmas.lean +++ b/src/Std/Data/DHashMap/Lemmas.lean @@ -2464,6 +2464,308 @@ theorem get!_inter_of_not_mem_left [EquivBEq α] [LawfulHashable α] [Inhabited end Const +section Diff + +variable (m₁ m₂ : DHashMap α β) + +variable {m₁ m₂} + +@[simp] +theorem diff_eq : m₁.diff m₂ = m₁ \ m₂ := by + simp only [SDiff.sdiff] + +/- contains -/ +@[simp] +theorem contains_diff [EquivBEq α] [LawfulHashable α] {k : α} : + (m₁ \ m₂).contains k = (m₁.contains k && !m₂.contains k) := + @Raw₀.contains_diff _ _ _ _ ⟨m₁.1, m₁.2.size_buckets_pos⟩ ⟨m₂.1, m₂.2.size_buckets_pos⟩ _ _ m₁.2 m₂.2 k + +/- mem -/ +@[simp] +theorem mem_diff_iff [EquivBEq α] [LawfulHashable α] {k : α} : + k ∈ m₁ \ m₂ ↔ k ∈ m₁ ∧ k ∉ m₂ := + @Raw₀.contains_diff_iff _ _ _ _ ⟨m₁.1, _⟩ ⟨m₂.1, _⟩ _ _ m₁.2 m₂.2 k + +theorem not_mem_diff_of_not_mem_left [EquivBEq α] [LawfulHashable α] {k : α} + (not_mem : k ∉ m₁) : + k ∉ m₁ \ m₂ := by + rw [← contains_eq_false_iff_not_mem] at not_mem ⊢ + exact @Raw₀.contains_diff_eq_false_of_contains_eq_false_left _ _ _ _ ⟨m₁.1, m₁.2.size_buckets_pos⟩ ⟨m₂.1, m₂.2.size_buckets_pos⟩ _ _ m₁.2 m₂.2 k not_mem + +theorem not_mem_diff_of_mem_right [EquivBEq α] [LawfulHashable α] {k : α} + (mem : k ∈ m₂) : + k ∉ m₁ \ m₂ := by + rw [← contains_eq_false_iff_not_mem] + exact @Raw₀.contains_diff_eq_false_of_contains_right _ _ _ _ ⟨m₁.1, m₁.2.size_buckets_pos⟩ ⟨m₂.1, m₂.2.size_buckets_pos⟩ _ _ m₁.2 m₂.2 k mem + +/- get? -/ +theorem get?_diff [LawfulBEq α] {k : α} : + (m₁ \ m₂).get? k = if k ∈ m₂ then none else m₁.get? k := + @Raw₀.get?_diff _ _ _ _ ⟨m₁.1, m₁.2.size_buckets_pos⟩ ⟨m₂.1, m₂.2.size_buckets_pos⟩ _ m₁.2 m₂.2 k + +theorem get?_diff_of_not_mem_right [LawfulBEq α] + {k : α} (not_mem : k ∉ m₂) : + (m₁ \ m₂).get? k = m₁.get? k := by + rw [← contains_eq_false_iff_not_mem] at not_mem + exact @Raw₀.get?_diff_of_contains_eq_false_right _ _ _ _ ⟨m₁.1, m₁.2.size_buckets_pos⟩ ⟨m₂.1, m₂.2.size_buckets_pos⟩ _ m₁.2 m₂.2 k not_mem + +theorem get?_diff_of_not_mem_left [LawfulBEq α] + {k : α} (not_mem : k ∉ m₁) : + (m₁ \ m₂).get? k = none := by + rw [← contains_eq_false_iff_not_mem] at not_mem + exact @Raw₀.get?_diff_of_contains_eq_false_left _ _ _ _ ⟨m₁.1, m₁.2.size_buckets_pos⟩ ⟨m₂.1, m₂.2.size_buckets_pos⟩ _ m₁.2 m₂.2 k not_mem + +theorem get?_diff_of_mem_right [LawfulBEq α] + {k : α} (mem : k ∈ m₂) : + (m₁ \ m₂).get? k = none := + @Raw₀.get?_diff_of_contains_right _ _ _ _ ⟨m₁.1, m₁.2.size_buckets_pos⟩ ⟨m₂.1, m₂.2.size_buckets_pos⟩ _ m₁.2 m₂.2 k mem + +/- get -/ +@[simp] theorem get_diff [LawfulBEq α] + {k : α} {h_mem : k ∈ m₁ \ m₂} : + (m₁ \ m₂).get k h_mem = + m₁.get k ((mem_diff_iff.1 h_mem).1) := by + rw [mem_iff_contains] at h_mem + exact @Raw₀.get_diff _ _ _ _ ⟨m₁.1, _⟩ ⟨m₂.1, _⟩ _ m₁.2 m₂.2 k h_mem + +/- getD -/ +theorem getD_diff [LawfulBEq α] {k : α} {fallback : β k} : + (m₁ \ m₂).getD k fallback = + if k ∈ m₂ then fallback else m₁.getD k fallback := + @Raw₀.getD_diff _ _ _ _ ⟨m₁.1, m₁.2.size_buckets_pos⟩ ⟨m₂.1, m₂.2.size_buckets_pos⟩ _ m₁.2 m₂.2 k fallback + +theorem getD_diff_of_not_mem_right [LawfulBEq α] + {k : α} {fallback : β k} (not_mem : k ∉ m₂) : + (m₁ \ m₂).getD k fallback = m₁.getD k fallback := by + rw [← contains_eq_false_iff_not_mem] at not_mem + exact @Raw₀.getD_diff_of_contains_eq_false_right _ _ _ _ ⟨m₁.1, m₁.2.size_buckets_pos⟩ ⟨m₂.1, m₂.2.size_buckets_pos⟩ _ m₁.2 m₂.2 k fallback not_mem + +theorem getD_diff_of_mem_right [LawfulBEq α] + {k : α} {fallback : β k} (mem : k ∈ m₂) : + (m₁ \ m₂).getD k fallback = fallback := + @Raw₀.getD_diff_of_contains_right _ _ _ _ ⟨m₁.1, m₁.2.size_buckets_pos⟩ ⟨m₂.1, m₂.2.size_buckets_pos⟩ _ m₁.2 m₂.2 k fallback mem + +theorem getD_diff_of_not_mem_left [LawfulBEq α] + {k : α} {fallback : β k} (not_mem : k ∉ m₁) : + (m₁ \ m₂).getD k fallback = fallback := by + rw [← contains_eq_false_iff_not_mem] at not_mem + exact @Raw₀.getD_diff_of_contains_eq_false_left _ _ _ _ ⟨m₁.1, m₁.2.size_buckets_pos⟩ ⟨m₂.1, m₂.2.size_buckets_pos⟩ _ m₁.2 m₂.2 k fallback not_mem + +/- get! -/ +theorem get!_diff [LawfulBEq α] {k : α} [Inhabited (β k)] : + (m₁ \ m₂).get! k = + if k ∈ m₂ then default else m₁.get! k := + @Raw₀.get!_diff _ _ _ _ ⟨m₁.1, m₁.2.size_buckets_pos⟩ ⟨m₂.1, m₂.2.size_buckets_pos⟩ _ m₁.2 m₂.2 k _ + +theorem get!_diff_of_not_mem_right [LawfulBEq α] + {k : α} [Inhabited (β k)] (not_mem : k ∉ m₂) : + (m₁ \ m₂).get! k = m₁.get! k := by + rw [← contains_eq_false_iff_not_mem] at not_mem + exact @Raw₀.get!_diff_of_contains_eq_false_right _ _ _ _ ⟨m₁.1, m₁.2.size_buckets_pos⟩ ⟨m₂.1, m₂.2.size_buckets_pos⟩ _ m₁.2 m₂.2 k _ not_mem + +theorem get!_diff_of_mem_right [LawfulBEq α] + {k : α} [Inhabited (β k)] (mem : k ∈ m₂) : + (m₁ \ m₂).get! k = default := + @Raw₀.get!_diff_of_contains_right _ _ _ _ ⟨m₁.1, m₁.2.size_buckets_pos⟩ ⟨m₂.1, m₂.2.size_buckets_pos⟩ _ m₁.2 m₂.2 k _ mem + +theorem get!_diff_of_not_mem_left [LawfulBEq α] + {k : α} [Inhabited (β k)] (not_mem : k ∉ m₁) : + (m₁ \ m₂).get! k = default := by + rw [← contains_eq_false_iff_not_mem] at not_mem + exact @Raw₀.get!_diff_of_contains_eq_false_left _ _ _ _ ⟨m₁.1, m₁.2.size_buckets_pos⟩ ⟨m₂.1, m₂.2.size_buckets_pos⟩ _ m₁.2 m₂.2 k _ not_mem + +/- getKey? -/ +theorem getKey?_diff [EquivBEq α] [LawfulHashable α] {k : α} : + (m₁ \ m₂).getKey? k = + if k ∈ m₂ then none else m₁.getKey? k := + @Raw₀.getKey?_diff _ _ _ _ ⟨m₁.1, m₁.2.size_buckets_pos⟩ ⟨m₂.1, m₂.2.size_buckets_pos⟩ _ _ m₁.2 m₂.2 k + +theorem getKey?_diff_of_not_mem_right [EquivBEq α] [LawfulHashable α] + {k : α} (not_mem : k ∉ m₂) : + (m₁ \ m₂).getKey? k = m₁.getKey? k := by + rw [← contains_eq_false_iff_not_mem] at not_mem + exact @Raw₀.getKey?_diff_of_contains_eq_false_right _ _ _ _ ⟨m₁.1, m₁.2.size_buckets_pos⟩ ⟨m₂.1, m₂.2.size_buckets_pos⟩ _ _ m₁.2 m₂.2 k not_mem + +theorem getKey?_diff_of_not_mem_left [EquivBEq α] [LawfulHashable α] + {k : α} (not_mem : k ∉ m₁) : + (m₁ \ m₂).getKey? k = none := by + rw [← contains_eq_false_iff_not_mem] at not_mem + exact @Raw₀.getKey?_diff_of_contains_eq_false_left _ _ _ _ ⟨m₁.1, m₁.2.size_buckets_pos⟩ ⟨m₂.1, m₂.2.size_buckets_pos⟩ _ _ m₁.2 m₂.2 k not_mem + +theorem getKey?_diff_of_mem_right [EquivBEq α] [LawfulHashable α] + {k : α} (mem : k ∈ m₂) : + (m₁ \ m₂).getKey? k = none := + @Raw₀.getKey?_diff_of_contains_right _ _ _ _ ⟨m₁.1, m₁.2.size_buckets_pos⟩ ⟨m₂.1, m₂.2.size_buckets_pos⟩ _ _ m₁.2 m₂.2 k mem + +/- getKey -/ +@[simp] theorem getKey_diff [EquivBEq α] [LawfulHashable α] + {k : α} {h_mem : k ∈ m₁ \ m₂} : + (m₁ \ m₂).getKey k h_mem = + m₁.getKey k ((mem_diff_iff.1 h_mem).1) := by + rw [mem_iff_contains] at h_mem + exact @Raw₀.getKey_diff _ _ _ _ ⟨m₁.1, _⟩ ⟨m₂.1, _⟩ _ _ m₁.2 m₂.2 k h_mem + +/- getKeyD -/ +theorem getKeyD_diff [EquivBEq α] [LawfulHashable α] {k fallback : α} : + (m₁ \ m₂).getKeyD k fallback = + if k ∈ m₂ then fallback else m₁.getKeyD k fallback := + @Raw₀.getKeyD_diff _ _ _ _ ⟨m₁.1, m₁.2.size_buckets_pos⟩ ⟨m₂.1, m₂.2.size_buckets_pos⟩ _ _ m₁.2 m₂.2 k fallback + +theorem getKeyD_diff_of_not_mem_right [EquivBEq α] [LawfulHashable α] + {k fallback : α} (not_mem : k ∉ m₂) : + (m₁ \ m₂).getKeyD k fallback = m₁.getKeyD k fallback := by + rw [← contains_eq_false_iff_not_mem] at not_mem + exact @Raw₀.getKeyD_diff_of_contains_eq_false_right _ _ _ _ ⟨m₁.1, m₁.2.size_buckets_pos⟩ ⟨m₂.1, m₂.2.size_buckets_pos⟩ _ _ m₁.2 m₂.2 k fallback not_mem + +theorem getKeyD_diff_of_mem_right [EquivBEq α] [LawfulHashable α] + {k fallback : α} (mem : k ∈ m₂) : + (m₁ \ m₂).getKeyD k fallback = fallback := + @Raw₀.getKeyD_diff_of_contains_right _ _ _ _ ⟨m₁.1, m₁.2.size_buckets_pos⟩ ⟨m₂.1, m₂.2.size_buckets_pos⟩ _ _ m₁.2 m₂.2 k fallback mem + +theorem getKeyD_diff_of_not_mem_left [EquivBEq α] [LawfulHashable α] + {k fallback : α} (not_mem : k ∉ m₁) : + (m₁ \ m₂).getKeyD k fallback = fallback := by + rw [← contains_eq_false_iff_not_mem] at not_mem + exact @Raw₀.getKeyD_diff_of_contains_eq_false_left _ _ _ _ ⟨m₁.1, m₁.2.size_buckets_pos⟩ ⟨m₂.1, m₂.2.size_buckets_pos⟩ _ _ m₁.2 m₂.2 k fallback not_mem + +/- getKey! -/ +theorem getKey!_diff [EquivBEq α] [LawfulHashable α] [Inhabited α] {k : α} : + (m₁ \ m₂).getKey! k = + if k ∈ m₂ then default else m₁.getKey! k := + @Raw₀.getKey!_diff _ _ _ _ ⟨m₁.1, m₁.2.size_buckets_pos⟩ ⟨m₂.1, m₂.2.size_buckets_pos⟩ _ _ _ m₁.2 m₂.2 k + +theorem getKey!_diff_of_not_mem_right [EquivBEq α] [LawfulHashable α] [Inhabited α] + {k : α} (not_mem : k ∉ m₂) : + (m₁ \ m₂).getKey! k = m₁.getKey! k := by + rw [← contains_eq_false_iff_not_mem] at not_mem + exact @Raw₀.getKey!_diff_of_contains_eq_false_right _ _ _ _ ⟨m₁.1, m₁.2.size_buckets_pos⟩ ⟨m₂.1, m₂.2.size_buckets_pos⟩ _ _ _ m₁.2 m₂.2 k not_mem + +theorem getKey!_diff_of_mem_right [EquivBEq α] [LawfulHashable α] [Inhabited α] + {k : α} (mem : k ∈ m₂) : + (m₁ \ m₂).getKey! k = default := + @Raw₀.getKey!_diff_of_contains_right _ _ _ _ ⟨m₁.1, m₁.2.size_buckets_pos⟩ ⟨m₂.1, m₂.2.size_buckets_pos⟩ _ _ _ m₁.2 m₂.2 k mem + +theorem getKey!_diff_of_not_mem_left [EquivBEq α] [LawfulHashable α] [Inhabited α] + {k : α} (not_mem : k ∉ m₁) : + (m₁ \ m₂).getKey! k = default := by + rw [← contains_eq_false_iff_not_mem] at not_mem + exact @Raw₀.getKey!_diff_of_contains_eq_false_left _ _ _ _ ⟨m₁.1, m₁.2.size_buckets_pos⟩ ⟨m₂.1, m₂.2.size_buckets_pos⟩ _ _ _ m₁.2 m₂.2 k not_mem + +/- size -/ +theorem size_diff_le_size_left [EquivBEq α] [LawfulHashable α] : + (m₁ \ m₂).size ≤ m₁.size := + @Raw₀.size_diff_le_size_left _ _ _ _ ⟨m₁.1, m₁.2.size_buckets_pos⟩ ⟨m₂.1, m₂.2.size_buckets_pos⟩ _ _ m₁.2 m₂.2 + +theorem size_diff_eq_size_left [EquivBEq α] [LawfulHashable α] + (h : ∀ (a : α), a ∈ m₁ → a ∉ m₂) : + (m₁ \ m₂).size = m₁.size := by + have : ∀ (a : α), m₁.contains a → m₂.contains a = false := by + intro a ha + rw [contains_iff_mem] at ha + rw [contains_eq_false_iff_not_mem] + exact h a ha + exact @Raw₀.size_diff_eq_size_left _ _ _ _ ⟨m₁.1, m₁.2.size_buckets_pos⟩ ⟨m₂.1, m₂.2.size_buckets_pos⟩ _ _ m₁.2 m₂.2 this + +theorem size_diff_add_size_inter_eq_size_left [EquivBEq α] [LawfulHashable α] : + (m₁ \ m₂).size + (m₁ ∩ m₂).size = m₁.size := + @Raw₀.size_diff_add_size_inter_eq_size_left _ _ _ _ ⟨m₁.1, m₁.2.size_buckets_pos⟩ ⟨m₂.1, m₂.2.size_buckets_pos⟩ _ _ m₁.2 m₂.2 + +/- isEmpty -/ +@[simp] +theorem isEmpty_diff_left [EquivBEq α] [LawfulHashable α] (h : m₁.isEmpty) : + (m₁ \ m₂).isEmpty = true := + @Raw₀.isEmpty_diff_left _ _ _ _ ⟨m₁.1, m₁.2.size_buckets_pos⟩ ⟨m₂.1, m₂.2.size_buckets_pos⟩ _ _ m₁.2 m₂.2 h + +theorem isEmpty_diff_iff [EquivBEq α] [LawfulHashable α] : + (m₁ \ m₂).isEmpty ↔ ∀ k, k ∈ m₁ → k ∈ m₂ := by + simpa only [mem_iff_contains] using + @Raw₀.isEmpty_diff_iff _ _ _ _ ⟨m₁.1, m₁.2.size_buckets_pos⟩ ⟨m₂.1, m₂.2.size_buckets_pos⟩ _ _ m₁.wf m₂.wf + +end Diff + +namespace Const + +variable {β : Type v} {m₁ m₂ : DHashMap α (fun _ => β)} + +/- get? -/ +theorem get?_diff [EquivBEq α] [LawfulHashable α] {k : α} : + Const.get? (m₁.diff m₂) k = + if k ∈ m₂ then none else Const.get? m₁ k := + @Raw₀.Const.get?_diff _ _ _ _ ⟨m₁.1, m₁.2.size_buckets_pos⟩ ⟨m₂.1, m₂.2.size_buckets_pos⟩ _ _ m₁.2 m₂.2 k + +theorem get?_diff_of_not_mem_right [EquivBEq α] [LawfulHashable α] + {k : α} (not_mem : k ∉ m₂) : + Const.get? (m₁.diff m₂) k = Const.get? m₁ k := by + rw [← contains_eq_false_iff_not_mem] at not_mem + exact @Raw₀.Const.get?_diff_of_contains_eq_false_right _ _ _ _ ⟨m₁.1, m₁.2.size_buckets_pos⟩ ⟨m₂.1, m₂.2.size_buckets_pos⟩ _ _ m₁.2 m₂.2 k not_mem + +theorem get?_diff_of_not_mem_left [EquivBEq α] [LawfulHashable α] + {k : α} (not_mem : ¬k ∈ m₁) : + Const.get? (m₁.diff m₂) k = none := by + rw [← contains_eq_false_iff_not_mem] at not_mem + exact @Raw₀.Const.get?_diff_of_contains_eq_false_left _ _ _ _ ⟨m₁.1, m₁.2.size_buckets_pos⟩ ⟨m₂.1, m₂.2.size_buckets_pos⟩ _ _ m₁.2 m₂.2 k not_mem + +theorem get?_diff_of_mem_right [EquivBEq α] [LawfulHashable α] + {k : α} (mem : k ∈ m₂) : + Const.get? (m₁.diff m₂) k = none := + @Raw₀.Const.get?_diff_of_contains_right _ _ _ _ ⟨m₁.1, m₁.2.size_buckets_pos⟩ ⟨m₂.1, m₂.2.size_buckets_pos⟩ _ _ m₁.2 m₂.2 k mem + +/- get -/ +@[simp] theorem get_diff [EquivBEq α] [LawfulHashable α] + {k : α} {h_mem : k ∈ m₁ \ m₂} : + Const.get (m₁.diff m₂) k h_mem = + Const.get m₁ k ((mem_diff_iff.1 h_mem).1) := by + rw [mem_iff_contains] at h_mem + exact @Raw₀.Const.get_diff _ _ _ _ ⟨m₁.1, _⟩ ⟨m₂.1, _⟩ _ _ m₁.2 m₂.2 k h_mem + +/- getD -/ +theorem getD_diff [EquivBEq α] [LawfulHashable α] {k : α} {fallback : β} : + Const.getD (m₁.diff m₂) k fallback = + if k ∈ m₂ then fallback else Const.getD m₁ k fallback := + @Raw₀.Const.getD_diff _ _ _ _ ⟨m₁.1, m₁.2.size_buckets_pos⟩ ⟨m₂.1, m₂.2.size_buckets_pos⟩ _ _ m₁.2 m₂.2 k fallback + +theorem getD_diff_of_not_mem_right [EquivBEq α] [LawfulHashable α] + {k : α} {fallback : β} (not_mem : k ∉ m₂) : + Const.getD (m₁.diff m₂) k fallback = Const.getD m₁ k fallback := by + rw [← contains_eq_false_iff_not_mem] at not_mem + exact @Raw₀.Const.getD_diff_of_contains_eq_false_right _ _ _ _ ⟨m₁.1, m₁.2.size_buckets_pos⟩ ⟨m₂.1, m₂.2.size_buckets_pos⟩ _ _ m₁.2 m₂.2 k fallback not_mem + +theorem getD_diff_of_mem_right [EquivBEq α] [LawfulHashable α] + {k : α} {fallback : β} (mem : k ∈ m₂) : + Const.getD (m₁.diff m₂) k fallback = fallback := + @Raw₀.Const.getD_diff_of_contains_right _ _ _ _ ⟨m₁.1, m₁.2.size_buckets_pos⟩ ⟨m₂.1, m₂.2.size_buckets_pos⟩ _ _ m₁.2 m₂.2 k fallback mem + +theorem getD_diff_of_not_mem_left [EquivBEq α] [LawfulHashable α] + {k : α} {fallback : β} (not_mem : ¬k ∈ m₁) : + Const.getD (m₁.diff m₂) k fallback = fallback := by + rw [← contains_eq_false_iff_not_mem] at not_mem + exact @Raw₀.Const.getD_diff_of_contains_eq_false_left _ _ _ _ ⟨m₁.1, m₁.2.size_buckets_pos⟩ ⟨m₂.1, m₂.2.size_buckets_pos⟩ _ _ m₁.2 m₂.2 k fallback not_mem + +/- get! -/ +theorem get!_diff [EquivBEq α] [LawfulHashable α] [Inhabited β] {k : α} : + Const.get! (m₁.diff m₂) k = + if k ∈ m₂ then default else Const.get! m₁ k := + @Raw₀.Const.get!_diff _ _ _ _ ⟨m₁.1, m₁.2.size_buckets_pos⟩ ⟨m₂.1, m₂.2.size_buckets_pos⟩ _ _ _ m₁.2 m₂.2 k + +theorem get!_diff_of_not_mem_right [EquivBEq α] [LawfulHashable α] [Inhabited β] + {k : α} (not_mem : k ∉ m₂) : + Const.get! (m₁.diff m₂) k = Const.get! m₁ k := by + rw [← contains_eq_false_iff_not_mem] at not_mem + exact @Raw₀.Const.get!_diff_of_contains_eq_false_right _ _ _ _ ⟨m₁.1, m₁.2.size_buckets_pos⟩ ⟨m₂.1, m₂.2.size_buckets_pos⟩ _ _ _ m₁.2 m₂.2 k not_mem + +theorem get!_diff_of_mem_right [EquivBEq α] [LawfulHashable α] [Inhabited β] + {k : α} (mem : k ∈ m₂) : + Const.get! (m₁.diff m₂) k = default := + @Raw₀.Const.get!_diff_of_contains_right _ _ _ _ ⟨m₁.1, m₁.2.size_buckets_pos⟩ ⟨m₂.1, m₂.2.size_buckets_pos⟩ _ _ _ m₁.2 m₂.2 k mem + +theorem get!_diff_of_not_mem_left [EquivBEq α] [LawfulHashable α] [Inhabited β] + {k : α} (not_mem : ¬k ∈ m₁) : + Const.get! (m₁.diff m₂) k = default := by + rw [← contains_eq_false_iff_not_mem] at not_mem + exact @Raw₀.Const.get!_diff_of_contains_eq_false_left _ _ _ _ ⟨m₁.1, m₁.2.size_buckets_pos⟩ ⟨m₂.1, m₂.2.size_buckets_pos⟩ _ _ _ m₁.2 m₂.2 k not_mem + +end Const + namespace Const variable {β : Type v} {m : DHashMap α (fun _ => β)} diff --git a/src/Std/Data/DHashMap/Raw.lean b/src/Std/Data/DHashMap/Raw.lean index 77c250495722..8abe1c01120b 100644 --- a/src/Std/Data/DHashMap/Raw.lean +++ b/src/Std/Data/DHashMap/Raw.lean @@ -522,7 +522,22 @@ This function always merges the smaller map into the larger map, so the expected instance [BEq α] [Hashable α] : Inter (Raw α β) := ⟨inter⟩ +/-- +Computes the difference of the given hash maps. + +This function always iterates through the smaller map, so the expected runtime is +`O(min(m₁.size, m₂.size))`. +-/ +@[inline] def diff [BEq α] [Hashable α] (m₁ m₂ : Raw α β) : Raw α β := + if h₁ : 0 < m₁.buckets.size then + if h₂ : 0 < m₂.buckets.size then + Raw₀.diff ⟨m₁, h₁⟩ ⟨m₂, h₂⟩ + else + m₁ + else + m₂ +instance [BEq α] [Hashable α] : SDiff (Raw α β) := ⟨diff⟩ section Unverified /-! We currently do not provide lemmas for the functions below. -/ @@ -786,6 +801,12 @@ theorem WF.inter [BEq α] [Hashable α] {m₁ m₂ : Raw α β} (h₁ : m₁.WF) simp [Std.DHashMap.Raw.inter, h₁.size_buckets_pos, h₂.size_buckets_pos] exact WF.inter₀ h₁ h₂ +theorem WF.diff₀ [BEq α] [Hashable α] {m₁ m₂ : Raw α β} (h₁ : m₁.WF) (h₂ : m₂.WF) : (Raw₀.diff ⟨m₁, h₁.size_buckets_pos⟩ ⟨m₂, h₂.size_buckets_pos⟩).val.WF := by + simp only [Raw₀.diff] + split + . exact @WF.filter₀ α β _ _ m₁ h₁.size_buckets_pos (fun k x => !Raw₀.contains ⟨m₂, h₂.size_buckets_pos⟩ k) h₁ + . exact (Raw₀.eraseManyEntries ⟨m₁, h₁.size_buckets_pos⟩ m₂).2 _ WF.erase₀ h₁ + end WF end Raw diff --git a/src/Std/Data/DHashMap/RawLemmas.lean b/src/Std/Data/DHashMap/RawLemmas.lean index e24813fcd59c..38cb038e5727 100644 --- a/src/Std/Data/DHashMap/RawLemmas.lean +++ b/src/Std/Data/DHashMap/RawLemmas.lean @@ -48,7 +48,7 @@ private meta def baseNames : Array Name := ``getKey?_eq, ``getKey_eq, ``getKey!_eq, ``getKeyD_eq, ``insertMany_eq, ``Const.insertMany_eq, ``Const.insertManyIfNewUnit_eq, ``ofList_eq, ``Const.ofList_eq, ``Const.unitOfList_eq, - ``alter_eq, ``Const.alter_eq, ``modify_eq, ``Const.modify_eq, ``union_eq, ``inter_eq] + ``alter_eq, ``Const.alter_eq, ``modify_eq, ``Const.modify_eq, ``union_eq, ``inter_eq, ``diff_eq] /-- Internal implementation detail of the hash map -/ scoped syntax "simp_to_raw" ("using" term)? : tactic @@ -2871,6 +2871,393 @@ theorem get!_inter_of_not_mem_left [EquivBEq α] [LawfulHashable α] [Inhabited end Const +section Diff + +variable {m₁ m₂ : Raw α β} + +/- contains -/ +@[simp] +theorem contains_diff [EquivBEq α] [LawfulHashable α] (h₁ : m₁.WF) (h₂ : m₂.WF) {k : α} : + (m₁ \ m₂).contains k = (m₁.contains k && !m₂.contains k) := by + simp only [SDiff.sdiff] + simp_to_raw using Raw₀.contains_diff + +/- mem -/ +@[simp] +theorem mem_diff_iff [EquivBEq α] [LawfulHashable α] (h₁ : m₁.WF) (h₂ : m₂.WF) {k : α} : + k ∈ m₁ \ m₂ ↔ k ∈ m₁ ∧ k ∉ m₂ := by + simp only [SDiff.sdiff, Membership.mem] + simp_to_raw using @Raw₀.contains_diff_iff _ _ _ _ ⟨m₁, _⟩ ⟨m₂, _⟩ _ _ + +theorem not_mem_diff_of_not_mem_left [EquivBEq α] [LawfulHashable α] (h₁ : m₁.WF) (h₂ : m₂.WF) {k : α} + (not_mem : k ∉ m₁) : + k ∉ m₁ \ m₂ := by + rw [← contains_eq_false_iff_not_mem] at not_mem ⊢ + revert not_mem + simp only [SDiff.sdiff] + simp_to_raw using Raw₀.contains_diff_eq_false_of_contains_eq_false_left + +theorem not_mem_diff_of_mem_right [EquivBEq α] [LawfulHashable α] (h₁ : m₁.WF) (h₂ : m₂.WF) {k : α} + (mem : k ∈ m₂) : + k ∉ m₁ \ m₂ := by + rw [← contains_eq_false_iff_not_mem] + rw [mem_iff_contains] at mem + revert mem + simp only [SDiff.sdiff] + simp_to_raw using Raw₀.contains_diff_eq_false_of_contains_right + +/- get? -/ +theorem get?_diff [LawfulBEq α] (h₁ : m₁.WF) (h₂ : m₂.WF) {k : α} : + (m₁ \ m₂).get? k = if k ∈ m₂ then none else m₁.get? k := by + simp only [Membership.mem, SDiff.sdiff] + simp_to_raw using Raw₀.get?_diff + +theorem get?_diff_of_not_mem_right [LawfulBEq α] (h₁ : m₁.WF) (h₂ : m₂.WF) + {k : α} (not_mem : k ∉ m₂) : + (m₁ \ m₂).get? k = m₁.get? k := by + rw [← contains_eq_false_iff_not_mem] at not_mem + revert not_mem + simp only [SDiff.sdiff] + simp_to_raw using Raw₀.get?_diff_of_contains_eq_false_right + +theorem get?_diff_of_not_mem_left [LawfulBEq α] (h₁ : m₁.WF) (h₂ : m₂.WF) + {k : α} (not_mem : k ∉ m₁) : + (m₁ \ m₂).get? k = none := by + rw [← contains_eq_false_iff_not_mem] at not_mem + revert not_mem + simp only [SDiff.sdiff] + simp_to_raw using Raw₀.get?_diff_of_contains_eq_false_left + +theorem get?_diff_of_mem_right [LawfulBEq α] (h₁ : m₁.WF) (h₂ : m₂.WF) + {k : α} (mem : k ∈ m₂) : + (m₁ \ m₂).get? k = none := by + rw [mem_iff_contains] at mem + revert mem + simp only [SDiff.sdiff] + simp_to_raw using Raw₀.get?_diff_of_contains_right + +/- get -/ +@[simp] theorem get_diff [LawfulBEq α] (h₁ : m₁.WF) (h₂ : m₂.WF) + {k : α} {h_mem : k ∈ m₁ \ m₂} : + (m₁ \ m₂).get k h_mem = + m₁.get k ((mem_diff_iff h₁ h₂).1 h_mem).1 := by + rw [mem_iff_contains] at h_mem + revert h_mem + simp only [SDiff.sdiff] + simp_to_raw using Raw₀.get_diff + +/- getD -/ +theorem getD_diff [LawfulBEq α] (h₁ : m₁.WF) (h₂ : m₂.WF) + {k : α} {fallback : β k} : + (m₁ \ m₂).getD k fallback = + if k ∈ m₂ then fallback else m₁.getD k fallback := by + simp only [Membership.mem, SDiff.sdiff] + simp_to_raw using Raw₀.getD_diff + +theorem getD_diff_of_not_mem_right [LawfulBEq α] (h₁ : m₁.WF) (h₂ : m₂.WF) + {k : α} {fallback : β k} (not_mem : k ∉ m₂) : + (m₁ \ m₂).getD k fallback = m₁.getD k fallback := by + rw [← contains_eq_false_iff_not_mem] at not_mem + revert not_mem + simp only [SDiff.sdiff] + simp_to_raw using Raw₀.getD_diff_of_contains_eq_false_right + +theorem getD_diff_of_mem_right [LawfulBEq α] (h₁ : m₁.WF) (h₂ : m₂.WF) + {k : α} {fallback : β k} (mem : k ∈ m₂) : + (m₁ \ m₂).getD k fallback = fallback := by + rw [mem_iff_contains] at mem + revert mem + simp only [SDiff.sdiff] + simp_to_raw using Raw₀.getD_diff_of_contains_right + +theorem getD_diff_of_not_mem_left [LawfulBEq α] (h₁ : m₁.WF) (h₂ : m₂.WF) + {k : α} {fallback : β k} (not_mem : k ∉ m₁) : + (m₁ \ m₂).getD k fallback = fallback := by + rw [← contains_eq_false_iff_not_mem] at not_mem + revert not_mem + simp only [SDiff.sdiff] + simp_to_raw using Raw₀.getD_diff_of_contains_eq_false_left + +/- get! -/ +theorem get!_diff [LawfulBEq α] (h₁ : m₁.WF) (h₂ : m₂.WF) + {k : α} [Inhabited (β k)] : + (m₁ \ m₂).get! k = if k ∈ m₂ then default else m₁.get! k := by + simp only [Membership.mem, SDiff.sdiff] + simp_to_raw using Raw₀.get!_diff + +theorem get!_diff_of_not_mem_right [LawfulBEq α] (h₁ : m₁.WF) (h₂ : m₂.WF) + {k : α} [Inhabited (β k)] (not_mem : k ∉ m₂) : + (m₁ \ m₂).get! k = m₁.get! k := by + rw [← contains_eq_false_iff_not_mem] at not_mem + revert not_mem + simp only [SDiff.sdiff] + simp_to_raw using Raw₀.get!_diff_of_contains_eq_false_right + +theorem get!_diff_of_mem_right [LawfulBEq α] (h₁ : m₁.WF) (h₂ : m₂.WF) + {k : α} [Inhabited (β k)] (mem : k ∈ m₂) : + (m₁ \ m₂).get! k = default := by + rw [mem_iff_contains] at mem + revert mem + simp only [SDiff.sdiff] + simp_to_raw using Raw₀.get!_diff_of_contains_right + +theorem get!_diff_of_not_mem_left [LawfulBEq α] (h₁ : m₁.WF) (h₂ : m₂.WF) + {k : α} [Inhabited (β k)] (not_mem : k ∉ m₁) : + (m₁ \ m₂).get! k = default := by + rw [← contains_eq_false_iff_not_mem] at not_mem + revert not_mem + simp only [SDiff.sdiff] + simp_to_raw using Raw₀.get!_diff_of_contains_eq_false_left + +/- getKey? -/ +theorem getKey?_diff [EquivBEq α] [LawfulHashable α] (h₁ : m₁.WF) (h₂ : m₂.WF) {k : α} : + (m₁ \ m₂).getKey? k = + if k ∈ m₂ then none else m₁.getKey? k := by + simp only [Membership.mem, SDiff.sdiff] + simp_to_raw using Raw₀.getKey?_diff + +theorem getKey?_diff_of_not_mem_right [EquivBEq α] [LawfulHashable α] (h₁ : m₁.WF) (h₂ : m₂.WF) + {k : α} (not_mem : k ∉ m₂) : + (m₁ \ m₂).getKey? k = m₁.getKey? k := by + rw [← contains_eq_false_iff_not_mem] at not_mem + revert not_mem + simp only [SDiff.sdiff] + simp_to_raw using Raw₀.getKey?_diff_of_contains_eq_false_right + +theorem getKey?_diff_of_not_mem_left [EquivBEq α] [LawfulHashable α] (h₁ : m₁.WF) (h₂ : m₂.WF) + {k : α} (not_mem : k ∉ m₁) : + (m₁ \ m₂).getKey? k = none := by + rw [← contains_eq_false_iff_not_mem] at not_mem + revert not_mem + simp only [SDiff.sdiff] + simp_to_raw using Raw₀.getKey?_diff_of_contains_eq_false_left + +theorem getKey?_diff_of_mem_right [EquivBEq α] [LawfulHashable α] (h₁ : m₁.WF) (h₂ : m₂.WF) + {k : α} (mem : k ∈ m₂) : + (m₁ \ m₂).getKey? k = none := by + rw [mem_iff_contains] at mem + revert mem + simp only [SDiff.sdiff] + simp_to_raw using Raw₀.getKey?_diff_of_contains_right + +/- getKey -/ +@[simp] theorem getKey_diff [EquivBEq α] [LawfulHashable α] (h₁ : m₁.WF) (h₂ : m₂.WF) + {k : α} {h_mem : k ∈ m₁ \ m₂} : + (m₁ \ m₂).getKey k h_mem = + m₁.getKey k ((mem_diff_iff h₁ h₂).1 h_mem).1 := by + rw [mem_iff_contains] at h_mem + revert h_mem + simp only [SDiff.sdiff] + simp_to_raw using Raw₀.getKey_diff + +/- getKeyD -/ +theorem getKeyD_diff [EquivBEq α] [LawfulHashable α] (h₁ : m₁.WF) (h₂ : m₂.WF) + {k fallback : α} : + (m₁ \ m₂).getKeyD k fallback = + if k ∈ m₂ then fallback else m₁.getKeyD k fallback := by + simp only [Membership.mem, SDiff.sdiff] + simp_to_raw using Raw₀.getKeyD_diff + +theorem getKeyD_diff_of_not_mem_right [EquivBEq α] [LawfulHashable α] (h₁ : m₁.WF) (h₂ : m₂.WF) + {k fallback : α} (not_mem : k ∉ m₂) : + (m₁ \ m₂).getKeyD k fallback = m₁.getKeyD k fallback := by + rw [← contains_eq_false_iff_not_mem] at not_mem + revert not_mem + simp only [SDiff.sdiff] + simp_to_raw using Raw₀.getKeyD_diff_of_contains_eq_false_right + +theorem getKeyD_diff_of_mem_right [EquivBEq α] [LawfulHashable α] (h₁ : m₁.WF) (h₂ : m₂.WF) + {k fallback : α} (mem : k ∈ m₂) : + (m₁ \ m₂).getKeyD k fallback = fallback := by + rw [mem_iff_contains] at mem + revert mem + simp only [SDiff.sdiff] + simp_to_raw using Raw₀.getKeyD_diff_of_contains_right + +theorem getKeyD_diff_of_not_mem_left [EquivBEq α] [LawfulHashable α] (h₁ : m₁.WF) (h₂ : m₂.WF) + {k fallback : α} (not_mem : k ∉ m₁) : + (m₁ \ m₂).getKeyD k fallback = fallback := by + rw [← contains_eq_false_iff_not_mem] at not_mem + revert not_mem + simp only [SDiff.sdiff] + simp_to_raw using Raw₀.getKeyD_diff_of_contains_eq_false_left + +/- getKey! -/ +theorem getKey!_diff [EquivBEq α] [LawfulHashable α] [Inhabited α] (h₁ : m₁.WF) (h₂ : m₂.WF) + {k : α} : + (m₁ \ m₂).getKey! k = + if k ∈ m₂ then default else m₁.getKey! k := by + simp only [Membership.mem, SDiff.sdiff] + simp_to_raw using Raw₀.getKey!_diff + +theorem getKey!_diff_of_not_mem_right [EquivBEq α] [LawfulHashable α] [Inhabited α] (h₁ : m₁.WF) (h₂ : m₂.WF) + {k : α} (not_mem : k ∉ m₂) : + (m₁ \ m₂).getKey! k = m₁.getKey! k := by + rw [← contains_eq_false_iff_not_mem] at not_mem + revert not_mem + simp only [SDiff.sdiff] + simp_to_raw using Raw₀.getKey!_diff_of_contains_eq_false_right + +theorem getKey!_diff_of_mem_right [EquivBEq α] [LawfulHashable α] [Inhabited α] (h₁ : m₁.WF) (h₂ : m₂.WF) + {k : α} (mem : k ∈ m₂) : + (m₁ \ m₂).getKey! k = default := by + rw [mem_iff_contains] at mem + revert mem + simp only [SDiff.sdiff] + simp_to_raw using Raw₀.getKey!_diff_of_contains_right + +theorem getKey!_diff_of_not_mem_left [EquivBEq α] [LawfulHashable α] [Inhabited α] (h₁ : m₁.WF) (h₂ : m₂.WF) + {k : α} (not_mem : k ∉ m₁) : + (m₁ \ m₂).getKey! k = default := by + rw [← contains_eq_false_iff_not_mem] at not_mem + revert not_mem + simp only [SDiff.sdiff] + simp_to_raw using Raw₀.getKey!_diff_of_contains_eq_false_left + +/- size -/ +theorem size_diff_le_size_left [EquivBEq α] [LawfulHashable α] (h₁ : m₁.WF) (h₂ : m₂.WF) : + (m₁ \ m₂).size ≤ m₁.size := by + simp only [SDiff.sdiff] + simp_to_raw using Raw₀.size_diff_le_size_left + +theorem size_diff_eq_size_left [EquivBEq α] [LawfulHashable α] (h₁ : m₁.WF) (h₂ : m₂.WF) + (h : ∀ (a : α), a ∈ m₁ → a ∉ m₂) : + (m₁ \ m₂).size = m₁.size := by + have : ∀ (a : α), m₁.contains a → m₂.contains a = false := by + intro a ha + rw [contains_iff_mem] at ha + rw [contains_eq_false_iff_not_mem] + exact h a ha + revert this + simp only [SDiff.sdiff] + simp_to_raw using Raw₀.size_diff_eq_size_left + +theorem size_diff_add_size_inter_eq_size_left [EquivBEq α] [LawfulHashable α] (h₁ : m₁.WF) (h₂ : m₂.WF) : + (m₁ \ m₂).size + (m₁ ∩ m₂).size = m₁.size := by + simp only [SDiff.sdiff, Inter.inter] + simp_to_raw using Raw₀.size_diff_add_size_inter_eq_size_left + +/- isEmpty -/ +@[simp] +theorem isEmpty_diff_left [EquivBEq α] [LawfulHashable α] (h₁ : m₁.WF) (h₂ : m₂.WF) (h : m₁.isEmpty) : + (m₁ \ m₂).isEmpty = true := by + revert h + simp only [SDiff.sdiff] + intro h + simp_to_raw using Raw₀.isEmpty_diff_left + +theorem isEmpty_diff_iff [EquivBEq α] [LawfulHashable α] (h₁ : m₁.WF) (h₂ : m₂.WF) : + (m₁ \ m₂).isEmpty ↔ ∀ k, k ∈ m₁ → k ∈ m₂ := by + conv => + rhs + rhs + rhs + rw [mem_iff_contains] + simp only [SDiff.sdiff, Membership.mem] + simp_to_raw using @Raw₀.isEmpty_diff_iff _ _ _ _ ⟨m₁, _⟩ ⟨m₂, _⟩ _ _ + +end Diff + +namespace Const + +variable {β : Type v} {m₁ m₂ : Raw α (fun _ => β)} + +/- get? -/ +theorem get?_diff [EquivBEq α] [LawfulHashable α] (h₁ : m₁.WF) (h₂ : m₂.WF) {k : α} : + Const.get? (m₁.diff m₂) k = if k ∈ m₂ then none else Const.get? m₁ k := by + simp only [Membership.mem] + simp_to_raw using Raw₀.Const.get?_diff + +theorem get?_diff_of_not_mem_right [EquivBEq α] [LawfulHashable α] (h₁ : m₁.WF) (h₂ : m₂.WF) + {k : α} (not_mem : k ∉ m₂) : + Const.get? (m₁.diff m₂) k = Const.get? m₁ k := by + rw [← contains_eq_false_iff_not_mem] at not_mem + revert not_mem + simp_to_raw using Raw₀.Const.get?_diff_of_contains_eq_false_right + +theorem get?_diff_of_not_mem_left [EquivBEq α] [LawfulHashable α] (h₁ : m₁.WF) (h₂ : m₂.WF) + {k : α} (not_mem : ¬k ∈ m₁) : + Const.get? (m₁.diff m₂) k = none := by + rw [← contains_eq_false_iff_not_mem] at not_mem + revert not_mem + simp_to_raw using Raw₀.Const.get?_diff_of_contains_eq_false_left + +theorem get?_diff_of_mem_right [EquivBEq α] [LawfulHashable α] (h₁ : m₁.WF) (h₂ : m₂.WF) + {k : α} (mem : k ∈ m₂) : + Const.get? (m₁.diff m₂) k = none := by + rw [mem_iff_contains] at mem + revert mem + simp_to_raw using Raw₀.Const.get?_diff_of_contains_right + +/- get -/ +@[simp] theorem get_diff [EquivBEq α] [LawfulHashable α] (h₁ : m₁.WF) (h₂ : m₂.WF) + {k : α} {h_mem : k ∈ m₁ \ m₂} : + Const.get (m₁.diff m₂) k h_mem = + Const.get m₁ k ((mem_diff_iff h₁ h₂).1 h_mem).1 := by + rw [mem_iff_contains] at h_mem + revert h_mem + simp only [SDiff.sdiff] + simp_to_raw using Raw₀.Const.get_diff + +/- getD -/ +theorem getD_diff [EquivBEq α] [LawfulHashable α] (h₁ : m₁.WF) (h₂ : m₂.WF) + {k : α} {fallback : β} : + Const.getD (m₁.diff m₂) k fallback = + if k ∈ m₂ then fallback else Const.getD m₁ k fallback := by + simp only [Membership.mem] + simp_to_raw using Raw₀.Const.getD_diff + +theorem getD_diff_of_not_mem_right [EquivBEq α] [LawfulHashable α] (h₁ : m₁.WF) (h₂ : m₂.WF) + {k : α} {fallback : β} (not_mem : k ∉ m₂) : + Const.getD (m₁.diff m₂) k fallback = Const.getD m₁ k fallback := by + rw [← contains_eq_false_iff_not_mem] at not_mem + revert not_mem + + simp_to_raw using Raw₀.Const.getD_diff_of_contains_eq_false_right + +theorem getD_diff_of_mem_right [EquivBEq α] [LawfulHashable α] (h₁ : m₁.WF) (h₂ : m₂.WF) + {k : α} {fallback : β} (mem : k ∈ m₂) : + Const.getD (m₁.diff m₂) k fallback = fallback := by + rw [mem_iff_contains] at mem + revert mem + simp_to_raw using Raw₀.Const.getD_diff_of_contains_right + +theorem getD_diff_of_not_mem_left [EquivBEq α] [LawfulHashable α] (h₁ : m₁.WF) (h₂ : m₂.WF) + {k : α} {fallback : β} (not_mem : ¬k ∈ m₁) : + Const.getD (m₁.diff m₂) k fallback = fallback := by + rw [← contains_eq_false_iff_not_mem] at not_mem + revert not_mem + simp_to_raw using Raw₀.Const.getD_diff_of_contains_eq_false_left + +/- get! -/ +theorem get!_diff [EquivBEq α] [LawfulHashable α] [Inhabited β] (h₁ : m₁.WF) (h₂ : m₂.WF) {k : α} : + Const.get! (m₁.diff m₂) k = if k ∈ m₂ then default else Const.get! m₁ k := by + simp only [Membership.mem] + simp_to_raw using Raw₀.Const.get!_diff + +theorem get!_diff_of_not_mem_right [EquivBEq α] [LawfulHashable α] [Inhabited β] (h₁ : m₁.WF) (h₂ : m₂.WF) + {k : α} (not_mem : k ∉ m₂) : + Const.get! (m₁.diff m₂) k = Const.get! m₁ k := by + rw [← contains_eq_false_iff_not_mem] at not_mem + revert not_mem + simp_to_raw using Raw₀.Const.get!_diff_of_contains_eq_false_right + +theorem get!_diff_of_mem_right [EquivBEq α] [LawfulHashable α] [Inhabited β] (h₁ : m₁.WF) (h₂ : m₂.WF) + {k : α} (mem : k ∈ m₂) : + Const.get! (m₁.diff m₂) k = default := by + rw [mem_iff_contains] at mem + revert mem + simp_to_raw using Raw₀.Const.get!_diff_of_contains_right + +theorem get!_diff_of_not_mem_left [EquivBEq α] [LawfulHashable α] [Inhabited β] (h₁ : m₁.WF) (h₂ : m₂.WF) + {k : α} (not_mem : ¬k ∈ m₁) : + Const.get! (m₁.diff m₂) k = default := by + rw [← contains_eq_false_iff_not_mem] at not_mem + revert not_mem + simp_to_raw using Raw₀.Const.get!_diff_of_contains_eq_false_left + +end Const + namespace Const variable {β : Type v} {m : Raw α (fun _ => β)} diff --git a/src/Std/Data/HashMap/Raw.lean b/src/Std/Data/HashMap/Raw.lean index 2c4ca703d606..c074ab902417 100644 --- a/src/Std/Data/HashMap/Raw.lean +++ b/src/Std/Data/HashMap/Raw.lean @@ -348,6 +348,9 @@ theorem WF.unitOfList [BEq α] [Hashable α] {l : List α} : (unitOfList l).WF : theorem WF.union [BEq α] [Hashable α] {m₁ m₂ : Raw α β} (h₁ : m₁.WF) (h₂ : m₂.WF) : (m₁.union m₂).WF := ⟨DHashMap.Raw.WF.union h₁.out h₂.out⟩ +theorem WF.inter [BEq α] [Hashable α] {m₁ m₂ : Raw α β} (h₁ : m₁.WF) (h₂ : m₂.WF) : (m₁.inter m₂).WF := + ⟨DHashMap.Raw.WF.inter h₁.out h₂.out⟩ + end Raw end HashMap From 150f3cb28a88299332b21c60cea2e96024a605b3 Mon Sep 17 00:00:00 2001 From: Wojciech Rozowski Date: Tue, 18 Nov 2025 15:57:03 +0100 Subject: [PATCH 75/88] Add all lemmas --- src/Std/Data/DHashMap/Raw.lean | 4 + src/Std/Data/HashMap/Basic.lean | 5 + src/Std/Data/HashMap/Lemmas.lean | 198 ++++++++++++++++++++++++++ src/Std/Data/HashMap/Raw.lean | 8 ++ src/Std/Data/HashMap/RawLemmas.lean | 209 ++++++++++++++++++++++++++++ src/Std/Data/HashSet/Basic.lean | 11 ++ src/Std/Data/HashSet/Lemmas.lean | 126 +++++++++++++++++ src/Std/Data/HashSet/Raw.lean | 17 +++ src/Std/Data/HashSet/RawLemmas.lean | 144 +++++++++++++++++++ 9 files changed, 722 insertions(+) diff --git a/src/Std/Data/DHashMap/Raw.lean b/src/Std/Data/DHashMap/Raw.lean index 8abe1c01120b..16709cf27286 100644 --- a/src/Std/Data/DHashMap/Raw.lean +++ b/src/Std/Data/DHashMap/Raw.lean @@ -807,6 +807,10 @@ theorem WF.diff₀ [BEq α] [Hashable α] {m₁ m₂ : Raw α β} (h₁ : m₁.W . exact @WF.filter₀ α β _ _ m₁ h₁.size_buckets_pos (fun k x => !Raw₀.contains ⟨m₂, h₂.size_buckets_pos⟩ k) h₁ . exact (Raw₀.eraseManyEntries ⟨m₁, h₁.size_buckets_pos⟩ m₂).2 _ WF.erase₀ h₁ +theorem WF.diff [BEq α] [Hashable α] {m₁ m₂ : Raw α β} (h₁ : m₁.WF) (h₂ : m₂.WF) : (m₁.diff m₂ : Raw α β).WF := by + simp [Std.DHashMap.Raw.diff, h₁.size_buckets_pos, h₂.size_buckets_pos] + exact WF.diff₀ h₁ h₂ + end WF end Raw diff --git a/src/Std/Data/HashMap/Basic.lean b/src/Std/Data/HashMap/Basic.lean index 941caf239a85..d1f066a251ba 100644 --- a/src/Std/Data/HashMap/Basic.lean +++ b/src/Std/Data/HashMap/Basic.lean @@ -274,6 +274,11 @@ instance [BEq α] [Hashable α] : Union (HashMap α β) := ⟨union⟩ instance [BEq α] [Hashable α] : Inter (HashMap α β) := ⟨inter⟩ +@[inherit_doc DHashMap.inter, inline] def diff [BEq α] [Hashable α] (m₁ m₂ : HashMap α β) : HashMap α β := + ⟨DHashMap.diff m₁.inner m₂.inner⟩ + +instance [BEq α] [Hashable α] : SDiff (HashMap α β) := ⟨diff⟩ + section Unverified /-! We currently do not provide lemmas for the functions below. -/ diff --git a/src/Std/Data/HashMap/Lemmas.lean b/src/Std/Data/HashMap/Lemmas.lean index d81fcff5f7ea..d3fef4142620 100644 --- a/src/Std/Data/HashMap/Lemmas.lean +++ b/src/Std/Data/HashMap/Lemmas.lean @@ -1769,6 +1769,204 @@ theorem isEmpty_inter_iff [EquivBEq α] [LawfulHashable α] : end Inter +section Diff + +variable {β : Type v} + +variable {m₁ m₂ : HashMap α β} + +@[simp] +theorem diff_eq : m₁.diff m₂ = m₁ \ m₂ := by + simp only [SDiff.sdiff] + +/- contains -/ +@[simp] +theorem contains_diff [EquivBEq α] [LawfulHashable α] {k : α} : + (m₁ \ m₂).contains k = (m₁.contains k && !m₂.contains k) := + @DHashMap.contains_diff _ _ _ _ m₁.inner m₂.inner _ _ k + +/- mem -/ +@[simp] +theorem mem_diff_iff [EquivBEq α] [LawfulHashable α] {k : α} : + k ∈ m₁ \ m₂ ↔ k ∈ m₁ ∧ k ∉ m₂ := + @DHashMap.mem_diff_iff _ _ _ _ m₁.inner m₂.inner _ _ k + +theorem not_mem_diff_of_not_mem_left [EquivBEq α] [LawfulHashable α] {k : α} + (not_mem : k ∉ m₁) : + k ∉ m₁ \ m₂ := + @DHashMap.not_mem_diff_of_not_mem_left _ _ _ _ m₁.inner m₂.inner _ _ k not_mem + +theorem not_mem_diff_of_mem_right [EquivBEq α] [LawfulHashable α] {k : α} + (mem : k ∈ m₂) : + k ∉ m₁ \ m₂ := + @DHashMap.not_mem_diff_of_mem_right _ _ _ _ m₁.inner m₂.inner _ _ k mem + +/- get? -/ +theorem get?_diff [EquivBEq α] [LawfulHashable α] {k : α} : + (m₁ \ m₂).get? k = + if k ∈ m₂ then none else m₁.get? k := + @DHashMap.Const.get?_diff _ _ _ _ m₁.inner m₂.inner _ _ k + +theorem get?_diff_of_not_mem_right [EquivBEq α] [LawfulHashable α] + {k : α} (not_mem : k ∉ m₂) : + (m₁ \ m₂).get? k = m₁.get? k := + @DHashMap.Const.get?_diff_of_not_mem_right _ _ _ _ m₁.inner m₂.inner _ _ k not_mem + +theorem get?_diff_of_not_mem_left [EquivBEq α] [LawfulHashable α] + {k : α} (not_mem : k ∉ m₁) : + (m₁ \ m₂).get? k = none := + @DHashMap.Const.get?_diff_of_not_mem_left _ _ _ _ m₁.inner m₂.inner _ _ k not_mem + +theorem get?_diff_of_mem_right [EquivBEq α] [LawfulHashable α] + {k : α} (mem : k ∈ m₂) : + (m₁ \ m₂).get? k = none := + @DHashMap.Const.get?_diff_of_mem_right _ _ _ _ m₁.inner m₂.inner _ _ k mem + +/- get -/ +@[simp] theorem get_diff [EquivBEq α] [LawfulHashable α] + {k : α} {h_mem : k ∈ m₁ \ m₂} : + (m₁ \ m₂).get k h_mem = + m₁.get k ((mem_diff_iff.1 h_mem).1) := + @DHashMap.Const.get_diff _ _ _ _ m₁.inner m₂.inner _ _ k h_mem + +/- getD -/ +theorem getD_diff [EquivBEq α] [LawfulHashable α] {k : α} {fallback : β} : + (m₁ \ m₂).getD k fallback = + if k ∈ m₂ then fallback else m₁.getD k fallback := + @DHashMap.Const.getD_diff _ _ _ _ m₁.inner m₂.inner _ _ k fallback + +theorem getD_diff_of_not_mem_right [EquivBEq α] [LawfulHashable α] + {k : α} {fallback : β} (not_mem : k ∉ m₂) : + (m₁ \ m₂).getD k fallback = m₁.getD k fallback := + @DHashMap.Const.getD_diff_of_not_mem_right _ _ _ _ m₁.inner m₂.inner _ _ k fallback not_mem + +theorem getD_diff_of_mem_right [EquivBEq α] [LawfulHashable α] + {k : α} {fallback : β} (mem : k ∈ m₂) : + (m₁ \ m₂).getD k fallback = fallback := + @DHashMap.Const.getD_diff_of_mem_right _ _ _ _ m₁.inner m₂.inner _ _ k fallback mem + +theorem getD_diff_of_not_mem_left [EquivBEq α] [LawfulHashable α] + {k : α} {fallback : β} (not_mem : k ∉ m₁) : + (m₁ \ m₂).getD k fallback = fallback := + @DHashMap.Const.getD_diff_of_not_mem_left _ _ _ _ m₁.inner m₂.inner _ _ k fallback not_mem + +/- get! -/ +theorem get!_diff [EquivBEq α] [LawfulHashable α] {k : α} [Inhabited β] : + (m₁ \ m₂).get! k = + if k ∈ m₂ then default else m₁.get! k := + @DHashMap.Const.get!_diff _ _ _ _ m₁.inner m₂.inner _ _ _ k + +theorem get!_diff_of_not_mem_right [EquivBEq α] [LawfulHashable α] + {k : α} [Inhabited β] (not_mem : k ∉ m₂) : + (m₁ \ m₂).get! k = m₁.get! k := + @DHashMap.Const.get!_diff_of_not_mem_right _ _ _ _ m₁.inner m₂.inner _ _ _ k not_mem + +theorem get!_diff_of_mem_right [EquivBEq α] [LawfulHashable α] + {k : α} [Inhabited β] (mem : k ∈ m₂) : + (m₁ \ m₂).get! k = default := + @DHashMap.Const.get!_diff_of_mem_right _ _ _ _ m₁.inner m₂.inner _ _ _ k mem + +theorem get!_diff_of_not_mem_left [EquivBEq α] [LawfulHashable α] + {k : α} [Inhabited β] (not_mem : k ∉ m₁) : + (m₁ \ m₂).get! k = default := + @DHashMap.Const.get!_diff_of_not_mem_left _ _ _ _ m₁.inner m₂.inner _ _ _ k not_mem + +/- getKey? -/ +theorem getKey?_diff [EquivBEq α] [LawfulHashable α] {k : α} : + (m₁ \ m₂).getKey? k = + if k ∈ m₂ then none else m₁.getKey? k := + @DHashMap.getKey?_diff _ _ _ _ m₁.inner m₂.inner _ _ k + +theorem getKey?_diff_of_not_mem_right [EquivBEq α] [LawfulHashable α] + {k : α} (not_mem : k ∉ m₂) : + (m₁ \ m₂).getKey? k = m₁.getKey? k := + @DHashMap.getKey?_diff_of_not_mem_right _ _ _ _ m₁.inner m₂.inner _ _ k not_mem + +theorem getKey?_diff_of_not_mem_left [EquivBEq α] [LawfulHashable α] + {k : α} (not_mem : k ∉ m₁) : + (m₁ \ m₂).getKey? k = none := + @DHashMap.getKey?_diff_of_not_mem_left _ _ _ _ m₁.inner m₂.inner _ _ k not_mem + +theorem getKey?_diff_of_mem_right [EquivBEq α] [LawfulHashable α] + {k : α} (mem : k ∈ m₂) : + (m₁ \ m₂).getKey? k = none := + @DHashMap.getKey?_diff_of_mem_right _ _ _ _ m₁.inner m₂.inner _ _ k mem + +/- getKey -/ +@[simp] theorem getKey_diff [EquivBEq α] [LawfulHashable α] + {k : α} {h_mem : k ∈ m₁ \ m₂} : + (m₁ \ m₂).getKey k h_mem = + m₁.getKey k ((mem_diff_iff.1 h_mem).1) := + @DHashMap.getKey_diff _ _ _ _ m₁.inner m₂.inner _ _ k h_mem + +/- getKeyD -/ +theorem getKeyD_diff [EquivBEq α] [LawfulHashable α] {k fallback : α} : + (m₁ \ m₂).getKeyD k fallback = + if k ∈ m₂ then fallback else m₁.getKeyD k fallback := + @DHashMap.getKeyD_diff _ _ _ _ m₁.inner m₂.inner _ _ k fallback + +theorem getKeyD_diff_of_not_mem_right [EquivBEq α] [LawfulHashable α] + {k fallback : α} (not_mem : k ∉ m₂) : + (m₁ \ m₂).getKeyD k fallback = m₁.getKeyD k fallback := + @DHashMap.getKeyD_diff_of_not_mem_right _ _ _ _ m₁.inner m₂.inner _ _ k fallback not_mem + +theorem getKeyD_diff_of_mem_right [EquivBEq α] [LawfulHashable α] + {k fallback : α} (mem : k ∈ m₂) : + (m₁ \ m₂).getKeyD k fallback = fallback := + @DHashMap.getKeyD_diff_of_mem_right _ _ _ _ m₁.inner m₂.inner _ _ k fallback mem + +theorem getKeyD_diff_of_not_mem_left [EquivBEq α] [LawfulHashable α] + {k fallback : α} (not_mem : k ∉ m₁) : + (m₁ \ m₂).getKeyD k fallback = fallback := + @DHashMap.getKeyD_diff_of_not_mem_left _ _ _ _ m₁.inner m₂.inner _ _ k fallback not_mem + +/- getKey! -/ +theorem getKey!_diff [EquivBEq α] [LawfulHashable α] [Inhabited α] {k : α} : + (m₁ \ m₂).getKey! k = + if k ∈ m₂ then default else m₁.getKey! k := + @DHashMap.getKey!_diff _ _ _ _ m₁.inner m₂.inner _ _ _ k + +theorem getKey!_diff_of_not_mem_right [EquivBEq α] [LawfulHashable α] [Inhabited α] + {k : α} (not_mem : k ∉ m₂) : + (m₁ \ m₂).getKey! k = m₁.getKey! k := + @DHashMap.getKey!_diff_of_not_mem_right _ _ _ _ m₁.inner m₂.inner _ _ _ k not_mem + +theorem getKey!_diff_of_mem_right [EquivBEq α] [LawfulHashable α] [Inhabited α] + {k : α} (mem : k ∈ m₂) : + (m₁ \ m₂).getKey! k = default := + @DHashMap.getKey!_diff_of_mem_right _ _ _ _ m₁.inner m₂.inner _ _ _ k mem + +theorem getKey!_diff_of_not_mem_left [EquivBEq α] [LawfulHashable α] [Inhabited α] + {k : α} (not_mem : k ∉ m₁) : + (m₁ \ m₂).getKey! k = default := + @DHashMap.getKey!_diff_of_not_mem_left _ _ _ _ m₁.inner m₂.inner _ _ _ k not_mem + +/- size -/ +theorem size_diff_le_size_left [EquivBEq α] [LawfulHashable α] : + (m₁ \ m₂).size ≤ m₁.size := + @DHashMap.size_diff_le_size_left _ _ _ _ m₁.inner m₂.inner _ _ + +theorem size_diff_eq_size_left [EquivBEq α] [LawfulHashable α] + (h : ∀ (a : α), a ∈ m₁ → a ∉ m₂) : + (m₁ \ m₂).size = m₁.size := + @DHashMap.size_diff_eq_size_left _ _ _ _ m₁.inner m₂.inner _ _ h + +theorem size_diff_add_size_inter_eq_size_left [EquivBEq α] [LawfulHashable α] : + (m₁ \ m₂).size + (m₁ ∩ m₂).size = m₁.size := + @DHashMap.size_diff_add_size_inter_eq_size_left _ _ _ _ m₁.inner m₂.inner _ _ + +/- isEmpty -/ +@[simp] +theorem isEmpty_diff_left [EquivBEq α] [LawfulHashable α] (h : m₁.isEmpty) : + (m₁ \ m₂).isEmpty = true := + @DHashMap.isEmpty_diff_left α _ _ _ m₁.inner m₂.inner _ _ h + +theorem isEmpty_diff_iff [EquivBEq α] [LawfulHashable α] : + (m₁ \ m₂).isEmpty ↔ ∀ k, k ∈ m₁ → k ∈ m₂ := + @DHashMap.isEmpty_diff_iff α _ _ _ m₁.inner m₂.inner _ _ + +end Diff + -- As `insertManyIfNewUnit` is really an implementation detail for `HashSet.insertMany`, -- we do not add `@[grind]` annotations to any of its lemmas. diff --git a/src/Std/Data/HashMap/Raw.lean b/src/Std/Data/HashMap/Raw.lean index c074ab902417..62b58015c8ae 100644 --- a/src/Std/Data/HashMap/Raw.lean +++ b/src/Std/Data/HashMap/Raw.lean @@ -237,10 +237,15 @@ instance {m : Type w → Type w'} : ForIn m (Raw α β) (α × β) where @[inherit_doc DHashMap.Raw.inter, inline] def inter [BEq α] [Hashable α] (m₁ m₂ : Raw α β) : Raw α β := ⟨DHashMap.Raw.inter m₁.inner m₂.inner⟩ +@[inherit_doc DHashMap.Raw.inter, inline] def diff [BEq α] [Hashable α] (m₁ m₂ : Raw α β) : Raw α β := + ⟨DHashMap.Raw.diff m₁.inner m₂.inner⟩ + instance [BEq α] [Hashable α] : Union (Raw α β) := ⟨union⟩ instance [BEq α] [Hashable α] : Inter (Raw α β) := ⟨inter⟩ +instance [BEq α] [Hashable α] : SDiff (Raw α β) := ⟨diff⟩ + section Unverified @[inline, inherit_doc DHashMap.Raw.filterMap] def filterMap {γ : Type w} (f : α → β → Option γ) @@ -351,6 +356,9 @@ theorem WF.union [BEq α] [Hashable α] {m₁ m₂ : Raw α β} (h₁ : m₁.WF) theorem WF.inter [BEq α] [Hashable α] {m₁ m₂ : Raw α β} (h₁ : m₁.WF) (h₂ : m₂.WF) : (m₁.inter m₂).WF := ⟨DHashMap.Raw.WF.inter h₁.out h₂.out⟩ +theorem WF.diff [BEq α] [Hashable α] {m₁ m₂ : Raw α β} (h₁ : m₁.WF) (h₂ : m₂.WF) : (m₁.diff m₂).WF := + ⟨DHashMap.Raw.WF.diff h₁.out h₂.out⟩ + end Raw end HashMap diff --git a/src/Std/Data/HashMap/RawLemmas.lean b/src/Std/Data/HashMap/RawLemmas.lean index aff224922a3d..5cd7cdf9dc41 100644 --- a/src/Std/Data/HashMap/RawLemmas.lean +++ b/src/Std/Data/HashMap/RawLemmas.lean @@ -1759,6 +1759,215 @@ theorem isEmpty_inter_iff [EquivBEq α] [LawfulHashable α] (h₁ : m₁.WF) (h end Inter +section Diff + +variable {β : Type v} + +variable {m₁ m₂ : Raw α β} + +@[simp] +theorem diff_eq : m₁.diff m₂ = m₁ \ m₂ := by + simp only [SDiff.sdiff] + +/- contains -/ +@[simp] +theorem contains_diff [EquivBEq α] [LawfulHashable α] (h₁ : m₁.WF) + (h₂ : m₂.WF) {k : α} : + (m₁ \ m₂).contains k = (m₁.contains k && !m₂.contains k) := + @DHashMap.Raw.contains_diff _ _ _ _ m₁.inner m₂.inner _ _ h₁.out h₂.out k + +/- mem -/ +@[simp] +theorem mem_diff_iff [EquivBEq α] [LawfulHashable α] (h₁ : m₁.WF) + (h₂ : m₂.WF) {k : α} : + k ∈ m₁ \ m₂ ↔ k ∈ m₁ ∧ k ∉ m₂ := + @DHashMap.Raw.mem_diff_iff _ _ _ _ m₁.inner m₂.inner _ _ h₁.out h₂.out k + +theorem not_mem_diff_of_not_mem_left [EquivBEq α] [LawfulHashable α] + (h₁ : m₁.WF) (h₂ : m₂.WF) {k : α} + (not_mem : k ∉ m₁) : + k ∉ m₁ \ m₂ := + @DHashMap.Raw.not_mem_diff_of_not_mem_left _ _ _ _ m₁.inner m₂.inner _ _ h₁.out h₂.out k not_mem + +theorem not_mem_diff_of_mem_right [EquivBEq α] [LawfulHashable α] + (h₁ : m₁.WF) (h₂ : m₂.WF) {k : α} + (mem : k ∈ m₂) : + k ∉ m₁ \ m₂ := + @DHashMap.Raw.not_mem_diff_of_mem_right _ _ _ _ m₁.inner m₂.inner _ _ h₁.out h₂.out k mem + +/- get? -/ +theorem get?_diff [EquivBEq α] [LawfulHashable α] (h₁ : m₁.WF) (h₂ : m₂.WF) {k : α} : + get? (m₁ \ m₂) k = + if k ∈ m₂ then none else get? m₁ k := + @DHashMap.Raw.Const.get?_diff _ _ _ _ m₁.inner m₂.inner _ _ h₁.out h₂.out k + +theorem get?_diff_of_not_mem_right [EquivBEq α] [LawfulHashable α] (h₁ : m₁.WF) (h₂ : m₂.WF) + {k : α} (not_mem : k ∉ m₂) : + get? (m₁ \ m₂) k = get? m₁ k := + @DHashMap.Raw.Const.get?_diff_of_not_mem_right _ _ _ _ m₁.inner m₂.inner _ _ h₁.out h₂.out k not_mem + +theorem get?_diff_of_not_mem_left [EquivBEq α] [LawfulHashable α] (h₁ : m₁.WF) (h₂ : m₂.WF) + {k : α} (not_mem : k ∉ m₁) : + get? (m₁ \ m₂) k = none := + @DHashMap.Raw.Const.get?_diff_of_not_mem_left _ _ _ _ m₁.inner m₂.inner _ _ h₁.out h₂.out k not_mem + +theorem get?_diff_of_mem_right [EquivBEq α] [LawfulHashable α] (h₁ : m₁.WF) (h₂ : m₂.WF) + {k : α} (mem : k ∈ m₂) : + get? (m₁ \ m₂) k = none := + @DHashMap.Raw.Const.get?_diff_of_mem_right _ _ _ _ m₁.inner m₂.inner _ _ h₁.out h₂.out k mem + +/- get -/ +@[simp] theorem get_diff [EquivBEq α] [LawfulHashable α] (h₁ : m₁.WF) (h₂ : m₂.WF) + {k : α} {h_mem : k ∈ m₁ \ m₂} : + get (m₁ \ m₂) k h_mem = + get m₁ k ((mem_diff_iff h₁ h₂).1 h_mem).1 := + @DHashMap.Raw.Const.get_diff _ _ _ _ m₁.inner m₂.inner _ _ h₁.out h₂.out k h_mem + +/- getD -/ +theorem getD_diff [EquivBEq α] [LawfulHashable α] (h₁ : m₁.WF) (h₂ : m₂.WF) + {k : α} {fallback : β} : + getD (m₁ \ m₂) k fallback = + if k ∈ m₂ then fallback else getD m₁ k fallback := + @DHashMap.Raw.Const.getD_diff _ _ _ _ m₁.inner m₂.inner _ _ h₁.out h₂.out k fallback + +theorem getD_diff_of_not_mem_right [EquivBEq α] [LawfulHashable α] (h₁ : m₁.WF) (h₂ : m₂.WF) + {k : α} {fallback : β} (not_mem : k ∉ m₂) : + getD (m₁ \ m₂) k fallback = getD m₁ k fallback := + @DHashMap.Raw.Const.getD_diff_of_not_mem_right _ _ _ _ m₁.inner m₂.inner _ _ h₁.out h₂.out k fallback not_mem + +theorem getD_diff_of_mem_right [EquivBEq α] [LawfulHashable α] (h₁ : m₁.WF) (h₂ : m₂.WF) + {k : α} {fallback : β} (mem : k ∈ m₂) : + getD (m₁ \ m₂) k fallback = fallback := + @DHashMap.Raw.Const.getD_diff_of_mem_right _ _ _ _ m₁.inner m₂.inner _ _ h₁.out h₂.out k fallback mem + +theorem getD_diff_of_not_mem_left [EquivBEq α] [LawfulHashable α] (h₁ : m₁.WF) (h₂ : m₂.WF) + {k : α} {fallback : β} (not_mem : k ∉ m₁) : + getD (m₁ \ m₂) k fallback = fallback := + @DHashMap.Raw.Const.getD_diff_of_not_mem_left _ _ _ _ m₁.inner m₂.inner _ _ h₁.out h₂.out k fallback not_mem + +/- get! -/ +theorem get!_diff [EquivBEq α] [LawfulHashable α] [Inhabited β] (h₁ : m₁.WF) (h₂ : m₂.WF) {k : α} : + get! (m₁ \ m₂) k = + if k ∈ m₂ then default else get! m₁ k := + @DHashMap.Raw.Const.get!_diff _ _ _ _ m₁.inner m₂.inner _ _ _ h₁.out h₂.out k + +theorem get!_diff_of_not_mem_right [EquivBEq α] [LawfulHashable α] [Inhabited β] (h₁ : m₁.WF) (h₂ : m₂.WF) + {k : α} (not_mem : k ∉ m₂) : + get! (m₁ \ m₂) k = get! m₁ k := + @DHashMap.Raw.Const.get!_diff_of_not_mem_right _ _ _ _ m₁.inner m₂.inner _ _ _ h₁.out h₂.out k not_mem + +theorem get!_diff_of_mem_right [EquivBEq α] [LawfulHashable α] [Inhabited β] (h₁ : m₁.WF) (h₂ : m₂.WF) + {k : α} (mem : k ∈ m₂) : + get! (m₁ \ m₂) k = default := + @DHashMap.Raw.Const.get!_diff_of_mem_right _ _ _ _ m₁.inner m₂.inner _ _ _ h₁.out h₂.out k mem + +theorem get!_diff_of_not_mem_left [EquivBEq α] [LawfulHashable α] [Inhabited β] (h₁ : m₁.WF) (h₂ : m₂.WF) + {k : α} (not_mem : k ∉ m₁) : + get! (m₁ \ m₂) k = default := + @DHashMap.Raw.Const.get!_diff_of_not_mem_left _ _ _ _ m₁.inner m₂.inner _ _ _ h₁.out h₂.out k not_mem + +/- getKey? -/ +theorem getKey?_diff [EquivBEq α] [LawfulHashable α] + (h₁ : m₁.WF) (h₂ : m₂.WF) {k : α} : + (m₁ \ m₂).getKey? k = + if k ∈ m₂ then none else m₁.getKey? k := + @DHashMap.Raw.getKey?_diff _ _ _ _ m₁.inner m₂.inner _ _ h₁.out h₂.out k + +theorem getKey?_diff_of_not_mem_right [EquivBEq α] [LawfulHashable α] + (h₁ : m₁.WF) (h₂ : m₂.WF) {k : α} (not_mem : k ∉ m₂) : + (m₁ \ m₂).getKey? k = m₁.getKey? k := + @DHashMap.Raw.getKey?_diff_of_not_mem_right _ _ _ _ m₁.inner m₂.inner _ _ h₁.out h₂.out k not_mem + +theorem getKey?_diff_of_not_mem_left [EquivBEq α] [LawfulHashable α] + (h₁ : m₁.WF) (h₂ : m₂.WF) {k : α} (not_mem : k ∉ m₁) : + (m₁ \ m₂).getKey? k = none := + @DHashMap.Raw.getKey?_diff_of_not_mem_left _ _ _ _ m₁.inner m₂.inner _ _ h₁.out h₂.out k not_mem + +theorem getKey?_diff_of_mem_right [EquivBEq α] [LawfulHashable α] + (h₁ : m₁.WF) (h₂ : m₂.WF) {k : α} (mem : k ∈ m₂) : + (m₁ \ m₂).getKey? k = none := + @DHashMap.Raw.getKey?_diff_of_mem_right _ _ _ _ m₁.inner m₂.inner _ _ h₁.out h₂.out k mem + +/- getKey -/ +@[simp] theorem getKey_diff [EquivBEq α] [LawfulHashable α] (h₁ : m₁.WF) (h₂ : m₂.WF) + {k : α} {h_mem : k ∈ m₁ \ m₂} : + (m₁ \ m₂).getKey k h_mem = + m₁.getKey k ((mem_diff_iff h₁ h₂).1 h_mem).1 := + @DHashMap.Raw.getKey_diff _ _ _ _ m₁.inner m₂.inner _ _ h₁.out h₂.out k h_mem + +/- getKeyD -/ +theorem getKeyD_diff [EquivBEq α] [LawfulHashable α] (h₁ : m₁.WF) + (h₂ : m₂.WF) {k fallback : α} : + (m₁ \ m₂).getKeyD k fallback = + if k ∈ m₂ then fallback else m₁.getKeyD k fallback := + @DHashMap.Raw.getKeyD_diff _ _ _ _ m₁.inner m₂.inner _ _ h₁.out h₂.out k fallback + +theorem getKeyD_diff_of_not_mem_right [EquivBEq α] [LawfulHashable α] (h₁ : m₁.WF) + (h₂ : m₂.WF) {k fallback : α} (not_mem : k ∉ m₂) : + (m₁ \ m₂).getKeyD k fallback = m₁.getKeyD k fallback := + @DHashMap.Raw.getKeyD_diff_of_not_mem_right _ _ _ _ m₁.inner m₂.inner _ _ h₁.out h₂.out k fallback not_mem + +theorem getKeyD_diff_of_mem_right [EquivBEq α] [LawfulHashable α] (h₁ : m₁.WF) + (h₂ : m₂.WF) {k fallback : α} (mem : k ∈ m₂) : + (m₁ \ m₂).getKeyD k fallback = fallback := + @DHashMap.Raw.getKeyD_diff_of_mem_right _ _ _ _ m₁.inner m₂.inner _ _ h₁.out h₂.out k fallback mem + +theorem getKeyD_diff_of_not_mem_left [EquivBEq α] [LawfulHashable α] (h₁ : m₁.WF) + (h₂ : m₂.WF) {k fallback : α} (not_mem : k ∉ m₁) : + (m₁ \ m₂).getKeyD k fallback = fallback := + @DHashMap.Raw.getKeyD_diff_of_not_mem_left _ _ _ _ m₁.inner m₂.inner _ _ h₁.out h₂.out k fallback not_mem + +/- getKey! -/ +theorem getKey!_diff [EquivBEq α] [LawfulHashable α] [Inhabited α] (h₁ : m₁.WF) + (h₂ : m₂.WF) {k : α} : + (m₁ \ m₂).getKey! k = + if k ∈ m₂ then default else m₁.getKey! k := + @DHashMap.Raw.getKey!_diff _ _ _ _ m₁.inner m₂.inner _ _ _ h₁.out h₂.out k + +theorem getKey!_diff_of_not_mem_right [EquivBEq α] [LawfulHashable α] [Inhabited α] (h₁ : m₁.WF) + (h₂ : m₂.WF) {k : α} (not_mem : k ∉ m₂) : + (m₁ \ m₂).getKey! k = m₁.getKey! k := + @DHashMap.Raw.getKey!_diff_of_not_mem_right _ _ _ _ m₁.inner m₂.inner _ _ _ h₁.out h₂.out k not_mem + +theorem getKey!_diff_of_mem_right [EquivBEq α] [LawfulHashable α] [Inhabited α] (h₁ : m₁.WF) + (h₂ : m₂.WF) {k : α} (mem : k ∈ m₂) : + (m₁ \ m₂).getKey! k = default := + @DHashMap.Raw.getKey!_diff_of_mem_right _ _ _ _ m₁.inner m₂.inner _ _ _ h₁.out h₂.out k mem + +theorem getKey!_diff_of_not_mem_left [EquivBEq α] [LawfulHashable α] [Inhabited α] (h₁ : m₁.WF) + (h₂ : m₂.WF) {k : α} (not_mem : k ∉ m₁) : + (m₁ \ m₂).getKey! k = default := + @DHashMap.Raw.getKey!_diff_of_not_mem_left _ _ _ _ m₁.inner m₂.inner _ _ _ h₁.out h₂.out k not_mem + +/- size -/ +theorem size_diff_le_size_left [EquivBEq α] [LawfulHashable α] + (h₁ : m₁.WF) (h₂ : m₂.WF) : + (m₁ \ m₂).size ≤ m₁.size := + @DHashMap.Raw.size_diff_le_size_left _ _ _ _ m₁.inner m₂.inner _ _ h₁.out h₂.out + +theorem size_diff_eq_size_left [EquivBEq α] [LawfulHashable α] + (h₁ : m₁.WF) (h₂ : m₂.WF) + (h : ∀ (a : α), a ∈ m₁ → a ∉ m₂) : + (m₁ \ m₂).size = m₁.size := + @DHashMap.Raw.size_diff_eq_size_left _ _ _ _ m₁.inner m₂.inner _ _ h₁.out h₂.out h + +theorem size_diff_add_size_inter_eq_size_left [EquivBEq α] [LawfulHashable α] + (h₁ : m₁.WF) (h₂ : m₂.WF) : + (m₁ \ m₂).size + (m₁ ∩ m₂).size = m₁.size := + @DHashMap.Raw.size_diff_add_size_inter_eq_size_left _ _ _ _ m₁.inner m₂.inner _ _ h₁.out h₂.out + +/- isEmpty -/ +@[simp] +theorem isEmpty_diff_left [EquivBEq α] [LawfulHashable α] (h₁ : m₁.WF) (h₂ : m₂.WF) (h : m₁.isEmpty) : + (m₁ \ m₂).isEmpty = true := + @DHashMap.Raw.isEmpty_diff_left _ _ _ _ m₁.inner m₂.inner _ _ h₁.out h₂.out h + +theorem isEmpty_diff_iff [EquivBEq α] [LawfulHashable α] (h₁ : m₁.WF) (h₂ : m₂.WF) : + (m₁ \ m₂).isEmpty ↔ ∀ k, k ∈ m₁ → k ∈ m₂ := + @DHashMap.Raw.isEmpty_diff_iff _ _ _ _ m₁.inner m₂.inner _ _ h₁.out h₂.out + +end Diff + 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) : diff --git a/src/Std/Data/HashSet/Basic.lean b/src/Std/Data/HashSet/Basic.lean index f59584ac4cb4..400c10d2bbe0 100644 --- a/src/Std/Data/HashSet/Basic.lean +++ b/src/Std/Data/HashSet/Basic.lean @@ -259,6 +259,17 @@ This function always iterates through the smaller set, so the expected runtime i instance [BEq α] [Hashable α] : Inter (HashSet α) := ⟨inter⟩ +/-- +Computes the difference of the given hash sets. + +This function always iterates through the smaller set, so the expected runtime is +`O(min(m₁.size, m₂.size))`. +-/ +@[inline] def diff [BEq α] [Hashable α] (m₁ m₂ : HashSet α) : HashSet α := + ⟨HashMap.diff m₁.inner m₂.inner⟩ + +instance [BEq α] [Hashable α] : SDiff (HashSet α) := ⟨diff⟩ + section Unverified /-! We currently do not provide lemmas for the functions below. -/ diff --git a/src/Std/Data/HashSet/Lemmas.lean b/src/Std/Data/HashSet/Lemmas.lean index 3477057f09bf..14c06ca4d10c 100644 --- a/src/Std/Data/HashSet/Lemmas.lean +++ b/src/Std/Data/HashSet/Lemmas.lean @@ -1062,6 +1062,132 @@ theorem isEmpty_inter_iff [EquivBEq α] [LawfulHashable α] : end Inter +section Diff + +variable {m₁ m₂ : HashSet α} + +@[simp] +theorem diff_eq : m₁.diff m₂ = m₁ \ m₂ := by + simp only [SDiff.sdiff] + +/- contains -/ +@[simp] +theorem contains_diff [EquivBEq α] [LawfulHashable α] {k : α} : + (m₁ \ m₂).contains k = (m₁.contains k && !m₂.contains k) := + @HashMap.contains_diff _ _ _ _ m₁.inner m₂.inner _ _ k + +/- mem -/ +@[simp] +theorem mem_diff_iff [EquivBEq α] [LawfulHashable α] {k : α} : + k ∈ m₁ \ m₂ ↔ k ∈ m₁ ∧ k ∉ m₂ := + @HashMap.mem_diff_iff _ _ _ _ m₁.inner m₂.inner _ _ k + +theorem not_mem_diff_of_not_mem_left [EquivBEq α] [LawfulHashable α] {k : α} + (not_mem : k ∉ m₁) : + k ∉ m₁ \ m₂ := + @HashMap.not_mem_diff_of_not_mem_left _ _ _ _ m₁.inner m₂.inner _ _ k not_mem + +theorem not_mem_diff_of_mem_right [EquivBEq α] [LawfulHashable α] {k : α} + (mem : k ∈ m₂) : + k ∉ m₁ \ m₂ := + @HashMap.not_mem_diff_of_mem_right _ _ _ _ m₁.inner m₂.inner _ _ k mem + +/- get? -/ +theorem get?_diff [EquivBEq α] [LawfulHashable α] {k : α} : + (m₁ \ m₂).get? k = + if k ∈ m₂ then none else m₁.get? k := + @HashMap.getKey?_diff _ _ _ _ m₁.inner m₂.inner _ _ k + +theorem get?_diff_of_not_mem_right [EquivBEq α] [LawfulHashable α] + {k : α} (not_mem : k ∉ m₂) : + (m₁ \ m₂).get? k = m₁.get? k := + @HashMap.getKey?_diff_of_not_mem_right _ _ _ _ m₁.inner m₂.inner _ _ k not_mem + +theorem get?_diff_of_not_mem_left [EquivBEq α] [LawfulHashable α] + {k : α} (not_mem : k ∉ m₁) : + (m₁ \ m₂).get? k = none := + @HashMap.getKey?_diff_of_not_mem_left _ _ _ _ m₁.inner m₂.inner _ _ k not_mem + +theorem get?_diff_of_mem_right [EquivBEq α] [LawfulHashable α] + {k : α} (mem : k ∈ m₂) : + (m₁ \ m₂).get? k = none := + @HashMap.getKey?_diff_of_mem_right _ _ _ _ m₁.inner m₂.inner _ _ k mem + +/- get -/ +@[simp] theorem get_diff [EquivBEq α] [LawfulHashable α] + {k : α} {h_mem : k ∈ m₁ \ m₂} : + (m₁ \ m₂).get k h_mem = + m₁.get k ((mem_diff_iff.1 h_mem).1) := + @HashMap.getKey_diff _ _ _ _ m₁.inner m₂.inner _ _ k h_mem + +/- getD -/ +theorem getD_diff [EquivBEq α] [LawfulHashable α] {k fallback : α} : + (m₁ \ m₂).getD k fallback = + if k ∈ m₂ then fallback else m₁.getD k fallback := + @HashMap.getKeyD_diff _ _ _ _ m₁.inner m₂.inner _ _ k fallback + +theorem getD_diff_of_not_mem_right [EquivBEq α] [LawfulHashable α] + {k fallback : α} (not_mem : k ∉ m₂) : + (m₁ \ m₂).getD k fallback = m₁.getD k fallback := + @HashMap.getKeyD_diff_of_not_mem_right _ _ _ _ m₁.inner m₂.inner _ _ k fallback not_mem + +theorem getD_diff_of_mem_right [EquivBEq α] [LawfulHashable α] + {k fallback : α} (mem : k ∈ m₂) : + (m₁ \ m₂).getD k fallback = fallback := + @HashMap.getKeyD_diff_of_mem_right _ _ _ _ m₁.inner m₂.inner _ _ k fallback mem + +theorem getD_diff_of_not_mem_left [EquivBEq α] [LawfulHashable α] + {k fallback : α} (not_mem : k ∉ m₁) : + (m₁ \ m₂).getD k fallback = fallback := + @HashMap.getKeyD_diff_of_not_mem_left _ _ _ _ m₁.inner m₂.inner _ _ k fallback not_mem + +/- get! -/ +theorem get!_diff [EquivBEq α] [LawfulHashable α] [Inhabited α] {k : α} : + (m₁ \ m₂).get! k = + if k ∈ m₂ then default else m₁.get! k := + @HashMap.getKey!_diff _ _ _ _ m₁.inner m₂.inner _ _ _ k + +theorem get!_diff_of_not_mem_right [EquivBEq α] [LawfulHashable α] [Inhabited α] + {k : α} (not_mem : k ∉ m₂) : + (m₁ \ m₂).get! k = m₁.get! k := + @HashMap.getKey!_diff_of_not_mem_right _ _ _ _ m₁.inner m₂.inner _ _ _ k not_mem + +theorem get!_diff_of_mem_right [EquivBEq α] [LawfulHashable α] [Inhabited α] + {k : α} (mem : k ∈ m₂) : + (m₁ \ m₂).get! k = default := + @HashMap.getKey!_diff_of_mem_right _ _ _ _ m₁.inner m₂.inner _ _ _ k mem + +theorem get!_diff_of_not_mem_left [EquivBEq α] [LawfulHashable α] [Inhabited α] + {k : α} (not_mem : k ∉ m₁) : + (m₁ \ m₂).get! k = default := + @HashMap.getKey!_diff_of_not_mem_left _ _ _ _ m₁.inner m₂.inner _ _ _ k not_mem + +/- size -/ +theorem size_diff_le_size_left [EquivBEq α] [LawfulHashable α] : + (m₁ \ m₂).size ≤ m₁.size := + @HashMap.size_diff_le_size_left _ _ _ _ m₁.inner m₂.inner _ _ + +theorem size_diff_eq_size_left [EquivBEq α] [LawfulHashable α] + (h : ∀ (a : α), a ∈ m₁ → a ∉ m₂) : + (m₁ \ m₂).size = m₁.size := + @HashMap.size_diff_eq_size_left _ _ _ _ m₁.inner m₂.inner _ _ h + +theorem size_diff_add_size_inter_eq_size_left [EquivBEq α] [LawfulHashable α] : + (m₁ \ m₂).size + (m₁ ∩ m₂).size = m₁.size := + @HashMap.size_diff_add_size_inter_eq_size_left _ _ _ _ m₁.inner m₂.inner _ _ + +/- isEmpty -/ +@[simp] +theorem isEmpty_diff_left [EquivBEq α] [LawfulHashable α] (h : m₁.isEmpty) : + (m₁ \ m₂).isEmpty = true := + @HashMap.isEmpty_diff_left α _ _ _ m₁.inner m₂.inner _ _ h + +theorem isEmpty_diff_iff [EquivBEq α] [LawfulHashable α] : + (m₁ \ m₂).isEmpty ↔ ∀ k, k ∈ m₁ → k ∈ m₂ := + @HashMap.isEmpty_diff_iff α _ _ _ m₁.inner m₂.inner _ _ + +end Diff + section @[simp, grind =] diff --git a/src/Std/Data/HashSet/Raw.lean b/src/Std/Data/HashSet/Raw.lean index 49824e3e7fb4..eb8ce8e028dc 100644 --- a/src/Std/Data/HashSet/Raw.lean +++ b/src/Std/Data/HashSet/Raw.lean @@ -243,6 +243,17 @@ This function always merges the smaller set into the larger set, so the expected instance [BEq α] [Hashable α] : Inter (Raw α) := ⟨inter⟩ +/-- +Computes the difference of the given hash sets. + +This function always iterates through the smaller, so the expected runtime is +`O(min(m₁.size, m₂.size))`. +-/ +@[inline] def diff [BEq α] [Hashable α] (m₁ m₂ : Raw α) : Raw α := + ⟨HashMap.Raw.diff m₁.inner m₂.inner⟩ + +instance [BEq α] [Hashable α] : SDiff (Raw α) := ⟨diff⟩ + section Unverified /-! We currently do not provide lemmas for the functions below. -/ @@ -328,6 +339,12 @@ theorem WF.ofList [BEq α] [Hashable α] {l : List α} : theorem WF.union [BEq α] [Hashable α] {m₁ m₂ : Raw α} (h₁ : m₁.WF) (h₂ : m₂.WF) : (m₁.union m₂).WF := ⟨HashMap.Raw.WF.union h₁.out h₂.out⟩ +theorem WF.inter [BEq α] [Hashable α] {m₁ m₂ : Raw α} (h₁ : m₁.WF) (h₂ : m₂.WF) : (m₁.inter m₂).WF := + ⟨HashMap.Raw.WF.inter h₁.out h₂.out⟩ + +theorem WF.diff [BEq α] [Hashable α] {m₁ m₂ : Raw α} (h₁ : m₁.WF) (h₂ : m₂.WF) : (m₁.diff m₂).WF := + ⟨HashMap.Raw.WF.diff h₁.out h₂.out⟩ + end Raw end HashSet diff --git a/src/Std/Data/HashSet/RawLemmas.lean b/src/Std/Data/HashSet/RawLemmas.lean index 3015d058b95c..da6305e85c1d 100644 --- a/src/Std/Data/HashSet/RawLemmas.lean +++ b/src/Std/Data/HashSet/RawLemmas.lean @@ -1129,6 +1129,150 @@ theorem isEmpty_inter_iff [EquivBEq α] [LawfulHashable α] (h₁ : m₁.WF) (h end Inter +section Diff + +variable {m₁ m₂ : Raw α} + +@[simp] +theorem diff_eq : m₁.diff m₂ = m₁ \ m₂ := by + simp only [SDiff.sdiff] + +/- contains -/ +@[simp] +theorem contains_diff [EquivBEq α] [LawfulHashable α] (h₁ : m₁.WF) + (h₂ : m₂.WF) {k : α} : + (m₁ \ m₂).contains k = (m₁.contains k && !m₂.contains k) := + @HashMap.Raw.contains_diff _ _ _ _ m₁.inner m₂.inner _ _ h₁.out h₂.out k + +/- mem -/ +@[simp] +theorem mem_diff_iff [EquivBEq α] [LawfulHashable α] (h₁ : m₁.WF) + (h₂ : m₂.WF) {k : α} : + k ∈ m₁ \ m₂ ↔ k ∈ m₁ ∧ k ∉ m₂ := + @HashMap.Raw.mem_diff_iff _ _ _ _ m₁.inner m₂.inner _ _ h₁.out h₂.out k + +theorem not_mem_diff_of_not_mem_left [EquivBEq α] [LawfulHashable α] + (h₁ : m₁.WF) (h₂ : m₂.WF) {k : α} + (not_mem : k ∉ m₁) : + k ∉ m₁ \ m₂ := + @HashMap.Raw.not_mem_diff_of_not_mem_left _ _ _ _ m₁.inner m₂.inner _ _ h₁.out h₂.out k not_mem + +theorem not_mem_diff_of_mem_right [EquivBEq α] [LawfulHashable α] + (h₁ : m₁.WF) (h₂ : m₂.WF) {k : α} + (mem : k ∈ m₂) : + k ∉ m₁ \ m₂ := + @HashMap.Raw.not_mem_diff_of_mem_right _ _ _ _ m₁.inner m₂.inner _ _ h₁.out h₂.out k mem + +/- get? -/ +theorem get?_diff [EquivBEq α] [LawfulHashable α] + (h₁ : m₁.WF) (h₂ : m₂.WF) + {k : α} : + (m₁ \ m₂).get? k = + if k ∈ m₂ then none else m₁.get? k := + @HashMap.Raw.getKey?_diff _ _ _ _ m₁.inner m₂.inner _ _ h₁.out h₂.out k + +theorem get?_diff_of_not_mem_right [EquivBEq α] [LawfulHashable α] + (h₁ : m₁.WF) (h₂ : m₂.WF) + {k : α} (not_mem : k ∉ m₂) : + (m₁ \ m₂).get? k = m₁.get? k := + @HashMap.Raw.getKey?_diff_of_not_mem_right _ _ _ _ m₁.inner m₂.inner _ _ h₁.out h₂.out k not_mem + +theorem get?_diff_of_not_mem_left [EquivBEq α] [LawfulHashable α] + (h₁ : m₁.WF) (h₂ : m₂.WF) + {k : α} (not_mem : k ∉ m₁) : + (m₁ \ m₂).get? k = none := + @HashMap.Raw.getKey?_diff_of_not_mem_left _ _ _ _ m₁.inner m₂.inner _ _ h₁.out h₂.out k not_mem + +theorem get?_diff_of_mem_right [EquivBEq α] [LawfulHashable α] + (h₁ : m₁.WF) (h₂ : m₂.WF) + {k : α} (mem : k ∈ m₂) : + (m₁ \ m₂).get? k = none := + @HashMap.Raw.getKey?_diff_of_mem_right _ _ _ _ m₁.inner m₂.inner _ _ h₁.out h₂.out k mem + +/- get -/ +@[simp] theorem get_diff [EquivBEq α] [LawfulHashable α] (h₁ : m₁.WF) (h₂ : m₂.WF) + {k : α} {h_mem : k ∈ m₁ \ m₂} : + (m₁ \ m₂).get k h_mem = + m₁.get k ((mem_diff_iff h₁ h₂).1 h_mem).1 := + @HashMap.Raw.getKey_diff _ _ _ _ m₁.inner m₂.inner _ _ h₁.out h₂.out k h_mem + +/- getD -/ +theorem getD_diff [EquivBEq α] [LawfulHashable α] (h₁ : m₁.WF) + (h₂ : m₂.WF) {k fallback : α} : + (m₁ \ m₂).getD k fallback = + if k ∈ m₂ then fallback else m₁.getD k fallback := + @HashMap.Raw.getKeyD_diff _ _ _ _ m₁.inner m₂.inner _ _ h₁.out h₂.out k fallback + +theorem getD_diff_of_not_mem_right [EquivBEq α] [LawfulHashable α] (h₁ : m₁.WF) + (h₂ : m₂.WF) {k fallback : α} (not_mem : k ∉ m₂) : + (m₁ \ m₂).getD k fallback = m₁.getD k fallback := + @HashMap.Raw.getKeyD_diff_of_not_mem_right _ _ _ _ m₁.inner m₂.inner _ _ h₁.out h₂.out k fallback not_mem + +theorem getD_diff_of_mem_right [EquivBEq α] [LawfulHashable α] (h₁ : m₁.WF) + (h₂ : m₂.WF) {k fallback : α} (mem : k ∈ m₂) : + (m₁ \ m₂).getD k fallback = fallback := + @HashMap.Raw.getKeyD_diff_of_mem_right _ _ _ _ m₁.inner m₂.inner _ _ h₁.out h₂.out k fallback mem + +theorem getD_diff_of_not_mem_left [EquivBEq α] [LawfulHashable α] (h₁ : m₁.WF) + (h₂ : m₂.WF) {k fallback : α} (not_mem : k ∉ m₁) : + (m₁ \ m₂).getD k fallback = fallback := + @HashMap.Raw.getKeyD_diff_of_not_mem_left _ _ _ _ m₁.inner m₂.inner _ _ h₁.out h₂.out k fallback not_mem + +/- get! -/ +theorem get!_diff [EquivBEq α] [LawfulHashable α] [Inhabited α] + (h₁ : m₁.WF) + (h₂ : m₂.WF) {k : α} : + (m₁ \ m₂).get! k = + if k ∈ m₂ then default else m₁.get! k := + @HashMap.Raw.getKey!_diff _ _ _ _ m₁.inner m₂.inner _ _ _ h₁.out h₂.out k + +theorem get!_diff_of_not_mem_right [Inhabited α] + [EquivBEq α] [LawfulHashable α] (h₁ : m₁.WF) (h₂ : m₂.WF) {k : α} + (not_mem : k ∉ m₂) : + (m₁ \ m₂).get! k = m₁.get! k := + @HashMap.Raw.getKey!_diff_of_not_mem_right _ _ _ _ m₁.inner m₂.inner _ _ _ h₁.out h₂.out k not_mem + +theorem get!_diff_of_mem_right [Inhabited α] + [EquivBEq α] [LawfulHashable α] (h₁ : m₁.WF) (h₂ : m₂.WF) {k : α} + (mem : k ∈ m₂) : + (m₁ \ m₂).get! k = default := + @HashMap.Raw.getKey!_diff_of_mem_right _ _ _ _ m₁.inner m₂.inner _ _ _ h₁.out h₂.out k mem + +theorem get!_diff_of_not_mem_left [Inhabited α] + [EquivBEq α] [LawfulHashable α] (h₁ : m₁.WF) (h₂ : m₂.WF) {k : α} + (not_mem : k ∉ m₁) : + (m₁ \ m₂).get! k = default := + @HashMap.Raw.getKey!_diff_of_not_mem_left _ _ _ _ m₁.inner m₂.inner _ _ _ h₁.out h₂.out k not_mem + +/- size -/ +theorem size_diff_le_size_left [EquivBEq α] [LawfulHashable α] (h₁ : m₁.WF) + (h₂ : m₂.WF) : + (m₁ \ m₂).size ≤ m₁.size := + @HashMap.Raw.size_diff_le_size_left _ _ _ _ m₁.inner m₂.inner _ _ h₁.out h₂.out + +theorem size_diff_eq_size_left [EquivBEq α] [LawfulHashable α] (h₁ : m₁.WF) + (h₂ : m₂.WF) + (h : ∀ (a : α), a ∈ m₁ → a ∉ m₂) : + (m₁ \ m₂).size = m₁.size := + @HashMap.Raw.size_diff_eq_size_left _ _ _ _ m₁.inner m₂.inner _ _ h₁.out h₂.out h + +theorem size_diff_add_size_inter_eq_size_left [EquivBEq α] [LawfulHashable α] + (h₁ : m₁.WF) (h₂ : m₂.WF) : + (m₁ \ m₂).size + (m₁ ∩ m₂).size = m₁.size := + @HashMap.Raw.size_diff_add_size_inter_eq_size_left _ _ _ _ m₁.inner m₂.inner _ _ h₁.out h₂.out + +/- isEmpty -/ +@[simp] +theorem isEmpty_diff_left [EquivBEq α] [LawfulHashable α] (h₁ : m₁.WF) (h₂ : m₂.WF) (h : m₁.isEmpty) : + (m₁ \ m₂).isEmpty = true := + @HashMap.Raw.isEmpty_diff_left _ _ _ _ m₁.inner m₂.inner _ _ h₁.out h₂.out h + +theorem isEmpty_diff_iff [EquivBEq α] [LawfulHashable α] (h₁ : m₁.WF) (h₂ : m₂.WF) : + (m₁ \ m₂).isEmpty ↔ ∀ k, k ∈ m₁ → k ∈ m₂ := + @HashMap.Raw.isEmpty_diff_iff _ _ _ _ m₁.inner m₂.inner _ _ h₁.out h₂.out + +end Diff + @[simp, grind =] theorem ofList_nil : ofList ([] : List α) = ∅ := From 8bd0307b60f0f155c056f36f6f42322d9ef6103e Mon Sep 17 00:00:00 2001 From: Wojciech Rozowski Date: Wed, 19 Nov 2025 10:24:19 +0100 Subject: [PATCH 76/88] Fix notation for WF lemmas --- src/Std/Data/DHashMap/Raw.lean | 12 ++++++------ src/Std/Data/HashMap/Raw.lean | 6 +++--- src/Std/Data/HashSet/Raw.lean | 6 +++--- 3 files changed, 12 insertions(+), 12 deletions(-) diff --git a/src/Std/Data/DHashMap/Raw.lean b/src/Std/Data/DHashMap/Raw.lean index 16709cf27286..17ba6a7c4abd 100644 --- a/src/Std/Data/DHashMap/Raw.lean +++ b/src/Std/Data/DHashMap/Raw.lean @@ -793,12 +793,12 @@ theorem WF.union₀ [BEq α] [Hashable α] {m₁ m₂ : Raw α β} (h₁ : m₁. . exact (Raw₀.insertManyIfNew ⟨m₂, h₂.size_buckets_pos⟩ m₁).2 _ WF.insertIfNew₀ h₂ . exact (Raw₀.insertMany ⟨m₁, h₁.size_buckets_pos⟩ m₂).2 _ WF.insert₀ h₁ -theorem WF.union [BEq α] [Hashable α] {m₁ m₂ : Raw α β} (h₁ : m₁.WF) (h₂ : m₂.WF) : (m₁.union m₂ : Raw α β).WF := by - simp [Std.DHashMap.Raw.union, h₁.size_buckets_pos, h₂.size_buckets_pos] +theorem WF.union [BEq α] [Hashable α] {m₁ m₂ : Raw α β} (h₁ : m₁.WF) (h₂ : m₂.WF) : (m₁ ∪ m₂ : Raw α β).WF := by + simp [Union.union, Std.DHashMap.Raw.union, h₁.size_buckets_pos, h₂.size_buckets_pos] exact WF.union₀ h₁ h₂ -theorem WF.inter [BEq α] [Hashable α] {m₁ m₂ : Raw α β} (h₁ : m₁.WF) (h₂ : m₂.WF) : (m₁.inter m₂ : Raw α β).WF := by - simp [Std.DHashMap.Raw.inter, h₁.size_buckets_pos, h₂.size_buckets_pos] +theorem WF.inter [BEq α] [Hashable α] {m₁ m₂ : Raw α β} (h₁ : m₁.WF) (h₂ : m₂.WF) : (m₁ ∩ m₂ : Raw α β).WF := by + simp [Inter.inter, Std.DHashMap.Raw.inter, h₁.size_buckets_pos, h₂.size_buckets_pos] exact WF.inter₀ h₁ h₂ theorem WF.diff₀ [BEq α] [Hashable α] {m₁ m₂ : Raw α β} (h₁ : m₁.WF) (h₂ : m₂.WF) : (Raw₀.diff ⟨m₁, h₁.size_buckets_pos⟩ ⟨m₂, h₂.size_buckets_pos⟩).val.WF := by @@ -807,8 +807,8 @@ theorem WF.diff₀ [BEq α] [Hashable α] {m₁ m₂ : Raw α β} (h₁ : m₁.W . exact @WF.filter₀ α β _ _ m₁ h₁.size_buckets_pos (fun k x => !Raw₀.contains ⟨m₂, h₂.size_buckets_pos⟩ k) h₁ . exact (Raw₀.eraseManyEntries ⟨m₁, h₁.size_buckets_pos⟩ m₂).2 _ WF.erase₀ h₁ -theorem WF.diff [BEq α] [Hashable α] {m₁ m₂ : Raw α β} (h₁ : m₁.WF) (h₂ : m₂.WF) : (m₁.diff m₂ : Raw α β).WF := by - simp [Std.DHashMap.Raw.diff, h₁.size_buckets_pos, h₂.size_buckets_pos] +theorem WF.diff [BEq α] [Hashable α] {m₁ m₂ : Raw α β} (h₁ : m₁.WF) (h₂ : m₂.WF) : (m₁ \ m₂ : Raw α β).WF := by + simp [SDiff.sdiff, Std.DHashMap.Raw.diff, h₁.size_buckets_pos, h₂.size_buckets_pos] exact WF.diff₀ h₁ h₂ end WF diff --git a/src/Std/Data/HashMap/Raw.lean b/src/Std/Data/HashMap/Raw.lean index 62b58015c8ae..930fd402aadf 100644 --- a/src/Std/Data/HashMap/Raw.lean +++ b/src/Std/Data/HashMap/Raw.lean @@ -350,13 +350,13 @@ theorem WF.ofList [BEq α] [Hashable α] {l : List (α × β)} : (ofList l).WF : theorem WF.unitOfList [BEq α] [Hashable α] {l : List α} : (unitOfList l).WF := ⟨DHashMap.Raw.WF.Const.unitOfList⟩ -theorem WF.union [BEq α] [Hashable α] {m₁ m₂ : Raw α β} (h₁ : m₁.WF) (h₂ : m₂.WF) : (m₁.union m₂).WF := +theorem WF.union [BEq α] [Hashable α] {m₁ m₂ : Raw α β} (h₁ : m₁.WF) (h₂ : m₂.WF) : (m₁ ∪ m₂).WF := ⟨DHashMap.Raw.WF.union h₁.out h₂.out⟩ -theorem WF.inter [BEq α] [Hashable α] {m₁ m₂ : Raw α β} (h₁ : m₁.WF) (h₂ : m₂.WF) : (m₁.inter m₂).WF := +theorem WF.inter [BEq α] [Hashable α] {m₁ m₂ : Raw α β} (h₁ : m₁.WF) (h₂ : m₂.WF) : (m₁ ∩ m₂).WF := ⟨DHashMap.Raw.WF.inter h₁.out h₂.out⟩ -theorem WF.diff [BEq α] [Hashable α] {m₁ m₂ : Raw α β} (h₁ : m₁.WF) (h₂ : m₂.WF) : (m₁.diff m₂).WF := +theorem WF.diff [BEq α] [Hashable α] {m₁ m₂ : Raw α β} (h₁ : m₁.WF) (h₂ : m₂.WF) : (m₁ \ m₂).WF := ⟨DHashMap.Raw.WF.diff h₁.out h₂.out⟩ end Raw diff --git a/src/Std/Data/HashSet/Raw.lean b/src/Std/Data/HashSet/Raw.lean index eb8ce8e028dc..d98c1eeb3fae 100644 --- a/src/Std/Data/HashSet/Raw.lean +++ b/src/Std/Data/HashSet/Raw.lean @@ -336,13 +336,13 @@ theorem WF.ofList [BEq α] [Hashable α] {l : List α} : (ofList l : Raw α).WF := ⟨HashMap.Raw.WF.unitOfList⟩ -theorem WF.union [BEq α] [Hashable α] {m₁ m₂ : Raw α} (h₁ : m₁.WF) (h₂ : m₂.WF) : (m₁.union m₂).WF := +theorem WF.union [BEq α] [Hashable α] {m₁ m₂ : Raw α} (h₁ : m₁.WF) (h₂ : m₂.WF) : (m₁ ∪ m₂).WF := ⟨HashMap.Raw.WF.union h₁.out h₂.out⟩ -theorem WF.inter [BEq α] [Hashable α] {m₁ m₂ : Raw α} (h₁ : m₁.WF) (h₂ : m₂.WF) : (m₁.inter m₂).WF := +theorem WF.inter [BEq α] [Hashable α] {m₁ m₂ : Raw α} (h₁ : m₁.WF) (h₂ : m₂.WF) : (m₁ ∩ m₂).WF := ⟨HashMap.Raw.WF.inter h₁.out h₂.out⟩ -theorem WF.diff [BEq α] [Hashable α] {m₁ m₂ : Raw α} (h₁ : m₁.WF) (h₂ : m₂.WF) : (m₁.diff m₂).WF := +theorem WF.diff [BEq α] [Hashable α] {m₁ m₂ : Raw α} (h₁ : m₁.WF) (h₂ : m₂.WF) : (m₁ \ m₂).WF := ⟨HashMap.Raw.WF.diff h₁.out h₂.out⟩ end Raw From 1689034575137f7a182386c2e9cb1c501cb3751e Mon Sep 17 00:00:00 2001 From: Wojciech Rozowski Date: Thu, 20 Nov 2025 17:44:32 +0000 Subject: [PATCH 77/88] save what I have so far --- src/Std/Data/DHashMap/Internal/RawLemmas.lean | 6 +- src/Std/Data/DHashMap/Internal/WF.lean | 8 +- src/Std/Data/Internal/List/Associative.lean | 147 ++++++++++-------- 3 files changed, 90 insertions(+), 71 deletions(-) diff --git a/src/Std/Data/DHashMap/Internal/RawLemmas.lean b/src/Std/Data/DHashMap/Internal/RawLemmas.lean index 5528b55b041d..e1ef803f02a3 100644 --- a/src/Std/Data/DHashMap/Internal/RawLemmas.lean +++ b/src/Std/Data/DHashMap/Internal/RawLemmas.lean @@ -3326,19 +3326,19 @@ variable {m₁ m₂ : Raw₀ α β} theorem contains_diff [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} : (m₁.diff m₂).contains k = (m₁.contains k && !m₂.contains k) := by - simp_to_model [diff, contains] using List.containsKey_diff_eq_false + simp_to_model [diff, contains] using List.containsKey_filter_contains_map_fst_eq_false theorem contains_diff_iff [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} : (m₁.diff m₂).contains k ↔ m₁.contains k ∧ ¬m₂.contains k := by - simp_to_model [diff, contains] using List.containsKey_diff_eq_false_iff + simp_to_model [diff, contains] using List.containsKey_filter_contains_map_fst_eq_false_iff theorem contains_diff_eq_false_of_contains_eq_false_left [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} (h : m₁.contains k = false) : (m₁.diff m₂).contains k = false := by revert h - simp_to_model [diff, contains] using List.containsKey_diff_eq_false_of_containsKey_eq_false_left + simp_to_model [diff, contains] using List.containsKey_filter_contains_map_fst_eq_false_eq_false_left theorem contains_diff_eq_false_of_contains_right [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} diff --git a/src/Std/Data/DHashMap/Internal/WF.lean b/src/Std/Data/DHashMap/Internal/WF.lean index 465f433d8d53..3d96c1c07d1f 100644 --- a/src/Std/Data/DHashMap/Internal/WF.lean +++ b/src/Std/Data/DHashMap/Internal/WF.lean @@ -1356,7 +1356,7 @@ theorem eraseMany_eq_eraseListₘ_toListModel [BEq α] [Hashable α] (m m₂ : R theorem toListModel_diffₘ [BEq α] [Hashable α] [EquivBEq α] [LawfulHashable α] {m₁ m₂ : Raw₀ α β} (h₁ : Raw.WFImp m₁.1) (h₂ : Raw.WFImp m₂.1) : Perm (toListModel (diffₘ m₁ m₂).1.buckets) - (List.filter (fun k => !containsKey k.fst (toListModel m₂.1.buckets)) (toListModel m₁.1.buckets)) := by + (List.filter (fun k => !List.contains ((toListModel m₂.1.buckets).map Sigma.fst) k.fst) (toListModel m₁.1.buckets)) := by rw [diffₘ] split · apply Perm.trans @@ -1367,9 +1367,10 @@ theorem toListModel_diffₘ [BEq α] [Hashable α] [EquivBEq α] [LawfulHashable ext p congr rw [containsₘ_eq_containsKey h₂] + rw [containsKey_eq_contains_map_fst] · rw [eraseManyEntries_eq_eraseListₘ] apply Perm.trans (toListModel_eraseListₘ h₁) - · apply eraseList_perm_filter_containsKey + · apply eraseList_perm_filter_contains_eq_false · exact h₁.distinct theorem diff_eq_diffₘ [BEq α] [Hashable α] (m₁ m₂ : Raw₀ α β) : @@ -1384,7 +1385,7 @@ theorem diff_eq_diffₘ [BEq α] [Hashable α] (m₁ m₂ : Raw₀ α β) : theorem toListModel_diff [BEq α] [Hashable α] [EquivBEq α] [LawfulHashable α] {m₁ m₂ : Raw₀ α β} (h₁ : Raw.WFImp m₁.1) (h₂ : Raw.WFImp m₂.1) : Perm (toListModel (m₁.diff m₂).1.buckets) - (List.filter (fun k => !containsKey k.fst (toListModel m₂.1.buckets)) (toListModel m₁.1.buckets)) := by + (List.filter (fun k => !List.contains ((toListModel m₂.1.buckets).map Sigma.fst) k.fst) (toListModel m₁.1.buckets)) := by rw [diff_eq_diffₘ] exact toListModel_diffₘ h₁ h₂ @@ -1396,7 +1397,6 @@ theorem wf_diff₀ [BEq α] [Hashable α] [EquivBEq α] [LawfulHashable α] · apply Raw.WF.filter₀ h'₁ · exact wf_eraseMany₀ ‹_› - /-! # `insertManyIfNew` -/ theorem wfImp_insertManyIfNew [BEq α] [Hashable α] [EquivBEq α] [LawfulHashable α] {ρ : Type w} diff --git a/src/Std/Data/Internal/List/Associative.lean b/src/Std/Data/Internal/List/Associative.lean index 388cdd6570a9..b173b4c88421 100644 --- a/src/Std/Data/Internal/List/Associative.lean +++ b/src/Std/Data/Internal/List/Associative.lean @@ -5349,16 +5349,16 @@ theorem getKey?_filter_containsKey_of_containsKey_right [BEq α] [EquivBEq α] · exact dl₁ /-- Internal implementation detail of the hash map -/ -def eraseList [BEq α] (l : List ((a : α) × β a)) (toErase : List α): List ((a : α) × β a) := +def eraseList [BEq α] (l : List ((a : α) × β a)) (toErase : List α) : List ((a : α) × β a) := match toErase with | .nil => l | .cons k toErase => eraseList (eraseKey k l) toErase -theorem eraseList_perm_filter_containsKey [BEq α] [EquivBEq α] (l toErase : List ((a : α) × β a)) (hl : DistinctKeys l): - List.Perm (eraseList l (toErase.map (·.1))) (l.filter (fun p => !containsKey p.fst toErase)) := by +theorem eraseList_perm_filter_contains_eq_false [BEq α] [EquivBEq α] (l toErase : List ((a : α) × β a)) (hl : DistinctKeys l): + List.Perm (eraseList l (toErase.map Sigma.fst)) (l.filter (fun p => !List.contains (toErase.map Sigma.fst) p.fst)) := by induction toErase generalizing l with | nil => - simp [eraseList, containsKey] + simp only [List.map_nil, eraseList, List.contains, List.elem, Bool.not_false] clear hl induction l case nil => simp @@ -5371,9 +5371,9 @@ theorem eraseList_perm_filter_containsKey [BEq α] [EquivBEq α] (l toErase : Li · apply ih · apply DistinctKeys.eraseKey hl · apply getEntry?_ext - · apply @DistinctKeys.filter α β _ (eraseKey kv.fst l) (fun p _ => !containsKey p tl) + · apply @DistinctKeys.filter α β _ (eraseKey kv.fst l) (fun p _ => !List.contains (tl.map Sigma.fst) p) · apply DistinctKeys.eraseKey hl - · apply @DistinctKeys.filter α β _ l (fun p _ => !containsKey p (kv :: tl)) hl + · apply @DistinctKeys.filter α β _ l (fun p _ => !List.contains ((kv :: tl).map Sigma.fst) p) hl · intro a rw [getEntry?_filter, getEntry?_filter] rw [getEntry?_eraseKey] @@ -5384,7 +5384,7 @@ theorem eraseList_perm_filter_containsKey [BEq α] [EquivBEq α] (l toErase : Li lhs ext p congr - rw [containsKey_cons] + rw [List.contains_cons] simp only [Option.filter_none, Bool.not_or] generalize heq2 : getEntry? a l = x cases x @@ -5393,7 +5393,7 @@ theorem eraseList_perm_filter_containsKey [BEq α] [EquivBEq α] (l toErase : Li rw [getEntry?_eq_some_iff] at heq2 · simp [Option.filter] intro hyp - rw [PartialEquivBEq.trans heq heq2.1] at hyp + rw [BEq.symm <| BEq.trans heq heq2.1] at hyp contradiction · exact hl · rename_i heq @@ -5403,13 +5403,13 @@ theorem eraseList_perm_filter_containsKey [BEq α] [EquivBEq α] (l toErase : Li · simp · rename_i val simp only [Option.filter, Bool.not_eq_eq_eq_not, Bool.not_true] - rw [containsKey_cons] + rw [List.contains_cons] rw [getEntry?_eq_some_iff] at heq2 - · suffices (kv.fst == val.fst) = false by simp [this] + · suffices (val.fst == kv.fst) = false by simp [this] apply Classical.byContradiction intro hyp simp only [Bool.not_eq_false] at hyp - rw [BEq.trans hyp (BEq.symm heq2.1)] at heq + rw [BEq.symm <| BEq.trans heq2.1 hyp] at heq contradiction · exact hl · exact hl @@ -5556,7 +5556,7 @@ theorem getEntry?_map [BEq α] [EquivBEq α] exact BEq.rfl theorem containsKey_filter_containsKey [BEq α] [EquivBEq α] {l₁ l₂ : List ((a : α) × β a)} {hl₁ : DistinctKeys l₁} {k : α} : - (containsKey k (List.filter (fun p => containsKey p.fst (l₂)) l₁)) = + (containsKey k (List.filter (fun p => containsKey p.fst l₂) l₁)) = (containsKey k l₁ && containsKey k l₂) := by rw [containsKey_eq_isSome_getEntry?, containsKey_eq_isSome_getEntry?] rw [getEntry?_filter] @@ -5577,9 +5577,9 @@ theorem containsKey_filter_containsKey_iff [BEq α] [EquivBEq α] {l₁ l₂ : L . simp . exact hl₁ -theorem containsKey_diff_eq_false [BEq α] [EquivBEq α] {l₁ l₂ : List ((a : α) × β a)} {hl₁ : DistinctKeys l₁} {k : α} : - (containsKey k (List.filter (fun p => !containsKey p.fst (l₂)) l₁)) = - (containsKey k l₁ && !containsKey k l₂) := by +theorem containsKey_filter_contains_eq_false [BEq α] [EquivBEq α] {l₁ : List ((a : α) × β a)} {l₂ : List α} {hl₁ : DistinctKeys l₁} {k : α} : + (containsKey k (List.filter (fun p => !List.contains l₂ p.fst ) l₁)) = + (containsKey k l₁ && !List.contains l₂ k) := by rw [containsKey_eq_isSome_getEntry?, containsKey_eq_isSome_getEntry?] rw [getEntry?_filter] generalize heq : (getEntry? k l₁) = x @@ -5588,14 +5588,21 @@ theorem containsKey_diff_eq_false [BEq α] [EquivBEq α] {l₁ l₂ : List ((a : simp case some kv => simp only [Option.isSome_filter, Option.any_some, Option.isSome_some, Bool.true_and] - rw [containsKey_congr] + rw [List.contains_congr] apply List.beq_of_getEntry?_eq_some heq . exact hl₁ -theorem containsKey_diff_eq_false_iff [BEq α] [EquivBEq α] {l₁ l₂ : List ((a : α) × β a)} {hl₁ : DistinctKeys l₁} {k : α} : - (containsKey k (List.filter (fun p => !containsKey p.fst l₂) l₁)) ↔ +theorem containsKey_filter_contains_map_fst_eq_false [BEq α] [EquivBEq α] {l₁ l₂ : List ((a : α) × β a)} {hl₁ : DistinctKeys l₁} {k : α} : + (containsKey k (List.filter (fun p => !List.contains (l₂.map Sigma.fst) p.fst ) l₁)) = + (containsKey k l₁ && !containsKey k l₂) := by + rw [containsKey_filter_contains_eq_false, ← containsKey_eq_contains_map_fst] + · exact hl₁ + +-- changed +theorem containsKey_filter_contains_map_fst_eq_false_iff [BEq α] [EquivBEq α] {l₁ l₂ : List ((a : α) × β a)} {hl₁ : DistinctKeys l₁} {k : α} : + (containsKey k (List.filter (fun p => !List.contains (l₂.map Sigma.fst) p.fst) l₁)) ↔ (containsKey k l₁ ∧ ¬ containsKey k l₂) := by - rw [containsKey_diff_eq_false] + rw [containsKey_filter_contains_map_fst_eq_false] . simp . exact hl₁ @@ -5603,12 +5610,13 @@ theorem getValueCast_diff [BEq α] [LawfulBEq α] {l₁ l₂ : List ((a : α) × β a)} {k : α} (dl₁ : DistinctKeys l₁) {h₁ h₂} : getValueCast k (List.filter (fun p => !containsKey p.fst l₂) l₁) h₁ = getValueCast k l₁ h₂ := by - suffices some (getValueCast k (List.filter (fun p => !containsKey p.fst l₂) l₁) h₁) = some (getValueCast k l₁ h₂) by injections - simp only [← getValueCast?_eq_some_getValueCast] - rw [getValueCast?_diff dl₁] - rw [containsKey_diff_eq_false_iff] at h₁ - simp [h₁.2] - exact dl₁ + sorry + -- suffices some (getValueCast k (List.filter (fun p => !containsKey p.fst l₂) l₁) h₁) = some (getValueCast k l₁ h₂) by injections + -- simp only [← getValueCast?_eq_some_getValueCast] + -- rw [getValueCast?_diff dl₁] + -- rw [containsKey_diff_eq_false_iff] at h₁ + -- simp [h₁.2] + -- exact dl₁ theorem getValueCast_filter_containsKey [BEq α] [LawfulBEq α] {l₁ l₂ : List ((a : α) × β a)} {k : α} @@ -5652,13 +5660,19 @@ theorem containsKey_filter_containsKey_eq_false_of_containsKey_eq_false_left [BE . simp [h] . exact hl₁ -theorem containsKey_diff_eq_false_of_containsKey_eq_false_left [BEq α] [EquivBEq α] {l₁ l₂ : List ((a : α) × β a)} {hl₁ : DistinctKeys l₁} {k : α} : - containsKey k l₁ = false → containsKey k (List.filter (fun p => !containsKey p.fst l₂) l₁) = false := by +theorem containsKey_filter_contains_eq_false_eq_false_left [BEq α] [EquivBEq α] {l₁: List ((a : α) × β a)} {l₂ : List α} {hl₁ : DistinctKeys l₁} {k : α} : + containsKey k l₁ = false → containsKey k (List.filter (fun p => !List.contains l₂ p.fst) l₁) = false := by intro h - rw [containsKey_diff_eq_false] + rw [containsKey_filter_contains_eq_false] . simp [h] . exact hl₁ +theorem containsKey_filter_contains_map_fst_eq_false_eq_false_left [BEq α] [EquivBEq α] {l₁ l₂ : List ((a : α) × β a)} {hl₁ : DistinctKeys l₁} {k : α} : + containsKey k l₁ = false → containsKey k (List.filter (fun p => !List.contains (l₂.map Sigma.fst) p.fst) l₁) = false := by + intro h + rw [List.containsKey_filter_contains_map_fst_eq_false, h, Bool.false_and] + exact hl₁ + theorem containsKey_filter_containsKey_eq_false_of_containsKey_eq_false_right [BEq α] [EquivBEq α] {l₁ l₂ : List ((a : α) × β a)} {hl₁ : DistinctKeys l₁} {k : α} : containsKey k l₂ = false → containsKey k (List.filter (fun p => containsKey p.fst l₂) l₁) = false := by intro h @@ -5668,10 +5682,11 @@ theorem containsKey_filter_containsKey_eq_false_of_containsKey_eq_false_right [B theorem containsKey_diff_of_containsKey_eq_false_right [BEq α] [EquivBEq α] {l₁ l₂ : List ((a : α) × β a)} {hl₁ : DistinctKeys l₁} {k : α} : containsKey k l₂ = true → containsKey k (List.filter (fun p => !containsKey p.fst l₂) l₁) = false := by - intro h - rw [containsKey_diff_eq_false] - . simp [h] - . exact hl₁ + sorry + -- intro h + -- rw [containsKey_diff_eq_false] + -- . simp [h] + -- . exact hl₁ theorem getValueCastD_diff [BEq α] [LawfulBEq α] {l₁ l₂ : List ((a : α) × β a)} {k : α} {fallback : β k} @@ -5694,11 +5709,12 @@ theorem getValueCastD_diff_of_containsKey_eq_false_left [BEq α] [LawfulBEq α] {l₁ l₂ : List ((a : α) × β a)} {k : α} {fallback} (dl₁ : DistinctKeys l₁) : containsKey k l₁ = false → getValueCastD k (List.filter (fun p => !containsKey p.fst l₂) l₁) fallback = fallback := by - intro h - apply getValueCastD_eq_fallback - apply containsKey_diff_eq_false_of_containsKey_eq_false_left - · exact dl₁ - · exact h + sorry + -- intro h + -- apply getValueCastD_eq_fallback + -- apply containsKey_diff_eq_false_of_containsKey_eq_false_left + -- · exact dl₁ + -- · exact h theorem getKey?_diff [BEq α] [EquivBEq α] {l₁ l₂ : List ((a : α) × β a)} {k : α} @@ -5768,14 +5784,15 @@ theorem getKey_diff [BEq α] [EquivBEq α] {l₁ l₂ : List ((a : α) × β a)} {k : α} (dl₁ : DistinctKeys l₁) {h₁ h₂} : getKey k (List.filter (fun p => !containsKey p.fst l₂) l₁) h₁ = getKey k l₁ h₂ := by - suffices some (getKey k (List.filter (fun p => !containsKey p.fst l₂) l₁) h₁) = some (getKey k l₁ h₂) by - injections - simp only [← getKey?_eq_some_getKey] - apply getKey?_diff_of_containsKey_eq_false_right dl₁ - · rw [containsKey_diff_eq_false_iff] at h₁ - · simp at h₁ - exact h₁.2 - · exact dl₁ + sorry + -- suffices some (getKey k (List.filter (fun p => !containsKey p.fst l₂) l₁) h₁) = some (getKey k l₁ h₂) by + -- injections + -- simp only [← getKey?_eq_some_getKey] + -- apply getKey?_diff_of_containsKey_eq_false_right dl₁ + -- · rw [containsKey_diff_eq_false_iff] at h₁ + -- · simp at h₁ + -- exact h₁.2 + -- · exact dl₁ theorem getKeyD_diff [BEq α] [EquivBEq α] {l₁ l₂ : List ((a : α) × β a)} {k : α} {fallback : α} @@ -5818,11 +5835,12 @@ theorem getKeyD_diff_of_containsKey_eq_false_left [BEq α] [EquivBEq α] {l₁ l₂ : List ((a : α) × β a)} {k : α} {fallback : α} (dl₁ : DistinctKeys l₁) : containsKey k l₁ = false → getKeyD k (List.filter (fun p => !containsKey p.fst l₂) l₁) fallback = fallback := by - intro h - apply getKeyD_eq_fallback - apply containsKey_diff_eq_false_of_containsKey_eq_false_left - · exact dl₁ - · exact h + sorry + -- intro h + -- apply getKeyD_eq_fallback + -- apply containsKey_diff_eq_false_of_containsKey_eq_false_left + -- · exact dl₁ + -- · exact h theorem length_diff_eq_length_left [BEq α] [EquivBEq α] {l₁ l₂ : List ((a : α) × β a)} (hl₁ : DistinctKeys l₁) @@ -5972,15 +5990,15 @@ theorem getValue?_diff_of_containsKey_right {β : Type v} [BEq α] [EquivBEq α] theorem getValue_diff {β : Type v} [BEq α] [EquivBEq α] {l₁ l₂ : List ((_ : α) × β)} {k : α} (dl₁ : DistinctKeys l₁) {h₁ h₂} : - getValue k (List.filter (fun p => !containsKey p.fst l₂) l₁) h₁ = getValue k l₁ h₂ := by - suffices some (getValue k (List.filter (fun p => !containsKey p.fst l₂) l₁) h₁) = some (getValue k l₁ h₂) by - injections - simp only [← getValue?_eq_some_getValue] - apply getValue?_diff_of_containsKey_eq_false_right dl₁ - · rw [containsKey_diff_eq_false_iff] at h₁ - · simp at h₁ - exact h₁.2 - · exact dl₁ + getValue k (List.filter (fun p => !containsKey p.fst l₂) l₁) h₁ = getValue k l₁ h₂ := by sorry + -- suffices some (getValue k (List.filter (fun p => !containsKey p.fst l₂) l₁) h₁) = some (getValue k l₁ h₂) by + -- injections + -- simp only [← getValue?_eq_some_getValue] + -- apply getValue?_diff_of_containsKey_eq_false_right dl₁ + -- · rw [containsKey_diff_eq_false_iff] at h₁ + -- · simp at h₁ + -- exact h₁.2 + -- · exact dl₁ theorem getValueD_diff {β : Type v} [BEq α] [EquivBEq α] {l₁ l₂ : List ((_ : α) × β)} {k : α} {fallback : β} @@ -6023,11 +6041,12 @@ theorem getValueD_diff_of_containsKey_eq_false_left {β : Type v} [BEq α] [Equi {l₁ l₂ : List ((_ : α) × β)} {k : α} {fallback : β} (dl₁ : DistinctKeys l₁) : containsKey k l₁ = false → getValueD k (List.filter (fun p => !containsKey p.fst l₂) l₁) fallback = fallback := by - intro h - apply getValueD_eq_fallback - apply containsKey_diff_eq_false_of_containsKey_eq_false_left - · exact dl₁ - · exact h + sorry + -- intro h + -- apply getValueD_eq_fallback + -- apply containsKey_diff_eq_false_of_containsKey_eq_false_left + -- · exact dl₁ + -- · exact h theorem getValueCastD_diff_of_containsKey_right [BEq α] [LawfulBEq α] {l₁ l₂ : List ((a : α) × β a)} {k : α} {fallback} From 0e2c1d82d3ab453840f75f225df26bfac1b08c55 Mon Sep 17 00:00:00 2001 From: Wojciech Rozowski Date: Fri, 21 Nov 2025 13:30:39 +0000 Subject: [PATCH 78/88] refactor progress --- src/Std/Data/DHashMap/Internal/RawLemmas.lean | 30 +- src/Std/Data/HashMap/Raw.lean | 1 - src/Std/Data/HashSet/Raw.lean | 1 - src/Std/Data/Internal/List/Associative.lean | 379 ++++++++++++------ 4 files changed, 281 insertions(+), 130 deletions(-) diff --git a/src/Std/Data/DHashMap/Internal/RawLemmas.lean b/src/Std/Data/DHashMap/Internal/RawLemmas.lean index 59639dc118b8..8013f9cf9cb9 100644 --- a/src/Std/Data/DHashMap/Internal/RawLemmas.lean +++ b/src/Std/Data/DHashMap/Internal/RawLemmas.lean @@ -3365,86 +3365,86 @@ theorem contains_diff_eq_false_of_contains_right [EquivBEq α] [LawfulHashable (h : m₂.contains k) : (m₁.diff m₂).contains k = false := by revert h - simp_to_model [diff, contains] using List.containsKey_diff_of_containsKey_eq_false_right + simp_to_model [diff, contains] using List.containsKey_filter_contains_map_fst_eq_false_of_contains_map_fst_right /- get? -/ theorem get?_diff [LawfulBEq α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} : (m₁.diff m₂).get? k = if m₂.contains k then none else m₁.get? k := by - simp_to_model [diff, get?, contains] using getValueCast?_diff + simp_to_model [diff, get?, contains] using List.getValueCast?_filter_contains_map_fst_eq_false theorem get?_diff_of_contains_eq_false_right [LawfulBEq α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} (h : m₂.contains k = false) : (m₁.diff m₂).get? k = m₁.get? k := by revert h - simp_to_model [diff, contains, get?] using List.getValueCast?_diff_of_containsKey_eq_false_right + simp_to_model [diff, contains, get?] using List.getValueCast?_filter_contains_map_fst_eq_false_of_containsKey_eq_false_right theorem get?_diff_of_contains_eq_false_left [LawfulBEq α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} (h : m₁.contains k = false) : (m₁.diff m₂).get? k = none := by revert h - simp_to_model [diff, contains, get?] using List.getValueCast?_diff_of_containsKey_eq_false_left + simp_to_model [diff, contains, get?] using List.getValueCast?_filter_contains_map_fst_eq_false_of_containsKey_eq_false_left theorem get?_diff_of_contains_right [LawfulBEq α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} (h : m₂.contains k) : (m₁.diff m₂).get? k = none := by revert h - simp_to_model [diff, get?, contains] using List.getValueCast?_diff_of_containsKey_right + simp_to_model [diff, get?, contains] using List.getValueCast?_filter_contains_map_fst_eq_false_of_containsKey_right /- get -/ theorem get_diff [LawfulBEq α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} {h_contains : (m₁.diff m₂).contains k} : (m₁.diff m₂).get k h_contains = m₁.get k ((contains_diff_iff h₁ h₂).1 h_contains).1 := by - simp_to_model [diff, get, contains] using List.getValueCast_diff + simp_to_model [diff, get, contains] using List.getValueCast_filter_contains_map_fst_eq_false /- getD -/ theorem getD_diff [LawfulBEq α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} {fallback : β k} : (m₁.diff m₂).getD k fallback = if m₂.contains k then fallback else m₁.getD k fallback := by - simp_to_model [diff, getD, contains] using List.getValueCastD_diff + simp_to_model [diff, getD, contains] using List.getValueCastD_filter_contains_map_fst_eq_false theorem getD_diff_of_contains_eq_false_right [LawfulBEq α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} {fallback : β k} (h : m₂.contains k = false) : (m₁.diff m₂).getD k fallback = m₁.getD k fallback := by revert h - simp_to_model [diff, contains, getD] using getValueCastD_diff_of_containsKey_eq_false_right + simp_to_model [diff, contains, getD] using List.getValueCastD_filter_contains_map_fst_eq_false_of_containsKey_eq_false_right theorem getD_diff_of_contains_right [LawfulBEq α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} {fallback : β k} (h : m₂.contains k) : (m₁.diff m₂).getD k fallback = fallback := by revert h - simp_to_model [diff, getD, contains] using List.getValueCastD_diff_of_containsKey_right + simp_to_model [diff, getD, contains] using List.getValueCastD_filter_contains_map_fst_eq_false_of_contains_map_fst_right theorem getD_diff_of_contains_eq_false_left [LawfulBEq α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} {fallback : β k} (h : m₁.contains k = false) : (m₁.diff m₂).getD k fallback = fallback := by revert h - simp_to_model [diff, getD, contains] using List.getValueCastD_diff_of_containsKey_eq_false_left + simp_to_model [diff, getD, contains] using List.getValueCastD_filter_contains_map_fst_eq_false_of_containsKey_eq_false_left /- get! -/ theorem get!_diff [LawfulBEq α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} [Inhabited (β k)] : (m₁.diff m₂).get! k = if m₂.contains k then default else m₁.get! k := by - simp_to_model [diff, get!, contains] using List.getValueCastD_diff + simp_to_model [diff, get!, contains] using List.getValueCastD_filter_contains_map_fst_eq_false theorem get!_diff_of_contains_eq_false_right [LawfulBEq α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} [Inhabited (β k)] (h : m₂.contains k = false) : (m₁.diff m₂).get! k = m₁.get! k := by revert h - simp_to_model [diff, contains, get!] using getValueCastD_diff_of_containsKey_eq_false_right + simp_to_model [diff, contains, get!] using List.getValueCastD_filter_contains_map_fst_eq_false_of_containsKey_eq_false_right theorem get!_diff_of_contains_right [LawfulBEq α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} [Inhabited (β k)] (h : m₂.contains k) : (m₁.diff m₂).get! k = default := by revert h - simp_to_model [diff, get!, contains] using List.getValueCastD_diff_of_containsKey_right + simp_to_model [diff, get!, contains] using List.getValueCastD_filter_contains_map_fst_eq_false_of_contains_map_fst_right theorem get!_diff_of_contains_eq_false_left [LawfulBEq α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} [Inhabited (β k)] (h : m₁.contains k = false) : (m₁.diff m₂).get! k = default := by revert h - simp_to_model [diff, get!, contains] using List.getValueCastD_diff_of_containsKey_eq_false_left + simp_to_model [diff, get!, contains] using List.getValueCastD_filter_contains_map_fst_eq_false_of_containsKey_eq_false_left /- getKey? -/ theorem getKey?_diff [EquivBEq α] [LawfulHashable α] @@ -3609,7 +3609,7 @@ theorem getD_diff_of_contains_right [EquivBEq α] [LawfulHashable α] (h₁ : m {k : α} {fallback : β} (h : m₂.contains k) : Const.getD (m₁.diff m₂) k fallback = fallback := by revert h - simp_to_model [diff, Const.getD, contains] using List.getValueD_diff_of_containsKey_right + simp_to_model [diff, Const.getD, contains] using List.getValueD_filter_contains_map_fst_eq_false_of_contains_map_fst_right theorem getD_diff_of_contains_eq_false_left [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} {fallback : β} (h : m₁.contains k = false) : diff --git a/src/Std/Data/HashMap/Raw.lean b/src/Std/Data/HashMap/Raw.lean index 1ac5b15ae221..930fd402aadf 100644 --- a/src/Std/Data/HashMap/Raw.lean +++ b/src/Std/Data/HashMap/Raw.lean @@ -350,7 +350,6 @@ theorem WF.ofList [BEq α] [Hashable α] {l : List (α × β)} : (ofList l).WF : theorem WF.unitOfList [BEq α] [Hashable α] {l : List α} : (unitOfList l).WF := ⟨DHashMap.Raw.WF.Const.unitOfList⟩ -theorem WF.union [BEq α] [Hashable α] {m₁ m₂ : Raw α β} (h₁ : m₁.WF) (h₂ : m₂.WF) : (m₁ ∪ m₂).WF := theorem WF.union [BEq α] [Hashable α] {m₁ m₂ : Raw α β} (h₁ : m₁.WF) (h₂ : m₂.WF) : (m₁ ∪ m₂).WF := ⟨DHashMap.Raw.WF.union h₁.out h₂.out⟩ diff --git a/src/Std/Data/HashSet/Raw.lean b/src/Std/Data/HashSet/Raw.lean index f972773573cd..d98c1eeb3fae 100644 --- a/src/Std/Data/HashSet/Raw.lean +++ b/src/Std/Data/HashSet/Raw.lean @@ -336,7 +336,6 @@ theorem WF.ofList [BEq α] [Hashable α] {l : List α} : (ofList l : Raw α).WF := ⟨HashMap.Raw.WF.unitOfList⟩ -theorem WF.union [BEq α] [Hashable α] {m₁ m₂ : Raw α} (h₁ : m₁.WF) (h₂ : m₂.WF) : (m₁ ∪ m₂).WF := theorem WF.union [BEq α] [Hashable α] {m₁ m₂ : Raw α} (h₁ : m₁.WF) (h₂ : m₂.WF) : (m₁ ∪ m₂).WF := ⟨HashMap.Raw.WF.union h₁.out h₂.out⟩ diff --git a/src/Std/Data/Internal/List/Associative.lean b/src/Std/Data/Internal/List/Associative.lean index f9b88f728f50..54f72c50e095 100644 --- a/src/Std/Data/Internal/List/Associative.lean +++ b/src/Std/Data/Internal/List/Associative.lean @@ -5233,68 +5233,127 @@ theorem getValueCast?_filter_containsKey [BEq α] [LawfulBEq α] case h_2 => rfl exact dl₁ -theorem getValueCast?_diff [BEq α] [LawfulBEq α] - {l₁ l₂ : List ((a : α) × β a)} {k : α} +/- `get?_diff` -/ +theorem getValueCast?_filter_contains_eq_false [BEq α] [LawfulBEq α] + {l₁ : List ((a : α) × β a)} {l₂ : List α} {k : α} (dl₁ : DistinctKeys l₁) : - getValueCast? k (List.filter (fun p => !containsKey p.fst l₂) l₁) = - if containsKey k l₂ = true then none else getValueCast? k l₁ := by + getValueCast? k (List.filter (fun p => !List.contains l₂ p.fst) l₁) = + if List.contains l₂ k = true then none else getValueCast? k l₁ := by simp only [getValueCast?_eq_getEntry?] split case isTrue h => - suffices (getEntry? k (List.filter (fun p => !containsKey p.fst l₂) l₁)) = none by simp [this] + suffices (getEntry? k (List.filter (fun p => !List.contains l₂ p.fst) l₁)) = none by + simp [Option.dmap] + split + case h_1 => rfl + case h_2 heq _ => + rw [heq] at this + contradiction rw [getEntry?_filter] simp [Option.filter] split case h_1 _ val heq => - rw [@containsKey_congr _ _ _ _ l₂ _ _ (PartialEquivBEq.symm <| @beq_of_getEntry?_eq_some α β _ l₁ k val heq)] at h - simp [h] + rw [@List.contains_congr _ _ _ l₂ _ _ (PartialEquivBEq.symm <| @beq_of_getEntry?_eq_some α β _ l₁ k val heq)] at h + split + case isTrue => rfl + case isFalse h₁ => + rw [List.contains_eq_mem] at h + rw [← decide_eq_false_iff_not] at h₁ + rw [h₁] at h + contradiction case h_2 => rfl exact dl₁ case isFalse h => apply Option.dmap_congr - . simp - . rw [getEntry?_filter] - . simp [Option.filter] + · simp + · rw [getEntry?_filter] + · simp [Option.filter] split case h_1 _ val heq => - rw [@containsKey_congr _ _ _ _ l₂ _ _ (PartialEquivBEq.symm <| @beq_of_getEntry?_eq_some α β _ l₁ k val heq)] at h - simp [h, heq] + rw [@List.contains_congr _ _ _ l₂ _ _ (PartialEquivBEq.symm <| @beq_of_getEntry?_eq_some α β _ l₁ k val heq)] at h + rw [List.contains_eq_mem] at h + rw [decide_eq_true_eq] at h + rw [heq] + split + case isTrue => contradiction + case isFalse => rfl case h_2 _ heq => simp [heq] - . exact dl₁ + · exact dl₁ -theorem getValueCast?_diff_of_containsKey_eq_false_right [BEq α] [LawfulBEq α] +theorem getValueCast?_filter_contains_map_fst_eq_false [BEq α] [LawfulBEq α] {l₁ l₂ : List ((a : α) × β a)} {k : α} (dl₁ : DistinctKeys l₁) : - containsKey k l₂ = false → getValueCast? k (List.filter (fun p => !containsKey p.fst l₂) l₁) = + getValueCast? k (List.filter (fun p => !List.contains (l₂.map Sigma.fst) p.fst) l₁) = + if containsKey k l₂ = true then none else getValueCast? k l₁ := by + rw [getValueCast?_filter_contains_eq_false, ← containsKey_eq_contains_map_fst] + exact dl₁ + +/- `get?_diff_of_contains_eq_false_right` -/ + +theorem getValueCast?_filter_contains_eq_false_of_contains_eq_false_right [BEq α] [LawfulBEq α] + {l₁ : List ((a : α) × β a)} {l₂ : List α} {k : α} + (dl₁ : DistinctKeys l₁) : + List.contains l₂ k = false → getValueCast? k (List.filter (fun p => !List.contains l₂ p.fst) l₁) = getValueCast? k l₁ := by intro h₁ - rw [getValueCast?_diff, h₁] + rw [getValueCast?_filter_contains_eq_false, h₁] all_goals simp [dl₁] -theorem getValueCast?_diff_of_containsKey_eq_false_left [BEq α] [LawfulBEq α] +theorem getValueCast?_filter_contains_map_fst_eq_false_of_containsKey_eq_false_right [BEq α] [LawfulBEq α] {l₁ l₂ : List ((a : α) × β a)} {k : α} (dl₁ : DistinctKeys l₁) : - containsKey k l₁ = false → getValueCast? k (List.filter (fun p => !containsKey p.fst l₂) l₁) = + containsKey k l₂ = false → getValueCast? k (List.filter (fun p => !List.contains (l₂.map Sigma.fst) p.fst) l₁) = + getValueCast? k l₁ := by + intro h₁ + rw [containsKey_eq_contains_map_fst] at h₁ + rw [getValueCast?_filter_contains_eq_false_of_contains_eq_false_right dl₁] + · exact h₁ + +/- `get?_diff_of_contains_eq_false_left` -/ +theorem getValueCast?_filter_contains_eq_false_of_contains_eq_false_left [BEq α] [LawfulBEq α] + {l₁ : List ((a : α) × β a)} {l₂ : List α} {k : α} + (dl₁ : DistinctKeys l₁) : + containsKey k l₁ = false → getValueCast? k (List.filter (fun p => !List.contains l₂ p.fst) l₁) = none := by intro h₁ - rw [getValueCast?_diff] + rw [getValueCast?_filter_contains_eq_false] · simp [getValueCast?_eq_none h₁] · exact dl₁ -theorem getValueCast?_diff_of_containsKey_right [BEq α] [LawfulBEq α] +theorem getValueCast?_filter_contains_map_fst_eq_false_of_containsKey_eq_false_left [BEq α] [LawfulBEq α] {l₁ l₂ : List ((a : α) × β a)} {k : α} (dl₁ : DistinctKeys l₁) : - containsKey k l₂ → getValueCast? k (List.filter (fun p => !containsKey p.fst l₂) l₁) = + containsKey k l₁ = false → getValueCast? k (List.filter (fun p => !List.contains (l₂.map Sigma.fst) p.fst) l₁) = + none := by + intro h₁ + rw [getValueCast?_filter_contains_eq_false_of_contains_eq_false_left dl₁] + · exact h₁ + +/- `get?_diff_of_contains_right` -/ +theorem getValueCast?_filter_contains_eq_false_of_contains_right [BEq α] [LawfulBEq α] + {l₁ : List ((a : α) × β a)} {l₂ : List α} {k : α} + (dl₁ : DistinctKeys l₁) : + List.contains l₂ k → getValueCast? k (List.filter (fun p => !List.contains l₂ p.fst) l₁) = none := by intro h₁ - rw [getValueCast?_diff] + rw [getValueCast?_filter_contains_eq_false] · simp only [ite_eq_left_iff, Bool.not_eq_true] intro h₂ rw [h₂] at h₁ contradiction · exact dl₁ +theorem getValueCast?_filter_contains_map_fst_eq_false_of_containsKey_right [BEq α] [LawfulBEq α] + {l₁ l₂ : List ((a : α) × β a)} {k : α} + (dl₁ : DistinctKeys l₁) : + containsKey k l₂ → getValueCast? k (List.filter (fun p => !List.contains (l₂.map Sigma.fst) p.fst) l₁) = + none := by + intro h₁ + rw [containsKey_eq_contains_map_fst] at h₁ + rw [getValueCast?_filter_contains_eq_false_of_contains_right dl₁] + · exact h₁ + theorem getKey?_filter_containsKey [BEq α] [EquivBEq α] {l₁ l₂ : List ((a : α) × β a)} {k : α} (dl₁ : DistinctKeys l₁) : @@ -5568,6 +5627,7 @@ theorem containsKey_filter_containsKey_iff [BEq α] [EquivBEq α] {l₁ l₂ : L . simp . exact hl₁ +/- `contains_diff` -/ theorem containsKey_filter_contains_eq_false [BEq α] [EquivBEq α] {l₁ : List ((a : α) × β a)} {l₂ : List α} {hl₁ : DistinctKeys l₁} {k : α} : (containsKey k (List.filter (fun p => !List.contains l₂ p.fst ) l₁)) = (containsKey k l₁ && !List.contains l₂ k) := by @@ -5589,25 +5649,40 @@ theorem containsKey_filter_contains_map_fst_eq_false [BEq α] [EquivBEq α] {l rw [containsKey_filter_contains_eq_false, ← containsKey_eq_contains_map_fst] · exact hl₁ --- changed +/- `contains_diff_iff` -/ +theorem containsKey_filter_contains_eq_false_iff [BEq α] [EquivBEq α] {l₁ : List ((a : α) × β a)} {l₂ : List α} {hl₁ : DistinctKeys l₁} {k : α} : + (containsKey k (List.filter (fun p => !List.contains l₂ p.fst) l₁)) ↔ + (containsKey k l₁ ∧ ¬ List.contains l₂ k) := by + rw [containsKey_filter_contains_eq_false] + . simp + . exact hl₁ + theorem containsKey_filter_contains_map_fst_eq_false_iff [BEq α] [EquivBEq α] {l₁ l₂ : List ((a : α) × β a)} {hl₁ : DistinctKeys l₁} {k : α} : (containsKey k (List.filter (fun p => !List.contains (l₂.map Sigma.fst) p.fst) l₁)) ↔ (containsKey k l₁ ∧ ¬ containsKey k l₂) := by - rw [containsKey_filter_contains_map_fst_eq_false] - . simp - . exact hl₁ + rw [containsKey_filter_contains_eq_false_iff, ← List.containsKey_eq_contains_map_fst] + · exact hl₁ + +/- `get_diff` -/ -theorem getValueCast_diff [BEq α] [LawfulBEq α] +theorem getValueCast_filter_contains_eq_false [BEq α] [LawfulBEq α] + {l₁ : List ((a : α) × β a)} {l₂ : List α} {k : α} + (dl₁ : DistinctKeys l₁) {h₁ h₂} : + getValueCast k (List.filter (fun p => !List.contains l₂ p.fst) l₁) h₁ = getValueCast k l₁ h₂ := by + suffices some (getValueCast k (List.filter (fun p => !List.contains l₂ p.fst) l₁) h₁) = some (getValueCast k l₁ h₂) by injections + simp only [← getValueCast?_eq_some_getValueCast] + rw [getValueCast?_filter_contains_eq_false] + rw [containsKey_filter_contains_eq_false_iff] at h₁ + simp only [h₁.2, Bool.false_eq_true, ↓reduceIte] + all_goals exact dl₁ + +theorem getValueCast_filter_contains_map_fst_eq_false [BEq α] [LawfulBEq α] {l₁ l₂ : List ((a : α) × β a)} {k : α} (dl₁ : DistinctKeys l₁) {h₁ h₂} : - getValueCast k (List.filter (fun p => !containsKey p.fst l₂) l₁) h₁ = getValueCast k l₁ h₂ := by - sorry - -- suffices some (getValueCast k (List.filter (fun p => !containsKey p.fst l₂) l₁) h₁) = some (getValueCast k l₁ h₂) by injections - -- simp only [← getValueCast?_eq_some_getValueCast] - -- rw [getValueCast?_diff dl₁] - -- rw [containsKey_diff_eq_false_iff] at h₁ - -- simp [h₁.2] - -- exact dl₁ + getValueCast k (List.filter (fun p => !List.contains (l₂.map Sigma.fst) p.fst) l₁) h₁ = getValueCast k l₁ h₂ := by + suffices some (getValueCast k (List.filter (fun p => !List.contains (l₂.map Sigma.fst) p.fst) l₁) h₁) = some (getValueCast k l₁ h₂) by injections + rw [getValueCast_filter_contains_eq_false] + · exact dl₁ theorem getValueCast_filter_containsKey [BEq α] [LawfulBEq α] {l₁ l₂ : List ((a : α) × β a)} {k : α} @@ -5661,8 +5736,9 @@ theorem containsKey_filter_contains_eq_false_eq_false_left [BEq α] [EquivBEq α theorem containsKey_filter_contains_map_fst_eq_false_eq_false_left [BEq α] [EquivBEq α] {l₁ l₂ : List ((a : α) × β a)} {hl₁ : DistinctKeys l₁} {k : α} : containsKey k l₁ = false → containsKey k (List.filter (fun p => !List.contains (l₂.map Sigma.fst) p.fst) l₁) = false := by intro h - rw [List.containsKey_filter_contains_map_fst_eq_false, h, Bool.false_and] - exact hl₁ + rw [containsKey_filter_contains_eq_false_eq_false_left] + · exact h + · exact hl₁ theorem containsKey_filter_containsKey_eq_false_of_containsKey_eq_false_right [BEq α] [EquivBEq α] {l₁ l₂ : List ((a : α) × β a)} {hl₁ : DistinctKeys l₁} {k : α} : containsKey k l₂ = false → containsKey k (List.filter (fun p => containsKey p.fst l₂) l₁) = false := by @@ -5697,30 +5773,50 @@ theorem congr_filter_containsKey_of_perm {l₁ l₂ l₃ l₄ : List ((a : α) · exact h₁ · apply perm_filter_containsKey_of_perm h₂ hd -theorem containsKey_diff_of_containsKey_eq_false_right [BEq α] [EquivBEq α] {l₁ l₂ : List ((a : α) × β a)} {hl₁ : DistinctKeys l₁} {k : α} : - containsKey k l₂ = true → containsKey k (List.filter (fun p => !containsKey p.fst l₂) l₁) = false := by - sorry - -- intro h - -- rw [containsKey_diff_eq_false] - -- . simp [h] - -- . exact hl₁ +/- `contains_diff_eq_false_of_contains_right` -/ +theorem containsKey_filter_contains_eq_false_of_contains_right [BEq α] [EquivBEq α] {l₁ : List ((a : α) × β a)} {l₂ : List α} {hl₁ : DistinctKeys l₁} {k : α} : + List.contains l₂ k → containsKey k (List.filter (fun p => !List.contains l₂ p.fst) l₁) = false := by + intro h + rw [containsKey_filter_contains_eq_false] + . simp [h] + . exact hl₁ -theorem getValueCastD_diff [BEq α] [LawfulBEq α] - {l₁ l₂ : List ((a : α) × β a)} {k : α} {fallback : β k} +theorem containsKey_filter_contains_map_fst_eq_false_of_contains_map_fst_right [BEq α] [EquivBEq α] {l₁ l₂ : List ((a : α) × β a)} {hl₁ : DistinctKeys l₁} {k : α} : + containsKey k l₂ → containsKey k (List.filter (fun p => !List.contains (l₂.map Sigma.fst) p.fst) l₁) = false := by + intro h + rw [containsKey_eq_contains_map_fst] at h + rw [containsKey_filter_contains_eq_false_of_contains_right] + · simp [h] + · exact hl₁ + +theorem getValueCastD_filter_contains_eq_false [BEq α] [LawfulBEq α] + {l₁ : List ((a : α) × β a)} {l₂ : List α} {k : α} {fallback : β k} (dl₁ : DistinctKeys l₁) : - getValueCastD k (List.filter (fun p => !containsKey p.fst l₂) l₁) fallback = - if containsKey k l₂ = true then fallback else getValueCastD k l₁ fallback := by + getValueCastD k (List.filter (fun p => !List.contains l₂ p.fst) l₁) fallback = + if List.contains l₂ k = true then fallback else getValueCastD k l₁ fallback := by split case isTrue h => apply getValueCastD_eq_fallback - apply containsKey_diff_of_containsKey_eq_false_right + apply containsKey_filter_contains_eq_false_of_contains_right · exact dl₁ · exact h case isFalse h => rw [getValueCastD_eq_getValueCast?, getValueCastD_eq_getValueCast?] congr 1 - rw [getValueCast?_diff dl₁] - simp [h] + rw [Bool.not_eq_true] at h + rw [getValueCast?_filter_contains_eq_false, h] + simp [Bool.false_eq_true, ↓reduceIte] + · exact dl₁ + +theorem getValueCastD_filter_contains_map_fst_eq_false [BEq α] [LawfulBEq α] + {l₁ l₂ : List ((a : α) × β a)} {k : α} {fallback : β k} + (dl₁ : DistinctKeys l₁) : + getValueCastD k (List.filter (fun p => !List.contains (l₂.map Sigma.fst) p.fst) l₁) fallback = +if containsKey k l₂ = true then fallback else getValueCastD k l₁ fallback := by + rw [getValueCastD_filter_contains_eq_false, containsKey_eq_contains_map_fst] + · exact dl₁ + + theorem getValueCastD_diff_of_containsKey_eq_false_left [BEq α] [LawfulBEq α] {l₁ l₂ : List ((a : α) × β a)} {k : α} {fallback} @@ -5816,17 +5912,18 @@ theorem getKeyD_diff [BEq α] [EquivBEq α] (dl₁ : DistinctKeys l₁) : getKeyD k (List.filter (fun p => !containsKey p.fst l₂) l₁) fallback = if containsKey k l₂ = true then fallback else getKeyD k l₁ fallback := by - split - case isTrue h => - apply getKeyD_eq_fallback - apply containsKey_diff_of_containsKey_eq_false_right - · exact dl₁ - · exact h - case isFalse h => - rw [getKeyD_eq_getKey?, getKeyD_eq_getKey?] - congr 1 - apply getKey?_diff_of_containsKey_eq_false_right dl₁ - simp only [h] + sorry + -- split + -- case isTrue h => + -- apply getKeyD_eq_fallback + -- apply containsKey_diff_of_containsKey_eq_false_right + -- · exact dl₁ + -- · exact h + -- case isFalse h => + -- rw [getKeyD_eq_getKey?, getKeyD_eq_getKey?] + -- congr 1 + -- apply getKey?_diff_of_containsKey_eq_false_right dl₁ + -- simp only [h] theorem getKeyD_diff_of_containsKey_eq_false_right [BEq α] [EquivBEq α] {l₁ l₂ : List ((a : α) × β a)} {k : α} {fallback : α} @@ -5842,11 +5939,12 @@ theorem getKeyD_diff_of_containsKey_right [BEq α] [EquivBEq α] {l₁ l₂ : List ((a : α) × β a)} {k : α} {fallback : α} (dl₁ : DistinctKeys l₁) : containsKey k l₂ = true → getKeyD k (List.filter (fun p => !containsKey p.fst l₂) l₁) fallback = fallback := by - intro h - apply getKeyD_eq_fallback - apply containsKey_diff_of_containsKey_eq_false_right - · exact dl₁ - · exact h + sorry + -- intro h + -- apply getKeyD_eq_fallback + -- apply containsKey_diff_of_containsKey_eq_false_right + -- · exact dl₁ + -- · exact h theorem getKeyD_diff_of_containsKey_eq_false_left [BEq α] [EquivBEq α] {l₁ l₂ : List ((a : α) × β a)} {k : α} {fallback : α} @@ -5859,6 +5957,24 @@ theorem getKeyD_diff_of_containsKey_eq_false_left [BEq α] [EquivBEq α] -- · exact dl₁ -- · exact h +theorem getValueCastD_filter_contains_eq_false_of_containsKey_eq_false_left {β : α → Type v} [BEq α] [LawfulBEq α] + {l₁ : List ((a : α) × β a)} {l₂ : List α} {k : α} {fallback : β k} + (dl₁ : DistinctKeys l₁) : + containsKey k l₁ = false → getValueCastD k (List.filter (fun p => !List.contains l₂ p.fst) l₁) fallback = fallback := by + intro h + apply getValueCastD_eq_fallback + apply List.containsKey_filter_contains_eq_false_eq_false_left + · exact dl₁ + · exact h + +theorem getValueCastD_filter_contains_map_fst_eq_false_of_containsKey_eq_false_left {β : α → Type v} [BEq α] [LawfulBEq α] + {l₁ l₂ : List ((a : α) × β a)} {k : α} {fallback : β k} + (dl₁ : DistinctKeys l₁) : + containsKey k l₁ = false → getValueCastD k (List.filter (fun p => !List.contains (l₂.map Sigma.fst) p.fst) l₁) fallback = fallback := by + apply getValueCastD_filter_contains_eq_false_of_containsKey_eq_false_left + · exact dl₁ + + theorem length_diff_eq_length_left [BEq α] [EquivBEq α] {l₁ l₂ : List ((a : α) × β a)} (hl₁ : DistinctKeys l₁) (w : ∀ (a : α), containsKey a l₁ → containsKey a l₂ = false) : @@ -5951,27 +6067,28 @@ theorem getValue?_diff {β : Type v} [BEq α] [EquivBEq α] (dl₁ : DistinctKeys l₁) : getValue? k (List.filter (fun p => !containsKey p.fst l₂) l₁) = if containsKey k l₂ = true then none else getValue? k l₁ := by - simp only [getValue?_eq_getEntry?] - split - case isTrue h => - simp only [Option.map_eq_none_iff, getEntry?_eq_none] - apply containsKey_diff_of_containsKey_eq_false_right - · exact dl₁ - · exact h - case isFalse h => - congr 1 - rw [getEntry?_filter] - · simp only [Option.filter] - split - case h_1 _ val heq => - simp [heq] - have := PartialEquivBEq.symm <| @beq_of_getEntry?_eq_some α (fun _ => β) _ l₁ k val heq - rw [← containsKey_congr this] - simp at h - exact h - case h_2 _ heq => - simp [heq] - · exact dl₁ + sorry + -- simp only [getValue?_eq_getEntry?] + -- split + -- case isTrue h => + -- simp only [Option.map_eq_none_iff, getEntry?_eq_none] + -- apply containsKey_diff_of_containsKey_eq_false_right + -- · exact dl₁ + -- · exact h + -- case isFalse h => + -- congr 1 + -- rw [getEntry?_filter] + -- · simp only [Option.filter] + -- split + -- case h_1 _ val heq => + -- simp [heq] + -- have := PartialEquivBEq.symm <| @beq_of_getEntry?_eq_some α (fun _ => β) _ l₁ k val heq + -- rw [← containsKey_congr this] + -- simp at h + -- exact h + -- case h_2 _ heq => + -- simp [heq] + -- · exact dl₁ theorem getValue?_diff_of_containsKey_eq_false_right {β : Type v} [BEq α] [EquivBEq α] {l₁ l₂ : List ((_ : α) × β)} {k : α} @@ -6022,17 +6139,18 @@ theorem getValueD_diff {β : Type v} [BEq α] [EquivBEq α] (dl₁ : DistinctKeys l₁) : getValueD k (List.filter (fun p => !containsKey p.fst l₂) l₁) fallback = if containsKey k l₂ = true then fallback else getValueD k l₁ fallback := by - split - case isTrue h => - apply getValueD_eq_fallback - apply containsKey_diff_of_containsKey_eq_false_right - · exact dl₁ - · exact h - case isFalse h => - rw [getValueD_eq_getValue?, getValueD_eq_getValue?] - congr 1 - apply getValue?_diff_of_containsKey_eq_false_right dl₁ - simp only [h] + sorry + -- split + -- case isTrue h => + -- apply getValueD_eq_fallback + -- apply containsKey_diff_of_containsKey_eq_false_right + -- · exact dl₁ + -- · exact h + -- case isFalse h => + -- rw [getValueD_eq_getValue?, getValueD_eq_getValue?] + -- congr 1 + -- apply getValue?_diff_of_containsKey_eq_false_right dl₁ + -- simp only [h] theorem getValueD_diff_of_containsKey_eq_false_right {β : Type v} [BEq α] [EquivBEq α] {l₁ l₂ : List ((_ : α) × β)} {k : α} {fallback : β} @@ -6044,16 +6162,27 @@ theorem getValueD_diff_of_containsKey_eq_false_right {β : Type v} [BEq α] [Equ congr 1 apply getValue?_diff_of_containsKey_eq_false_right dl₁ h -theorem getValueD_diff_of_containsKey_right {β : Type v} [BEq α] [EquivBEq α] - {l₁ l₂ : List ((_ : α) × β)} {k : α} {fallback : β} +theorem getValueD_filter_contains_eq_false_of_contains_right {β : Type v} [BEq α] [EquivBEq α] + {l₁ : List ((_ : α) × β)} {l₂ : List α} {k : α} {fallback : β} (dl₁ : DistinctKeys l₁) : - containsKey k l₂ = true → getValueD k (List.filter (fun p => !containsKey p.fst l₂) l₁) fallback = fallback := by + List.contains l₂ k → getValueD k (List.filter (fun p => !List.contains l₂ p.fst) l₁) fallback = fallback := by intro h apply getValueD_eq_fallback - apply containsKey_diff_of_containsKey_eq_false_right + apply List.containsKey_filter_contains_eq_false_of_contains_right · exact dl₁ · exact h +theorem getValueD_filter_contains_map_fst_eq_false_of_contains_map_fst_right {β : Type v} [BEq α] [EquivBEq α] + {l₁ l₂ : List ((_ : α) × β)} {k : α} {fallback : β} + (dl₁ : DistinctKeys l₁) : + containsKey k l₂ → getValueD k (List.filter (fun p => !List.contains (l₂.map Sigma.fst) p.fst) l₁) fallback = fallback := by + intro hyp + rw [containsKey_eq_contains_map_fst] at hyp + apply getValueD_filter_contains_eq_false_of_contains_right + · exact dl₁ + · exact hyp + + theorem getValueD_diff_of_containsKey_eq_false_left {β : Type v} [BEq α] [EquivBEq α] {l₁ l₂ : List ((_ : α) × β)} {k : α} {fallback : β} (dl₁ : DistinctKeys l₁) : @@ -6065,16 +6194,26 @@ theorem getValueD_diff_of_containsKey_eq_false_left {β : Type v} [BEq α] [Equi -- · exact dl₁ -- · exact h -theorem getValueCastD_diff_of_containsKey_right [BEq α] [LawfulBEq α] - {l₁ l₂ : List ((a : α) × β a)} {k : α} {fallback} +theorem getValueCastD_filter_contains_eq_false_of_contains_right {β : α → Type v} [BEq α] [LawfulBEq α] + {l₁ : List ((a : α) × β a)} {l₂ : List α} {k : α} {fallback : β k} (dl₁ : DistinctKeys l₁) : - containsKey k l₂ = true → getValueCastD k (List.filter (fun p => !containsKey p.fst l₂) l₁) fallback = fallback := by + List.contains l₂ k → getValueCastD k (List.filter (fun p => !List.contains l₂ p.fst) l₁) fallback = fallback := by intro h apply getValueCastD_eq_fallback - apply containsKey_diff_of_containsKey_eq_false_right + apply List.containsKey_filter_contains_eq_false_of_contains_right · exact dl₁ · exact h +theorem getValueCastD_filter_contains_map_fst_eq_false_of_contains_map_fst_right {β : α → Type v} [BEq α] [LawfulBEq α] + {l₁ l₂ : List ((a : α) × β a)} {k : α} {fallback : β k} + (dl₁ : DistinctKeys l₁) : + containsKey k l₂ → getValueCastD k (List.filter (fun p => !List.contains (l₂.map Sigma.fst) p.fst) l₁) fallback = fallback := by + intro hyp + rw [containsKey_eq_contains_map_fst] at hyp + apply getValueCastD_filter_contains_eq_false_of_contains_right + · exact dl₁ + · exact hyp + theorem List.getValue?_filter_containsKey {β : Type v} [BEq α] [EquivBEq α] {l₁ l₂ : List ((_ : α) × β)} {k : α} (dl₁ : DistinctKeys l₁) : @@ -6179,17 +6318,31 @@ theorem getValueCastD_filter_containsKey_of_containsKey_eq_false_right [BEq α] . exact dl₁ . exact h -theorem getValueCastD_diff_of_containsKey_eq_false_right [BEq α] [LawfulBEq α] - {l₁ l₂ : List ((a : α) × β a)} {k : α} {fallback} +theorem getValueCastD_filter_contains_eq_false_of_contains_eq_false_right [BEq α] [LawfulBEq α] + {l₁ : List ((a : α) × β a)} {l₂ : List α} {k : α} {fallback} (dl₁ : DistinctKeys l₁) : - containsKey k l₂ = false → - getValueCastD k (List.filter (fun p => !containsKey p.fst l₂) l₁) fallback = + List.contains l₂ k = false → + getValueCastD k (List.filter (fun p => !List.contains l₂ p.fst) l₁) fallback = getValueCastD k l₁ fallback := by intro h rw [getValueCastD_eq_getValueCast?, getValueCastD_eq_getValueCast?] congr 1 - rw [getValueCast?_diff dl₁] - simp [h] + rw [getValueCast?_filter_contains_eq_false dl₁] + split + · rename_i hyp + rw [hyp] at h + contradiction + · rfl + +theorem getValueCastD_filter_contains_map_fst_eq_false_of_containsKey_eq_false_right [BEq α] [LawfulBEq α] + {l₁ l₂ : List ((a : α) × β a)} {k : α} {fallback} + (dl₁ : DistinctKeys l₁) : + containsKey k l₂ = false → + getValueCastD k (List.filter (fun p => !List.contains (l₂.map Sigma.fst) p.fst) l₁) fallback = + getValueCastD k l₁ fallback := by + intro hyp + rw [List.containsKey_eq_contains_map_fst] at hyp + apply getValueCastD_filter_contains_eq_false_of_contains_eq_false_right dl₁ hyp theorem getKeyD_filter_containsKey [BEq α] [EquivBEq α] {l₁ l₂ : List ((a : α) × β a)} {k : α} {fallback : α} From bf29a9b2d0828ffb26b53328bf385e71e53982a2 Mon Sep 17 00:00:00 2001 From: Wojciech Rozowski Date: Fri, 21 Nov 2025 13:54:52 +0000 Subject: [PATCH 79/88] further progress --- src/Std/Data/DHashMap/Internal/RawLemmas.lean | 10 +- src/Std/Data/Internal/List/Associative.lean | 105 ++++++++++++------ 2 files changed, 79 insertions(+), 36 deletions(-) diff --git a/src/Std/Data/DHashMap/Internal/RawLemmas.lean b/src/Std/Data/DHashMap/Internal/RawLemmas.lean index 8013f9cf9cb9..07434835fd28 100644 --- a/src/Std/Data/DHashMap/Internal/RawLemmas.lean +++ b/src/Std/Data/DHashMap/Internal/RawLemmas.lean @@ -3450,32 +3450,32 @@ theorem get!_diff_of_contains_eq_false_left [LawfulBEq α] (h₁ : m₁.val.WF) theorem getKey?_diff [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} : (m₁.diff m₂).getKey? k = if m₂.contains k then none else m₁.getKey? k := by - simp_to_model [diff, contains, getKey?] using List.getKey?_diff + simp_to_model [diff, contains, getKey?] using List.getKey?_filter_contains_map_fst_eq_false theorem getKey?_diff_of_contains_eq_false_right [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} (h : m₂.contains k = false) : (m₁.diff m₂).getKey? k = m₁.getKey? k := by revert h - simp_to_model [contains, getKey?, diff] using List.getKey?_diff_of_containsKey_eq_false_right + simp_to_model [contains, getKey?, diff] using List.getKey?_filter_contains_map_fst_eq_false_of_containsKey_eq_false_right theorem getKey?_diff_of_contains_eq_false_left [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} (h : m₁.contains k = false) : (m₁.diff m₂).getKey? k = none := by revert h - simp_to_model [contains, getKey?, diff] using List.getKey?_diff_of_containsKey_eq_false_left + simp_to_model [contains, getKey?, diff] using List.getKey?_filter_contains_map_fst_eq_false_of_containsKey_eq_false_left theorem getKey?_diff_of_contains_right [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} (h : m₂.contains k) : (m₁.diff m₂).getKey? k = none := by revert h - simp_to_model [contains, getKey?, diff] using List.getKey?_diff_of_containsKey_right + simp_to_model [contains, getKey?, diff] using List.getKey?_filter_contains_map_fst_eq_false_of_containsKey_right /- getKey -/ theorem getKey_diff [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} {h_contains : (m₁.diff m₂).contains k} : (m₁.diff m₂).getKey k h_contains = m₁.getKey k ((contains_diff_iff h₁ h₂).1 h_contains).1 := by - simp_to_model [diff, contains, getKey] using List.getKey_diff + simp_to_model [diff, contains, getKey] using List.getKey_filter_contains_map_fst_eq_false /- getKeyD -/ theorem getKeyD_diff [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) diff --git a/src/Std/Data/Internal/List/Associative.lean b/src/Std/Data/Internal/List/Associative.lean index 54f72c50e095..1d79b4497bce 100644 --- a/src/Std/Data/Internal/List/Associative.lean +++ b/src/Std/Data/Internal/List/Associative.lean @@ -5829,22 +5829,22 @@ theorem getValueCastD_diff_of_containsKey_eq_false_left [BEq α] [LawfulBEq α] -- · exact dl₁ -- · exact h -theorem getKey?_diff [BEq α] [EquivBEq α] - {l₁ l₂ : List ((a : α) × β a)} {k : α} +theorem getKey?_filter_contains_eq_false [BEq α] [EquivBEq α] + {l₁ : List ((a : α) × β a)} {l₂ : List α} {k : α} (dl₁ : DistinctKeys l₁) : - getKey? k (List.filter (fun p => !containsKey p.fst l₂) l₁) = - if containsKey k l₂ = true then none else getKey? k l₁ := by + getKey? k (List.filter (fun p => !List.contains l₂ p.fst) l₁) = + if List.contains l₂ k = true then none else getKey? k l₁ := by simp only [getKey?_eq_getEntry?] split case isTrue h => - suffices (getEntry? k (List.filter (fun p => !containsKey p.fst l₂) l₁)) = none by + suffices (getEntry? k (List.filter (fun p => !List.contains l₂ p.fst) l₁)) = none by simp [this] rw [getEntry?_filter] · simp [Option.filter] split case h_1 _ kv heq => have key_eq := beq_of_getEntry?_eq_some heq - rw [@containsKey_congr _ _ _ _ l₂ _ _ (PartialEquivBEq.symm key_eq)] at h + rw [List.contains_congr (BEq.symm key_eq)] at h simp [h] case h_2 => rfl · exact dl₁ @@ -5857,55 +5857,97 @@ theorem getKey?_diff [BEq α] [EquivBEq α] rw [heq] have key_eq := beq_of_getEntry?_eq_some heq simp at h - rw [@containsKey_congr _ _ _ _ l₂ _ _ (PartialEquivBEq.symm key_eq)] at h + rw [List.contains_congr (BEq.symm key_eq)] at h simp [h] case h_2 _ heq => simp only [heq] · exact dl₁ -theorem getKey?_diff_of_containsKey_eq_false_right [BEq α] [EquivBEq α] +theorem getKey?_filter_contains_map_fst_eq_false [BEq α] [EquivBEq α] {l₁ l₂ : List ((a : α) × β a)} {k : α} (dl₁ : DistinctKeys l₁) : - containsKey k l₂ = false → getKey? k (List.filter (fun p => !containsKey p.fst l₂) l₁) = + getKey? k (List.filter (fun p => !List.contains (l₂.map Sigma.fst) p.fst) l₁) = +if containsKey k l₂ = true then none else getKey? k l₁ := by + rw [getKey?_filter_contains_eq_false, ← List.containsKey_eq_contains_map_fst] + exact dl₁ + +theorem getKey?_filter_contains_eq_false_of_contains_eq_false_right [BEq α] [EquivBEq α] + {l₁ : List ((a : α) × β a)} {l₂ : List α} {k : α} + (dl₁ : DistinctKeys l₁) : + List.contains l₂ k = false → getKey? k (List.filter (fun p => !List.contains l₂ p.fst) l₁) = getKey? k l₁ := by intro h - rw [getKey?_diff, h] + rw [getKey?_filter_contains_eq_false, h] · simp only [Bool.false_eq_true, ↓reduceIte] · exact dl₁ -theorem getKey?_diff_of_containsKey_eq_false_left [BEq α] [EquivBEq α] +theorem getKey?_filter_contains_map_fst_eq_false_of_containsKey_eq_false_right [BEq α] [EquivBEq α] {l₁ l₂ : List ((a : α) × β a)} {k : α} (dl₁ : DistinctKeys l₁) : - containsKey k l₁ = false → getKey? k (List.filter (fun p => !containsKey p.fst l₂) l₁) = none := by + containsKey k l₂ = false → getKey? k (List.filter (fun p => !List.contains (l₂.map Sigma.fst) p.fst) l₁) = + getKey? k l₁ := by + intro h + rw [getKey?_filter_contains_eq_false_of_contains_eq_false_right] + · exact dl₁ + · rwa [List.containsKey_eq_contains_map_fst] at h + +theorem getKey?_filter_contains_eq_false_of_containsKey_eq_false_left [BEq α] [EquivBEq α] + {l₁ : List ((a : α) × β a)} {l₂ : List α} {k : α} + (dl₁ : DistinctKeys l₁) : + containsKey k l₁ = false → getKey? k (List.filter (fun p => !List.contains l₂ p.fst) l₁) = none := by intro h - rw [getKey?_diff] + rw [getKey?_filter_contains_eq_false] · split · rfl · rw [getKey?_eq_none h] · exact dl₁ -theorem getKey?_diff_of_containsKey_right [BEq α] [EquivBEq α] +theorem getKey?_filter_contains_map_fst_eq_false_of_containsKey_eq_false_left [BEq α] [EquivBEq α] {l₁ l₂ : List ((a : α) × β a)} {k : α} (dl₁ : DistinctKeys l₁) : - containsKey k l₂ = true → getKey? k (List.filter (fun p => !containsKey p.fst l₂) l₁) = none := by + containsKey k l₁ = false → getKey? k (List.filter (fun p => !List.contains (l₂.map Sigma.fst) p.fst) l₁) = none := by intro h - rw [getKey?_diff, h] + rw [getKey?_filter_contains_eq_false_of_containsKey_eq_false_left] + · exact dl₁ + · exact h + +theorem getKey?_filter_contains_eq_false_of_containsKey_right [BEq α] [EquivBEq α] + {l₁ : List ((a : α) × β a)} {l₂ : List α} {k : α} + (dl₁ : DistinctKeys l₁) : + List.contains l₂ k → getKey? k (List.filter (fun p => !List.contains l₂ p.fst) l₁) = none := by + intro h + rw [getKey?_filter_contains_eq_false, h] · simp only [↓reduceIte] · exact dl₁ -theorem getKey_diff [BEq α] [EquivBEq α] +theorem getKey?_filter_contains_map_fst_eq_false_of_containsKey_right [BEq α] [EquivBEq α] + {l₁ l₂ : List ((a : α) × β a)} {k : α} + (dl₁ : DistinctKeys l₁) : + containsKey k l₂ → getKey? k (List.filter (fun p => !List.contains (l₂.map Sigma.fst) p.fst) l₁) = none := by + intro h + rw [getKey?_filter_contains_eq_false_of_containsKey_right] + · exact dl₁ + · rwa [List.containsKey_eq_contains_map_fst] at h + +theorem getKey_filter_contains_eq_false [BEq α] [EquivBEq α] + {l₁ : List ((a : α) × β a)} {l₂ : List α} {k : α} + (dl₁ : DistinctKeys l₁) {h₁ h₂} : + getKey k (List.filter (fun p => !List.contains l₂ p.fst) l₁) h₁ = getKey k l₁ h₂ := by + suffices some (getKey k (List.filter (fun p => !List.contains l₂ p.fst) l₁) h₁) = some (getKey k l₁ h₂) by + injections + simp only [← getKey?_eq_some_getKey] + · apply getKey?_filter_contains_eq_false_of_contains_eq_false_right + · exact dl₁ + · rw [containsKey_filter_contains_eq_false_iff] at h₁ + · replace h₁ := h₁.2 + rwa [Bool.not_eq_true] at h₁ + · exact dl₁ + +theorem getKey_filter_contains_map_fst_eq_false [BEq α] [EquivBEq α] {l₁ l₂ : List ((a : α) × β a)} {k : α} (dl₁ : DistinctKeys l₁) {h₁ h₂} : - getKey k (List.filter (fun p => !containsKey p.fst l₂) l₁) h₁ = getKey k l₁ h₂ := by - sorry - -- suffices some (getKey k (List.filter (fun p => !containsKey p.fst l₂) l₁) h₁) = some (getKey k l₁ h₂) by - -- injections - -- simp only [← getKey?_eq_some_getKey] - -- apply getKey?_diff_of_containsKey_eq_false_right dl₁ - -- · rw [containsKey_diff_eq_false_iff] at h₁ - -- · simp at h₁ - -- exact h₁.2 - -- · exact dl₁ + getKey k (List.filter (fun p => !List.contains (l₂.map Sigma.fst) p.fst) l₁) h₁ = getKey k l₁ h₂ := by + apply getKey_filter_contains_eq_false dl₁ theorem getKeyD_diff [BEq α] [EquivBEq α] {l₁ l₂ : List ((a : α) × β a)} {k : α} {fallback : α} @@ -5930,10 +5972,11 @@ theorem getKeyD_diff_of_containsKey_eq_false_right [BEq α] [EquivBEq α] (dl₁ : DistinctKeys l₁) : containsKey k l₂ = false → getKeyD k (List.filter (fun p => !containsKey p.fst l₂) l₁) fallback = getKeyD k l₁ fallback := by - intro h - rw [getKeyD_eq_getKey?, getKeyD_eq_getKey?] - congr 1 - apply getKey?_diff_of_containsKey_eq_false_right dl₁ h + sorry + -- intro h + -- rw [getKeyD_eq_getKey?, getKeyD_eq_getKey?] + -- congr 1 + -- apply getKey?_diff_of_containsKey_eq_false_right dl₁ h theorem getKeyD_diff_of_containsKey_right [BEq α] [EquivBEq α] {l₁ l₂ : List ((a : α) × β a)} {k : α} {fallback : α} From 54fb11904d93eff59bcd6ff07bcb679f03020d1f Mon Sep 17 00:00:00 2001 From: Wojciech Rozowski Date: Fri, 21 Nov 2025 17:16:01 +0000 Subject: [PATCH 80/88] Further progress --- src/Std/Data/DHashMap/Internal/RawLemmas.lean | 24 +-- src/Std/Data/Internal/List/Associative.lean | 182 ++++++++++++------ 2 files changed, 135 insertions(+), 71 deletions(-) diff --git a/src/Std/Data/DHashMap/Internal/RawLemmas.lean b/src/Std/Data/DHashMap/Internal/RawLemmas.lean index 07434835fd28..fc8498a5219f 100644 --- a/src/Std/Data/DHashMap/Internal/RawLemmas.lean +++ b/src/Std/Data/DHashMap/Internal/RawLemmas.lean @@ -3482,49 +3482,49 @@ theorem getKeyD_diff [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k fallback : α} : (m₁.diff m₂).getKeyD k fallback = if m₂.contains k then fallback else m₁.getKeyD k fallback := by - simp_to_model [diff, getKeyD, contains] using List.getKeyD_diff + simp_to_model [diff, getKeyD, contains] using List.getKeyD_filter_contains_map_fst_eq_false theorem getKeyD_diff_of_contains_eq_false_right [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k fallback : α} (h : m₂.contains k = false) : (m₁.diff m₂).getKeyD k fallback = m₁.getKeyD k fallback := by revert h - simp_to_model [contains, diff, getKeyD] using List.getKeyD_diff_of_containsKey_eq_false_right + simp_to_model [contains, diff, getKeyD] using List.getKeyD_filter_contains_map_fst_eq_false_of_contains_eq_false_right theorem getKeyD_diff_of_contains_right [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k fallback : α} (h : m₂.contains k) : (m₁.diff m₂).getKeyD k fallback = fallback := by revert h - simp_to_model [diff, getKeyD, contains] using List.getKeyD_diff_of_containsKey_right + simp_to_model [diff, getKeyD, contains] using List.getKeyD_filter_contains_map_fst_eq_false_of_contains_right theorem getKeyD_diff_of_contains_eq_false_left [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k fallback : α} (h : m₁.contains k = false) : (m₁.diff m₂).getKeyD k fallback = fallback := by revert h - simp_to_model [diff, getKeyD, contains] using List.getKeyD_diff_of_containsKey_eq_false_left + simp_to_model [diff, getKeyD, contains] using getKeyD_filter_contains_map_fst_eq_false_of_contains_eq_false_left /- getKey! -/ theorem getKey!_diff [EquivBEq α] [LawfulHashable α] [Inhabited α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} : (m₁.diff m₂).getKey! k = if m₂.contains k then default else m₁.getKey! k := by - simp_to_model [diff, getKey!, contains] using List.getKeyD_diff + simp_to_model [diff, getKey!, contains] using List.getKeyD_filter_contains_map_fst_eq_false theorem getKey!_diff_of_contains_eq_false_right [EquivBEq α] [LawfulHashable α] [Inhabited α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} (h : m₂.contains k = false) : (m₁.diff m₂).getKey! k = m₁.getKey! k := by revert h - simp_to_model [diff, getKey!, contains] using List.getKeyD_diff_of_containsKey_eq_false_right + simp_to_model [diff, getKey!, contains] using List.getKeyD_filter_contains_map_fst_eq_false_of_contains_eq_false_right theorem getKey!_diff_of_contains_right [EquivBEq α] [LawfulHashable α] [Inhabited α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} (h : m₂.contains k) : (m₁.diff m₂).getKey! k = default := by revert h - simp_to_model [diff, getKey!, contains] using List.getKeyD_diff_of_containsKey_right + simp_to_model [diff, getKey!, contains] using List.getKeyD_filter_contains_map_fst_eq_false_of_contains_right theorem getKey!_diff_of_contains_eq_false_left [EquivBEq α] [LawfulHashable α] [Inhabited α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} (h : m₁.contains k = false) : (m₁.diff m₂).getKey! k = default := by revert h - simp_to_model [diff, getKey!, contains] using List.getKeyD_diff_of_containsKey_eq_false_left + simp_to_model [diff, getKey!, contains] using List.getKeyD_filter_contains_map_fst_eq_false_of_contains_eq_false_left /- size -/ theorem size_diff_le_size_left [EquivBEq α] [LawfulHashable α] @@ -3537,12 +3537,12 @@ theorem size_diff_eq_size_left [EquivBEq α] [LawfulHashable α] (h : ∀ (a : α), m₁.contains a → m₂.contains a = false) : (m₁.diff m₂).1.size = m₁.1.size := by revert h - simp_to_model [diff, size, contains] using List.length_diff_eq_length_left + simp_to_model [diff, size, contains] using List.length_filter_contains_map_fst_eq_false_eq_length_left theorem size_diff_add_size_inter_eq_size_left [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) : (m₁.diff m₂).1.size + (m₁.inter m₂).1.size = m₁.1.size := by - simp_to_model [diff, inter, size] using List.size_diff_add_size_inter_eq_size_left + simp_to_model [diff, inter, size] using List.size_filter_contains_map_fst_eq_false_add_size_inter_eq_size_left /- isEmpty -/ @[simp] @@ -3550,11 +3550,11 @@ theorem isEmpty_diff_left [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) (h : m₁.1.isEmpty) : (m₁.diff m₂).1.isEmpty = true := by revert h - simp_to_model [isEmpty, diff] using List.isEmpty_diff_left + simp_to_model [isEmpty, diff] using List.isEmpty_filter_contains_map_fst_eq_false_left theorem isEmpty_diff_iff [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) : (m₁.diff m₂).1.isEmpty ↔ ∀ k, m₁.contains k → m₂.contains k := by - simp_to_model [diff, contains, isEmpty] using List.isEmpty_diff_iff + simp_to_model [diff, contains, isEmpty] using List.isEmpty_filter_contains_map_fst_eq_false_iff end Diff diff --git a/src/Std/Data/Internal/List/Associative.lean b/src/Std/Data/Internal/List/Associative.lean index 1d79b4497bce..1c5f026ee28d 100644 --- a/src/Std/Data/Internal/List/Associative.lean +++ b/src/Std/Data/Internal/List/Associative.lean @@ -5816,8 +5816,6 @@ if containsKey k l₂ = true then fallback else getValueCastD k l₁ fallback := rw [getValueCastD_filter_contains_eq_false, containsKey_eq_contains_map_fst] · exact dl₁ - - theorem getValueCastD_diff_of_containsKey_eq_false_left [BEq α] [LawfulBEq α] {l₁ l₂ : List ((a : α) × β a)} {k : α} {fallback} (dl₁ : DistinctKeys l₁) : @@ -5949,56 +5947,85 @@ theorem getKey_filter_contains_map_fst_eq_false [BEq α] [EquivBEq α] getKey k (List.filter (fun p => !List.contains (l₂.map Sigma.fst) p.fst) l₁) h₁ = getKey k l₁ h₂ := by apply getKey_filter_contains_eq_false dl₁ -theorem getKeyD_diff [BEq α] [EquivBEq α] - {l₁ l₂ : List ((a : α) × β a)} {k : α} {fallback : α} +theorem getKeyD_filter_contains_eq_false [BEq α] [EquivBEq α] + {l₁ : List ((a : α) × β a)} {l₂ : List α} {k : α} {fallback : α} (dl₁ : DistinctKeys l₁) : - getKeyD k (List.filter (fun p => !containsKey p.fst l₂) l₁) fallback = - if containsKey k l₂ = true then fallback else getKeyD k l₁ fallback := by - sorry - -- split - -- case isTrue h => - -- apply getKeyD_eq_fallback - -- apply containsKey_diff_of_containsKey_eq_false_right - -- · exact dl₁ - -- · exact h - -- case isFalse h => - -- rw [getKeyD_eq_getKey?, getKeyD_eq_getKey?] - -- congr 1 - -- apply getKey?_diff_of_containsKey_eq_false_right dl₁ - -- simp only [h] + getKeyD k (List.filter (fun p => !List.contains l₂ p.fst) l₁) fallback = + if List.contains l₂ k then fallback else getKeyD k l₁ fallback := by + split + case isTrue h => + apply getKeyD_eq_fallback + apply containsKey_filter_contains_eq_false_of_contains_right + · exact dl₁ + · exact h + case isFalse h => + rw [getKeyD_eq_getKey?, getKeyD_eq_getKey?] + congr 1 + apply getKey?_filter_contains_eq_false_of_contains_eq_false_right dl₁ + simp only [h] -theorem getKeyD_diff_of_containsKey_eq_false_right [BEq α] [EquivBEq α] +theorem getKeyD_filter_contains_map_fst_eq_false [BEq α] [EquivBEq α] {l₁ l₂ : List ((a : α) × β a)} {k : α} {fallback : α} (dl₁ : DistinctKeys l₁) : - containsKey k l₂ = false → getKeyD k (List.filter (fun p => !containsKey p.fst l₂) l₁) fallback = + getKeyD k (List.filter (fun p => !List.contains (l₂.map Sigma.fst) p.fst) l₁) fallback = + if containsKey k l₂ then fallback else getKeyD k l₁ fallback := by + rw [getKeyD_filter_contains_eq_false, containsKey_eq_contains_map_fst] + · exact dl₁ + +theorem getKeyD_filter_contains_eq_false_of_contains_eq_false_right [BEq α] [EquivBEq α] + {l₁ : List ((a : α) × β a)} {l₂ : List α} {k : α} {fallback : α} + (dl₁ : DistinctKeys l₁) : + List.contains l₂ k = false → getKeyD k (List.filter (fun p => !List.contains l₂ p.fst) l₁) fallback = getKeyD k l₁ fallback := by - sorry - -- intro h - -- rw [getKeyD_eq_getKey?, getKeyD_eq_getKey?] - -- congr 1 - -- apply getKey?_diff_of_containsKey_eq_false_right dl₁ h + intro h + rw [getKeyD_eq_getKey?, getKeyD_eq_getKey?] + congr 1 + apply getKey?_filter_contains_eq_false_of_contains_eq_false_right dl₁ h -theorem getKeyD_diff_of_containsKey_right [BEq α] [EquivBEq α] +theorem getKeyD_filter_contains_map_fst_eq_false_of_contains_eq_false_right [BEq α] [EquivBEq α] {l₁ l₂ : List ((a : α) × β a)} {k : α} {fallback : α} (dl₁ : DistinctKeys l₁) : - containsKey k l₂ = true → getKeyD k (List.filter (fun p => !containsKey p.fst l₂) l₁) fallback = fallback := by - sorry - -- intro h - -- apply getKeyD_eq_fallback - -- apply containsKey_diff_of_containsKey_eq_false_right - -- · exact dl₁ - -- · exact h + containsKey k l₂ = false → getKeyD k (List.filter (fun p => !List.contains (l₂.map Sigma.fst) p.fst) l₁) fallback = + getKeyD k l₁ fallback := by + intro h + rw [getKeyD_filter_contains_eq_false_of_contains_eq_false_right dl₁] + rwa [containsKey_eq_contains_map_fst] at h + +theorem getKeyD_filter_contains_eq_false_of_contains_right [BEq α] [EquivBEq α] + {l₁ : List ((a : α) × β a)} {l₂ : List α} {k : α} {fallback : α} + (dl₁ : DistinctKeys l₁) : + List.contains l₂ k → getKeyD k (List.filter (fun p => !List.contains l₂ p.fst) l₁) fallback = fallback := by + intro h + apply getKeyD_eq_fallback + rw [containsKey_filter_contains_eq_false_of_contains_right h] + · exact dl₁ -theorem getKeyD_diff_of_containsKey_eq_false_left [BEq α] [EquivBEq α] +theorem getKeyD_filter_contains_map_fst_eq_false_of_contains_right [BEq α] [EquivBEq α] {l₁ l₂ : List ((a : α) × β a)} {k : α} {fallback : α} (dl₁ : DistinctKeys l₁) : - containsKey k l₁ = false → getKeyD k (List.filter (fun p => !containsKey p.fst l₂) l₁) fallback = fallback := by - sorry - -- intro h - -- apply getKeyD_eq_fallback - -- apply containsKey_diff_eq_false_of_containsKey_eq_false_left - -- · exact dl₁ - -- · exact h + containsKey k l₂ → getKeyD k (List.filter (fun p => !List.contains (l₂.map Sigma.fst) p.fst) l₁) fallback = fallback := by + rw [containsKey_eq_contains_map_fst] + apply getKeyD_filter_contains_eq_false_of_contains_right + exact dl₁ + +theorem getKeyD_filter_contains_eq_false_of_contains_eq_false_left [BEq α] [EquivBEq α] + {l₁ : List ((a : α) × β a)} {l₂ : List α} {k : α} {fallback} + (dl₁ : DistinctKeys l₁) : + containsKey k l₁ = false → getKeyD k (List.filter (fun p => !List.contains l₂ p.fst) l₁) fallback = fallback := by + intro h + apply getKeyD_eq_fallback + apply containsKey_filter_contains_eq_false_eq_false_left + · exact dl₁ + · exact h + +theorem getKeyD_filter_contains_map_fst_eq_false_of_contains_eq_false_left [BEq α] [EquivBEq α] + {l₁ l₂ : List ((a : α) × β a)} {k : α} {fallback} + (dl₁ : DistinctKeys l₁) : + containsKey k l₁ = false → getKeyD k (List.filter (fun p => !List.contains (l₂.map Sigma.fst) p.fst) l₁) fallback = fallback := by + intro h + apply getKeyD_filter_contains_eq_false_of_contains_eq_false_left + · exact dl₁ + · exact h theorem getValueCastD_filter_contains_eq_false_of_containsKey_eq_false_left {β : α → Type v} [BEq α] [LawfulBEq α] {l₁ : List ((a : α) × β a)} {l₂ : List α} {k : α} {fallback : β k} @@ -6017,11 +6044,10 @@ theorem getValueCastD_filter_contains_map_fst_eq_false_of_containsKey_eq_false_l apply getValueCastD_filter_contains_eq_false_of_containsKey_eq_false_left · exact dl₁ - -theorem length_diff_eq_length_left [BEq α] [EquivBEq α] - {l₁ l₂ : List ((a : α) × β a)} (hl₁ : DistinctKeys l₁) - (w : ∀ (a : α), containsKey a l₁ → containsKey a l₂ = false) : - (l₁.filter fun p => !containsKey p.fst l₂).length = l₁.length := by +theorem length_filter_contains_eq_false_eq_length_left [BEq α] [EquivBEq α] + {l₁ : List ((a : α) × β a)} {l₂ : List α} (hl₁ : DistinctKeys l₁) + (w : ∀ (a : α), containsKey a l₁ → List.contains l₂ a = false) : + (l₁.filter fun p => !List.contains l₂ p.fst).length = l₁.length := by rw [← List.countP_eq_length_filter] induction l₁ case nil => simp @@ -6040,10 +6066,18 @@ theorem length_diff_eq_length_left [BEq α] [EquivBEq α] apply w rw [containsKey_cons, BEq.rfl, Bool.true_or] -theorem size_diff_add_size_inter_eq_size_left [BEq α] [EquivBEq α] - {l₁ l₂ : List ((a : α) × β a)} (dl₁ : DistinctKeys l₁) : - (List.filter (fun p => !containsKey p.fst l₂) l₁).length + - (List.filter (fun p => containsKey p.fst l₂) l₁).length = l₁.length := by +theorem length_filter_contains_map_fst_eq_false_eq_length_left [BEq α] [EquivBEq α] + {l₁ l₂ : List ((a : α) × β a)} (hl₁ : DistinctKeys l₁) + (w : ∀ (a : α), containsKey a l₁ → containsKey a l₂ = false) : + (l₁.filter fun p => !List.contains (l₂.map Sigma.fst) p.fst).length = l₁.length := by + apply length_filter_contains_eq_false_eq_length_left hl₁ + intro a mem + rw [← List.containsKey_eq_contains_map_fst, w a mem] + +theorem size_filter_contains_eq_false_add_size_inter_eq_size_left [BEq α] [EquivBEq α] + {l₁ : List ((a : α) × β a)} {l₂ : List α} (dl₁ : DistinctKeys l₁) : + (List.filter (fun p => !List.contains l₂ p.fst) l₁).length + + (List.filter (fun p => List.contains l₂ p.fst) l₁).length = l₁.length := by induction l₁ with | nil => simp | cons h t ih => @@ -6060,17 +6094,36 @@ theorem size_diff_add_size_inter_eq_size_left [BEq α] [EquivBEq α] simp only [hc, ↓reduceIte, List.length_cons] omega -theorem isEmpty_diff_left [BEq α] [EquivBEq α] - {l₁ l₂ : List ((a : α) × β a)} : - l₁.isEmpty → (List.filter (fun p => !containsKey p.fst l₂) l₁).isEmpty := by +theorem size_filter_contains_map_fst_eq_false_add_size_inter_eq_size_left [BEq α] [EquivBEq α] + {l₁ l₂ : List ((a : α) × β a)} (dl₁ : DistinctKeys l₁) : + (List.filter (fun p => !List.contains (l₂.map Sigma.fst) p.fst) l₁).length + + (List.filter (fun p => containsKey p.fst l₂) l₁).length = l₁.length := by + conv => + lhs + rhs + congr + lhs + ext p + rw [containsKey_eq_contains_map_fst] + apply size_filter_contains_eq_false_add_size_inter_eq_size_left dl₁ + +theorem isEmpty_filter_contains_eq_false_left [BEq α] [EquivBEq α] + {l₁ : List ((a : α) × β a)} {l₂ : List α}: + l₁.isEmpty → (List.filter (fun p => !List.contains l₂ p.fst) l₁).isEmpty := by intro hyp simp at hyp simp [hyp] -theorem isEmpty_diff_iff [BEq α] [EquivBEq α] - {l₁ l₂ : List ((a : α) × β a)} (dl₁ : DistinctKeys l₁) : - (List.filter (fun p => !containsKey p.fst l₂) l₁).isEmpty = true ↔ - ∀ (k : α), containsKey k l₁ = true → containsKey k l₂ = true := by +theorem isEmpty_filter_contains_map_fst_eq_false_left [BEq α] [EquivBEq α] + {l₁ l₂ : List ((a : α) × β a)} : + l₁.isEmpty → (List.filter (fun p => !List.contains (l₂.map Sigma.fst) p.fst) l₁).isEmpty := by + apply isEmpty_filter_contains_eq_false_left + + +theorem isEmpty_filter_contains_eq_false_iff [BEq α] [EquivBEq α] + {l₁ : List ((a : α) × β a)} {l₂ : List α} (dl₁ : DistinctKeys l₁) : + (List.filter (fun p => !List.contains l₂ p.fst) l₁).isEmpty = true ↔ + ∀ (k : α), containsKey k l₁ = true → List.contains l₂ k := by constructor case mpr => intro hyp @@ -6096,8 +6149,8 @@ theorem isEmpty_diff_iff [BEq α] [EquivBEq α] simp only [Bool.not_eq_eq_eq_not, Bool.not_true] at heq cases mem case inl heq2 => - rw [containsKey_congr (PartialEquivBEq.symm heq2)] - simp [heq] + rw [Bool.not_eq_false] at heq + rwa [List.contains_congr heq2] at heq case inr heq2 => apply ih · rw [List.distinctKeys_cons_iff] at dl₁ @@ -6105,6 +6158,17 @@ theorem isEmpty_diff_iff [BEq α] [EquivBEq α] · exact hyp · exact heq2 +theorem isEmpty_filter_contains_map_fst_eq_false_iff [BEq α] [EquivBEq α] + {l₁ l₂ : List ((a : α) × β a)} (dl₁ : DistinctKeys l₁) : + (List.filter (fun p => !List.contains (l₂.map Sigma.fst) p.fst) l₁).isEmpty = true ↔ + ∀ (k : α), containsKey k l₁ = true → containsKey k l₂ := by + conv => + rhs + ext k + rhs + rw [containsKey_eq_contains_map_fst] + apply isEmpty_filter_contains_eq_false_iff dl₁ + theorem getValue?_diff {β : Type v} [BEq α] [EquivBEq α] {l₁ l₂ : List ((_ : α) × β)} {k : α} (dl₁ : DistinctKeys l₁) : From c3ce6449076393aa485433937baf7c0d3a1da510 Mon Sep 17 00:00:00 2001 From: Wojciech Rozowski Date: Fri, 21 Nov 2025 17:46:24 +0000 Subject: [PATCH 81/88] Refactor progress --- src/Std/Data/DHashMap/Internal/RawLemmas.lean | 4 +- src/Std/Data/Internal/List/Associative.lean | 189 +++++++++++------- 2 files changed, 124 insertions(+), 69 deletions(-) diff --git a/src/Std/Data/DHashMap/Internal/RawLemmas.lean b/src/Std/Data/DHashMap/Internal/RawLemmas.lean index fc8498a5219f..355bb61c016b 100644 --- a/src/Std/Data/DHashMap/Internal/RawLemmas.lean +++ b/src/Std/Data/DHashMap/Internal/RawLemmas.lean @@ -3565,13 +3565,13 @@ variable {β : Type v} {m₁ m₂ : Raw₀ α (fun _ => β)} /- get? -/ theorem get?_diff [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} : Const.get? (m₁.diff m₂) k = if m₂.contains k then none else Const.get? m₁ k := by - simp_to_model [diff, Const.get?, contains] using List.getValue?_diff + simp_to_model [diff, Const.get?, contains] using List.getValue?_filter_contains_map_fst_eq_false theorem get?_diff_of_contains_eq_false_right [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} (h : m₂.contains k = false) : Const.get? (m₁.diff m₂) k = Const.get? m₁ k := by revert h - simp_to_model [diff, contains, Const.get?] using List.getValue?_diff_of_containsKey_eq_false_right + simp_to_model [diff, contains, Const.get?] using List.getValue?_filter_contains_map_fst_eq_false_of_contains_eq_false_right theorem get?_diff_of_contains_eq_false_left [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} (h : m₁.contains k = false) : diff --git a/src/Std/Data/Internal/List/Associative.lean b/src/Std/Data/Internal/List/Associative.lean index 1c5f026ee28d..76d016635401 100644 --- a/src/Std/Data/Internal/List/Associative.lean +++ b/src/Std/Data/Internal/List/Associative.lean @@ -6169,105 +6169,160 @@ theorem isEmpty_filter_contains_map_fst_eq_false_iff [BEq α] [EquivBEq α] rw [containsKey_eq_contains_map_fst] apply isEmpty_filter_contains_eq_false_iff dl₁ -theorem getValue?_diff {β : Type v} [BEq α] [EquivBEq α] - {l₁ l₂ : List ((_ : α) × β)} {k : α} +theorem getValue?_filter_contains_eq_false {β : Type v} [BEq α] [EquivBEq α] + {l₁ : List ((_ : α) × β)} {l₂ : List α} {k : α} (dl₁ : DistinctKeys l₁) : - getValue? k (List.filter (fun p => !containsKey p.fst l₂) l₁) = - if containsKey k l₂ = true then none else getValue? k l₁ := by - sorry - -- simp only [getValue?_eq_getEntry?] - -- split - -- case isTrue h => - -- simp only [Option.map_eq_none_iff, getEntry?_eq_none] - -- apply containsKey_diff_of_containsKey_eq_false_right - -- · exact dl₁ - -- · exact h - -- case isFalse h => - -- congr 1 - -- rw [getEntry?_filter] - -- · simp only [Option.filter] - -- split - -- case h_1 _ val heq => - -- simp [heq] - -- have := PartialEquivBEq.symm <| @beq_of_getEntry?_eq_some α (fun _ => β) _ l₁ k val heq - -- rw [← containsKey_congr this] - -- simp at h - -- exact h - -- case h_2 _ heq => - -- simp [heq] - -- · exact dl₁ - -theorem getValue?_diff_of_containsKey_eq_false_right {β : Type v} [BEq α] [EquivBEq α] + getValue? k (List.filter (fun p => !List.contains l₂ p.fst) l₁) = + if List.contains l₂ k then none else getValue? k l₁ := by + simp only [getValue?_eq_getEntry?] + split + case isTrue h => + simp only [Option.map_eq_none_iff, getEntry?_eq_none] + apply containsKey_filter_contains_eq_false_of_contains_right + · exact dl₁ + · exact h + case isFalse h => + congr 1 + rw [getEntry?_filter] + · simp only [Option.filter] + split + case h_1 _ val heq => + simp only [Bool.not_eq_eq_eq_not, Bool.not_true, heq, ite_eq_left_iff, Bool.not_eq_false, + reduceCtorEq, imp_false, Bool.not_eq_true] + have := @beq_of_getEntry?_eq_some α (fun _ => β) _ l₁ k val heq + rw [Bool.not_eq_true] at h + rw [List.contains_congr this, h] + case h_2 _ heq => + simp [heq] + · exact dl₁ + +theorem getValue?_filter_contains_map_fst_eq_false {β : Type v} [BEq α] [EquivBEq α] {l₁ l₂ : List ((_ : α) × β)} {k : α} (dl₁ : DistinctKeys l₁) : - containsKey k l₂ = false → getValue? k (List.filter (fun p => !containsKey p.fst l₂) l₁) = + getValue? k (List.filter (fun p => !List.contains (l₂.map Sigma.fst) p.fst) l₁) = +if containsKey k l₂ then none else getValue? k l₁ := by + rw [getValue?_filter_contains_eq_false, ← List.containsKey_eq_contains_map_fst] + exact dl₁ + +theorem getValue?_filter_contains_eq_false_of_contains_eq_false_right {β : Type v} [BEq α] [EquivBEq α] + {l₁ : List ((_ : α) × β)} {l₂ : List α} {k : α} + (dl₁ : DistinctKeys l₁) : + List.contains l₂ k = false → getValue? k (List.filter (fun p => !List.contains l₂ p.fst) l₁) = getValue? k l₁ := by intro h - rw [getValue?_diff, h] + rw [getValue?_filter_contains_eq_false, h] · simp only [Bool.false_eq_true, ↓reduceIte] · exact dl₁ -theorem getValue?_diff_of_containsKey_eq_false_left {β : Type v} [BEq α] [EquivBEq α] +theorem getValue?_filter_contains_map_fst_eq_false_of_contains_eq_false_right {β : Type v} [BEq α] [EquivBEq α] {l₁ l₂ : List ((_ : α) × β)} {k : α} (dl₁ : DistinctKeys l₁) : - containsKey k l₁ = false → getValue? k (List.filter (fun p => !containsKey p.fst l₂) l₁) = none := by + containsKey k l₂ = false → getValue? k (List.filter (fun p => !List.contains (l₂.map Sigma.fst) p.fst) l₁) = + getValue? k l₁ := by + intro h + apply getValue?_filter_contains_eq_false_of_contains_eq_false_right dl₁ + rwa [List.containsKey_eq_contains_map_fst] at h + +theorem getValue?_filter_contains_eq_false_of_containsKey_eq_false_left {β : Type v} [BEq α] [EquivBEq α] + {l₁ : List ((_ : α) × β)} {l₂ : List α} {k : α} + (dl₁ : DistinctKeys l₁) : + containsKey k l₁ = false → getValue? k (List.filter (fun p => !List.contains l₂ p.fst) l₁) = none := by intro h - rw [getValue?_diff] + rw [getValue?_filter_contains_eq_false] · split · rfl · rw [@getValue?_eq_none α β _ l₁ k] exact h · exact dl₁ -theorem getValue?_diff_of_containsKey_right {β : Type v} [BEq α] [EquivBEq α] +theorem getValue?_filter_contains_map_fst_eq_false_of_containsKey_eq_false_left {β : Type v} [BEq α] [EquivBEq α] {l₁ l₂ : List ((_ : α) × β)} {k : α} (dl₁ : DistinctKeys l₁) : - containsKey k l₂ = true → getValue? k (List.filter (fun p => !containsKey p.fst l₂) l₁) = none := by + containsKey k l₁ = false → getValue? k (List.filter (fun p => !List.contains (l₂.map Sigma.fst) p.fst) l₁) = none := by + apply getValue?_filter_contains_eq_false_of_containsKey_eq_false_left dl₁ + +theorem getValue?_filter_contains_eq_false_of_contains_right {β : Type v} [BEq α] [EquivBEq α] + {l₁ : List ((_ : α) × β)} {l₂ : List α} {k : α} + (dl₁ : DistinctKeys l₁) : + List.contains l₂ k → getValue? k (List.filter (fun p => !List.contains l₂ p.fst) l₁) = none := by intro h - rw [getValue?_diff, h] + rw [getValue?_filter_contains_eq_false, h] · simp only [↓reduceIte] · exact dl₁ -theorem getValue_diff {β : Type v} [BEq α] [EquivBEq α] +theorem getValue?_filter_contains_map_fst_eq_false_of_contains_right {β : Type v} [BEq α] [EquivBEq α] {l₁ l₂ : List ((_ : α) × β)} {k : α} + (dl₁ : DistinctKeys l₁) : + containsKey k l₂ → getValue? k (List.filter (fun p => !List.contains (l₂.map Sigma.fst) p.fst) l₁) = none := by + intro h + rw [getValue?_filter_contains_eq_false_of_contains_right dl₁] + rwa [List.containsKey_eq_contains_map_fst] at h + +theorem getValue_filter_contains_eq_false {β : Type v} [BEq α] [EquivBEq α] + {l₁ : List ((_ : α) × β)} {l₂ : List α} {k : α} (dl₁ : DistinctKeys l₁) {h₁ h₂} : - getValue k (List.filter (fun p => !containsKey p.fst l₂) l₁) h₁ = getValue k l₁ h₂ := by sorry - -- suffices some (getValue k (List.filter (fun p => !containsKey p.fst l₂) l₁) h₁) = some (getValue k l₁ h₂) by - -- injections - -- simp only [← getValue?_eq_some_getValue] - -- apply getValue?_diff_of_containsKey_eq_false_right dl₁ - -- · rw [containsKey_diff_eq_false_iff] at h₁ - -- · simp at h₁ - -- exact h₁.2 - -- · exact dl₁ - -theorem getValueD_diff {β : Type v} [BEq α] [EquivBEq α] - {l₁ l₂ : List ((_ : α) × β)} {k : α} {fallback : β} + getValue k (List.filter (fun p => !List.contains l₂ p.fst) l₁) h₁ = getValue k l₁ h₂ := by + suffices some (getValue k (List.filter (fun p => !List.contains l₂ p.fst) l₁) h₁) = some (getValue k l₁ h₂) by + injections + simp only [← getValue?_eq_some_getValue] + apply getValue?_filter_contains_eq_false_of_contains_eq_false_right + · rw [containsKey_filter_contains_eq_false_iff] at h₁ + · exact dl₁ + · exact dl₁ + · rw [containsKey_filter_contains_eq_false_iff] at h₁ + · replace h₁ := h₁.2 + simp [h₁] + · exact dl₁ + +theorem getValue_filter_contains_map_fst_eq_false {β : Type v} [BEq α] [EquivBEq α] + {l₁ l₂ : List ((_ : α) × β)} {k : α} + (dl₁ : DistinctKeys l₁) {h₁ h₂} : + getValue k (List.filter (fun p => !List.contains (l₂.map Sigma.fst) p.fst) l₁) h₁ = getValue k l₁ h₂ := by + apply getValue_filter_contains_eq_false dl₁ + +theorem getValueD_filter_contains_eq_false {β : Type v} [BEq α] [EquivBEq α] + {l₁ : List ((_ : α) × β)} {l₂ : List α} {k : α} {fallback : β} (dl₁ : DistinctKeys l₁) : - getValueD k (List.filter (fun p => !containsKey p.fst l₂) l₁) fallback = - if containsKey k l₂ = true then fallback else getValueD k l₁ fallback := by - sorry - -- split - -- case isTrue h => - -- apply getValueD_eq_fallback - -- apply containsKey_diff_of_containsKey_eq_false_right - -- · exact dl₁ - -- · exact h - -- case isFalse h => - -- rw [getValueD_eq_getValue?, getValueD_eq_getValue?] - -- congr 1 - -- apply getValue?_diff_of_containsKey_eq_false_right dl₁ - -- simp only [h] - -theorem getValueD_diff_of_containsKey_eq_false_right {β : Type v} [BEq α] [EquivBEq α] + getValueD k (List.filter (fun p => !List.contains l₂ p.fst) l₁) fallback = + if List.contains l₂ k then fallback else getValueD k l₁ fallback := by + split + case isTrue h => + apply getValueD_eq_fallback + apply containsKey_filter_contains_eq_false_of_contains_right + · exact dl₁ + · exact h + case isFalse h => + rw [getValueD_eq_getValue?, getValueD_eq_getValue?] + congr 1 + apply getValue?_filter_contains_eq_false_of_contains_eq_false_right dl₁ + simp [h] + +theorem getValueD_filter_contains_map_fst_eq_false {β : Type v} [BEq α] [EquivBEq α] {l₁ l₂ : List ((_ : α) × β)} {k : α} {fallback : β} (dl₁ : DistinctKeys l₁) : - containsKey k l₂ = false → getValueD k (List.filter (fun p => !containsKey p.fst l₂) l₁) fallback = + getValueD k (List.filter (fun p => !List.contains (l₂.map Sigma.fst) p.fst) l₁) fallback = + if containsKey k l₂ then fallback else getValueD k l₁ fallback := by + rw [containsKey_eq_contains_map_fst] + apply getValueD_filter_contains_eq_false dl₁ + +theorem getValueD_filter_contains_eq_false_of_containsKey_eq_false_right {β : Type v} [BEq α] [EquivBEq α] + {l₁ : List ((_ : α) × β)} {l₂ : List α} {k : α} {fallback : β} + (dl₁ : DistinctKeys l₁) : + List.contains l₂ k = false → getValueD k (List.filter (fun p => !List.contains l₂ p.fst) l₁) fallback = getValueD k l₁ fallback := by intro h rw [getValueD_eq_getValue?, getValueD_eq_getValue?] congr 1 - apply getValue?_diff_of_containsKey_eq_false_right dl₁ h + apply getValue?_filter_contains_eq_false_of_contains_eq_false_right dl₁ h + +theorem getValueD_filter_contains_map_fst_eq_false_of_containsKey_eq_false_right {β : Type v} [BEq α] [EquivBEq α] + {l₁ l₂ : List ((_ : α) × β)} {k : α} {fallback : β} + (dl₁ : DistinctKeys l₁) : + containsKey k l₂ = false → getValueD k (List.filter (fun p => !List.contains (l₂.map Sigma.fst) p.fst) l₁) fallback = + getValueD k l₁ fallback := by + intro h + apply getValueD_filter_contains_eq_false_of_containsKey_eq_false_right dl₁ + rwa [containsKey_eq_contains_map_fst] at h theorem getValueD_filter_contains_eq_false_of_contains_right {β : Type v} [BEq α] [EquivBEq α] {l₁ : List ((_ : α) × β)} {l₂ : List α} {k : α} {fallback : β} From 12e76c7cb9b7f5c594a1a75e81aa15d3ac4f7eb5 Mon Sep 17 00:00:00 2001 From: Wojciech Rozowski Date: Fri, 21 Nov 2025 17:53:32 +0000 Subject: [PATCH 82/88] Lemmas pass --- src/Std/Data/DHashMap/Internal/RawLemmas.lean | 20 +++++++++---------- src/Std/Data/Internal/List/Associative.lean | 20 +++++++++++-------- 2 files changed, 22 insertions(+), 18 deletions(-) diff --git a/src/Std/Data/DHashMap/Internal/RawLemmas.lean b/src/Std/Data/DHashMap/Internal/RawLemmas.lean index 355bb61c016b..4d125454c5b6 100644 --- a/src/Std/Data/DHashMap/Internal/RawLemmas.lean +++ b/src/Std/Data/DHashMap/Internal/RawLemmas.lean @@ -3577,33 +3577,33 @@ theorem get?_diff_of_contains_eq_false_left [EquivBEq α] [LawfulHashable α] (h {k : α} (h : m₁.contains k = false) : Const.get? (m₁.diff m₂) k = none := by revert h - simp_to_model [diff, Const.get?, contains] using List.getValue?_diff_of_containsKey_eq_false_left + simp_to_model [diff, Const.get?, contains] using List.getValue?_filter_contains_map_fst_eq_false_of_containsKey_eq_false_left theorem get?_diff_of_contains_right [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} (h : m₂.contains k) : Const.get? (m₁.diff m₂) k = none := by revert h - simp_to_model [diff, Const.get?, contains] using List.getValue?_diff_of_containsKey_right + simp_to_model [diff, Const.get?, contains] using List.getValue?_filter_contains_map_fst_eq_false_of_contains_right /- get -/ theorem get_diff [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} {h_contains : (m₁.diff m₂).contains k} : Const.get (m₁.diff m₂) k h_contains = Const.get m₁ k ((contains_diff_iff h₁ h₂).1 h_contains).1 := by - simp_to_model [diff, Const.get, contains] using List.getValue_diff + simp_to_model [diff, Const.get, contains] using List.getValue_filter_contains_eq_false /- getD -/ theorem getD_diff [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} {fallback : β} : Const.getD (m₁.diff m₂) k fallback = if m₂.contains k then fallback else Const.getD m₁ k fallback := by - simp_to_model [diff, Const.getD, contains] using List.getValueD_diff + simp_to_model [diff, Const.getD, contains] using List.getValueD_filter_contains_map_fst_eq_false theorem getD_diff_of_contains_eq_false_right [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} {fallback : β} (h : m₂.contains k = false) : Const.getD (m₁.diff m₂) k fallback = Const.getD m₁ k fallback := by revert h - simp_to_model [diff, contains, Const.getD] using List.getValueD_diff_of_containsKey_eq_false_right + simp_to_model [diff, contains, Const.getD] using List.getValueD_filter_contains_map_fst_eq_false_of_containsKey_eq_false_right theorem getD_diff_of_contains_right [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} {fallback : β} (h : m₂.contains k) : @@ -3615,30 +3615,30 @@ theorem getD_diff_of_contains_eq_false_left [EquivBEq α] [LawfulHashable α] (h {k : α} {fallback : β} (h : m₁.contains k = false) : Const.getD (m₁.diff m₂) k fallback = fallback := by revert h - simp_to_model [diff, Const.getD, contains] using List.getValueD_diff_of_containsKey_eq_false_left + simp_to_model [diff, Const.getD, contains] using List. /- get! -/ theorem get!_diff [EquivBEq α] [LawfulHashable α] [Inhabited β] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} : Const.get! (m₁.diff m₂) k = if m₂.contains k then default else Const.get! m₁ k := by - simp_to_model [diff, Const.get!, contains] using List.getValueD_diff + simp_to_model [diff, Const.get!, contains] using List.getValueD_filter_contains_map_fst_eq_false theorem get!_diff_of_contains_eq_false_right [EquivBEq α] [LawfulHashable α] [Inhabited β] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} (h : m₂.contains k = false) : Const.get! (m₁.diff m₂) k = Const.get! m₁ k := by revert h - simp_to_model [diff, contains, Const.get!] using List.getValueD_diff_of_containsKey_eq_false_right + simp_to_model [diff, contains, Const.get!] using List.getValueD_filter_contains_map_fst_eq_false_of_containsKey_eq_false_right theorem get!_diff_of_contains_right [EquivBEq α] [LawfulHashable α] [Inhabited β] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} (h : m₂.contains k) : Const.get! (m₁.diff m₂) k = default := by revert h - simp_to_model [diff, Const.get!, contains] using List.getValueD_diff_of_containsKey_right + simp_to_model [diff, Const.get!, contains] using List.getValueD_filter_contains_map_fst_eq_false_of_contains_map_fst_right theorem get!_diff_of_contains_eq_false_left [EquivBEq α] [LawfulHashable α] [Inhabited β] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} (h : m₁.contains k = false) : Const.get! (m₁.diff m₂) k = default := by revert h - simp_to_model [diff, Const.get!, contains] using List.getValueD_diff_of_containsKey_eq_false_left + simp_to_model [diff, Const.get!, contains] using List.getValueD_filter_contains_map_fst_eq_false_of_containsKey_eq_false_left end Const diff --git a/src/Std/Data/Internal/List/Associative.lean b/src/Std/Data/Internal/List/Associative.lean index 76d016635401..15d047cf7ef8 100644 --- a/src/Std/Data/Internal/List/Associative.lean +++ b/src/Std/Data/Internal/List/Associative.lean @@ -6344,17 +6344,21 @@ theorem getValueD_filter_contains_map_fst_eq_false_of_contains_map_fst_right {β · exact dl₁ · exact hyp - theorem getValueD_diff_of_containsKey_eq_false_left {β : Type v} [BEq α] [EquivBEq α] + {l₁ : List ((_ : α) × β)} {l₂ : List α} {k : α} {fallback : β} + (dl₁ : DistinctKeys l₁) : + containsKey k l₁ = false → getValueD k (List.filter (fun p => !List.contains l₂ p.fst) l₁) fallback = fallback := by + intro h + apply getValueD_eq_fallback + apply containsKey_filter_contains_eq_false_eq_false_left + · exact dl₁ + · exact h + +theorem getValueD_filter_contains_map_fst_eq_false_of_containsKey_eq_false_left {β : Type v} [BEq α] [EquivBEq α] {l₁ l₂ : List ((_ : α) × β)} {k : α} {fallback : β} (dl₁ : DistinctKeys l₁) : - containsKey k l₁ = false → getValueD k (List.filter (fun p => !containsKey p.fst l₂) l₁) fallback = fallback := by - sorry - -- intro h - -- apply getValueD_eq_fallback - -- apply containsKey_diff_eq_false_of_containsKey_eq_false_left - -- · exact dl₁ - -- · exact h + containsKey k l₁ = false → getValueD k (List.filter (fun p => !List.contains (l₂.map Sigma.fst) p.fst) l₁) fallback = fallback := by + apply getValueD_diff_of_containsKey_eq_false_left dl₁ theorem getValueCastD_filter_contains_eq_false_of_contains_right {β : α → Type v} [BEq α] [LawfulBEq α] {l₁ : List ((a : α) × β a)} {l₂ : List α} {k : α} {fallback : β k} From 005b5e7ee2c157460428d38ed8bc496471067d8a Mon Sep 17 00:00:00 2001 From: Wojciech Rozowski Date: Fri, 21 Nov 2025 17:54:27 +0000 Subject: [PATCH 83/88] Clean up `Associative.lean` --- src/Std/Data/Internal/List/Associative.lean | 25 ++------------------- 1 file changed, 2 insertions(+), 23 deletions(-) diff --git a/src/Std/Data/Internal/List/Associative.lean b/src/Std/Data/Internal/List/Associative.lean index 15d047cf7ef8..21f2a2283384 100644 --- a/src/Std/Data/Internal/List/Associative.lean +++ b/src/Std/Data/Internal/List/Associative.lean @@ -5233,7 +5233,6 @@ theorem getValueCast?_filter_containsKey [BEq α] [LawfulBEq α] case h_2 => rfl exact dl₁ -/- `get?_diff` -/ theorem getValueCast?_filter_contains_eq_false [BEq α] [LawfulBEq α] {l₁ : List ((a : α) × β a)} {l₂ : List α} {k : α} (dl₁ : DistinctKeys l₁) : @@ -5289,8 +5288,6 @@ theorem getValueCast?_filter_contains_map_fst_eq_false [BEq α] [LawfulBEq α] rw [getValueCast?_filter_contains_eq_false, ← containsKey_eq_contains_map_fst] exact dl₁ -/- `get?_diff_of_contains_eq_false_right` -/ - theorem getValueCast?_filter_contains_eq_false_of_contains_eq_false_right [BEq α] [LawfulBEq α] {l₁ : List ((a : α) × β a)} {l₂ : List α} {k : α} (dl₁ : DistinctKeys l₁) : @@ -5310,7 +5307,6 @@ theorem getValueCast?_filter_contains_map_fst_eq_false_of_containsKey_eq_false_r rw [getValueCast?_filter_contains_eq_false_of_contains_eq_false_right dl₁] · exact h₁ -/- `get?_diff_of_contains_eq_false_left` -/ theorem getValueCast?_filter_contains_eq_false_of_contains_eq_false_left [BEq α] [LawfulBEq α] {l₁ : List ((a : α) × β a)} {l₂ : List α} {k : α} (dl₁ : DistinctKeys l₁) : @@ -5330,7 +5326,6 @@ theorem getValueCast?_filter_contains_map_fst_eq_false_of_containsKey_eq_false_l rw [getValueCast?_filter_contains_eq_false_of_contains_eq_false_left dl₁] · exact h₁ -/- `get?_diff_of_contains_right` -/ theorem getValueCast?_filter_contains_eq_false_of_contains_right [BEq α] [LawfulBEq α] {l₁ : List ((a : α) × β a)} {l₂ : List α} {k : α} (dl₁ : DistinctKeys l₁) : @@ -5627,7 +5622,6 @@ theorem containsKey_filter_containsKey_iff [BEq α] [EquivBEq α] {l₁ l₂ : L . simp . exact hl₁ -/- `contains_diff` -/ theorem containsKey_filter_contains_eq_false [BEq α] [EquivBEq α] {l₁ : List ((a : α) × β a)} {l₂ : List α} {hl₁ : DistinctKeys l₁} {k : α} : (containsKey k (List.filter (fun p => !List.contains l₂ p.fst ) l₁)) = (containsKey k l₁ && !List.contains l₂ k) := by @@ -5649,7 +5643,6 @@ theorem containsKey_filter_contains_map_fst_eq_false [BEq α] [EquivBEq α] {l rw [containsKey_filter_contains_eq_false, ← containsKey_eq_contains_map_fst] · exact hl₁ -/- `contains_diff_iff` -/ theorem containsKey_filter_contains_eq_false_iff [BEq α] [EquivBEq α] {l₁ : List ((a : α) × β a)} {l₂ : List α} {hl₁ : DistinctKeys l₁} {k : α} : (containsKey k (List.filter (fun p => !List.contains l₂ p.fst) l₁)) ↔ (containsKey k l₁ ∧ ¬ List.contains l₂ k) := by @@ -5663,8 +5656,6 @@ theorem containsKey_filter_contains_map_fst_eq_false_iff [BEq α] [EquivBEq α] rw [containsKey_filter_contains_eq_false_iff, ← List.containsKey_eq_contains_map_fst] · exact hl₁ -/- `get_diff` -/ - theorem getValueCast_filter_contains_eq_false [BEq α] [LawfulBEq α] {l₁ : List ((a : α) × β a)} {l₂ : List α} {k : α} (dl₁ : DistinctKeys l₁) {h₁ h₂} : @@ -5773,7 +5764,6 @@ theorem congr_filter_containsKey_of_perm {l₁ l₂ l₃ l₄ : List ((a : α) · exact h₁ · apply perm_filter_containsKey_of_perm h₂ hd -/- `contains_diff_eq_false_of_contains_right` -/ theorem containsKey_filter_contains_eq_false_of_contains_right [BEq α] [EquivBEq α] {l₁ : List ((a : α) × β a)} {l₂ : List α} {hl₁ : DistinctKeys l₁} {k : α} : List.contains l₂ k → containsKey k (List.filter (fun p => !List.contains l₂ p.fst) l₁) = false := by intro h @@ -5816,17 +5806,6 @@ if containsKey k l₂ = true then fallback else getValueCastD k l₁ fallback := rw [getValueCastD_filter_contains_eq_false, containsKey_eq_contains_map_fst] · exact dl₁ -theorem getValueCastD_diff_of_containsKey_eq_false_left [BEq α] [LawfulBEq α] - {l₁ l₂ : List ((a : α) × β a)} {k : α} {fallback} - (dl₁ : DistinctKeys l₁) : - containsKey k l₁ = false → getValueCastD k (List.filter (fun p => !containsKey p.fst l₂) l₁) fallback = fallback := by - sorry - -- intro h - -- apply getValueCastD_eq_fallback - -- apply containsKey_diff_eq_false_of_containsKey_eq_false_left - -- · exact dl₁ - -- · exact h - theorem getKey?_filter_contains_eq_false [BEq α] [EquivBEq α] {l₁ : List ((a : α) × β a)} {l₂ : List α} {k : α} (dl₁ : DistinctKeys l₁) : @@ -6344,7 +6323,7 @@ theorem getValueD_filter_contains_map_fst_eq_false_of_contains_map_fst_right {β · exact dl₁ · exact hyp -theorem getValueD_diff_of_containsKey_eq_false_left {β : Type v} [BEq α] [EquivBEq α] +theorem getValueD_filter_contains_eq_false_of_containsKey_eq_false_left {β : Type v} [BEq α] [EquivBEq α] {l₁ : List ((_ : α) × β)} {l₂ : List α} {k : α} {fallback : β} (dl₁ : DistinctKeys l₁) : containsKey k l₁ = false → getValueD k (List.filter (fun p => !List.contains l₂ p.fst) l₁) fallback = fallback := by @@ -6358,7 +6337,7 @@ theorem getValueD_filter_contains_map_fst_eq_false_of_containsKey_eq_false_left {l₁ l₂ : List ((_ : α) × β)} {k : α} {fallback : β} (dl₁ : DistinctKeys l₁) : containsKey k l₁ = false → getValueD k (List.filter (fun p => !List.contains (l₂.map Sigma.fst) p.fst) l₁) fallback = fallback := by - apply getValueD_diff_of_containsKey_eq_false_left dl₁ + apply getValueD_filter_contains_eq_false_of_containsKey_eq_false_left dl₁ theorem getValueCastD_filter_contains_eq_false_of_contains_right {β : α → Type v} [BEq α] [LawfulBEq α] {l₁ : List ((a : α) × β a)} {l₂ : List α} {k : α} {fallback : β k} From 318b54ff50377c26ad9f236862a06b3635b995cc Mon Sep 17 00:00:00 2001 From: Wojciech Rozowski Date: Fri, 21 Nov 2025 17:55:24 +0000 Subject: [PATCH 84/88] Fix typo --- src/Std/Data/DHashMap/Internal/RawLemmas.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/src/Std/Data/DHashMap/Internal/RawLemmas.lean b/src/Std/Data/DHashMap/Internal/RawLemmas.lean index 4d125454c5b6..f8280079405b 100644 --- a/src/Std/Data/DHashMap/Internal/RawLemmas.lean +++ b/src/Std/Data/DHashMap/Internal/RawLemmas.lean @@ -3615,7 +3615,7 @@ theorem getD_diff_of_contains_eq_false_left [EquivBEq α] [LawfulHashable α] (h {k : α} {fallback : β} (h : m₁.contains k = false) : Const.getD (m₁.diff m₂) k fallback = fallback := by revert h - simp_to_model [diff, Const.getD, contains] using List. + simp_to_model [diff, Const.getD, contains] using List.getValueD_filter_contains_map_fst_eq_false_of_containsKey_eq_false_left /- get! -/ theorem get!_diff [EquivBEq α] [LawfulHashable α] [Inhabited β] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} : From eb88d10d5ca8d7166707f3decbfe68794745c5b3 Mon Sep 17 00:00:00 2001 From: Wojciech Rozowski Date: Thu, 27 Nov 2025 11:51:02 +0000 Subject: [PATCH 85/88] finish renaming --- src/Std/Data/DHashMap/Internal/RawLemmas.lean | 94 +++--- src/Std/Data/DHashMap/Internal/WF.lean | 2 +- src/Std/Data/Internal/List/Associative.lean | 282 +++++++++--------- 3 files changed, 189 insertions(+), 189 deletions(-) diff --git a/src/Std/Data/DHashMap/Internal/RawLemmas.lean b/src/Std/Data/DHashMap/Internal/RawLemmas.lean index f8280079405b..559ddbe6f00d 100644 --- a/src/Std/Data/DHashMap/Internal/RawLemmas.lean +++ b/src/Std/Data/DHashMap/Internal/RawLemmas.lean @@ -3346,185 +3346,185 @@ variable {m₁ m₂ : Raw₀ α β} theorem contains_diff [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} : (m₁.diff m₂).contains k = (m₁.contains k && !m₂.contains k) := by - simp_to_model [diff, contains] using List.containsKey_filter_contains_map_fst_eq_false + simp_to_model [diff, contains] using List.containsKey_filter_not_contains_map_fst theorem contains_diff_iff [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} : (m₁.diff m₂).contains k ↔ m₁.contains k ∧ ¬m₂.contains k := by - simp_to_model [diff, contains] using List.containsKey_filter_contains_map_fst_eq_false_iff + simp_to_model [diff, contains] using List.containsKey_filter_not_contains_map_fst_iff theorem contains_diff_eq_false_of_contains_eq_false_left [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} (h : m₁.contains k = false) : (m₁.diff m₂).contains k = false := by revert h - simp_to_model [diff, contains] using List.containsKey_filter_contains_map_fst_eq_false_eq_false_left + simp_to_model [diff, contains] using List.containsKey_filter_not_contains_map_fst_eq_false_left theorem contains_diff_eq_false_of_contains_right [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} (h : m₂.contains k) : (m₁.diff m₂).contains k = false := by revert h - simp_to_model [diff, contains] using List.containsKey_filter_contains_map_fst_eq_false_of_contains_map_fst_right + simp_to_model [diff, contains] using List.containsKey_filter_not_contains_map_fst_of_contains_map_fst_right /- get? -/ theorem get?_diff [LawfulBEq α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} : (m₁.diff m₂).get? k = if m₂.contains k then none else m₁.get? k := by - simp_to_model [diff, get?, contains] using List.getValueCast?_filter_contains_map_fst_eq_false + simp_to_model [diff, get?, contains] using List.getValueCast?_filter_not_contains_map_fst theorem get?_diff_of_contains_eq_false_right [LawfulBEq α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} (h : m₂.contains k = false) : (m₁.diff m₂).get? k = m₁.get? k := by revert h - simp_to_model [diff, contains, get?] using List.getValueCast?_filter_contains_map_fst_eq_false_of_containsKey_eq_false_right + simp_to_model [diff, contains, get?] using List.getValueCast?_filter_not_contains_map_fst_of_containsKey_eq_false_right theorem get?_diff_of_contains_eq_false_left [LawfulBEq α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} (h : m₁.contains k = false) : (m₁.diff m₂).get? k = none := by revert h - simp_to_model [diff, contains, get?] using List.getValueCast?_filter_contains_map_fst_eq_false_of_containsKey_eq_false_left + simp_to_model [diff, contains, get?] using List.getValueCast?_filter_not_contains_map_fst_of_containsKey_eq_false_left theorem get?_diff_of_contains_right [LawfulBEq α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} (h : m₂.contains k) : (m₁.diff m₂).get? k = none := by revert h - simp_to_model [diff, get?, contains] using List.getValueCast?_filter_contains_map_fst_eq_false_of_containsKey_right + simp_to_model [diff, get?, contains] using List.getValueCast?_filter_not_contains_map_fst_of_containsKey_right /- get -/ theorem get_diff [LawfulBEq α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} {h_contains : (m₁.diff m₂).contains k} : (m₁.diff m₂).get k h_contains = m₁.get k ((contains_diff_iff h₁ h₂).1 h_contains).1 := by - simp_to_model [diff, get, contains] using List.getValueCast_filter_contains_map_fst_eq_false + simp_to_model [diff, get, contains] using List.getValueCast_filter_not_contains_map_fst /- getD -/ theorem getD_diff [LawfulBEq α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} {fallback : β k} : (m₁.diff m₂).getD k fallback = if m₂.contains k then fallback else m₁.getD k fallback := by - simp_to_model [diff, getD, contains] using List.getValueCastD_filter_contains_map_fst_eq_false + simp_to_model [diff, getD, contains] using List.getValueCastD_filter_not_contains_map_fst theorem getD_diff_of_contains_eq_false_right [LawfulBEq α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} {fallback : β k} (h : m₂.contains k = false) : (m₁.diff m₂).getD k fallback = m₁.getD k fallback := by revert h - simp_to_model [diff, contains, getD] using List.getValueCastD_filter_contains_map_fst_eq_false_of_containsKey_eq_false_right + simp_to_model [diff, contains, getD] using List.getValueCastD_filter_not_contains_map_fst_of_containsKey_eq_false_right theorem getD_diff_of_contains_right [LawfulBEq α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} {fallback : β k} (h : m₂.contains k) : (m₁.diff m₂).getD k fallback = fallback := by revert h - simp_to_model [diff, getD, contains] using List.getValueCastD_filter_contains_map_fst_eq_false_of_contains_map_fst_right + simp_to_model [diff, getD, contains] using List.getValueCastD_filter_not_contains_map_fst_of_contains_map_fst_right theorem getD_diff_of_contains_eq_false_left [LawfulBEq α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} {fallback : β k} (h : m₁.contains k = false) : (m₁.diff m₂).getD k fallback = fallback := by revert h - simp_to_model [diff, getD, contains] using List.getValueCastD_filter_contains_map_fst_eq_false_of_containsKey_eq_false_left + simp_to_model [diff, getD, contains] using List.getValueCastD_filter_not_contains_map_fst_of_containsKey_eq_false_left /- get! -/ theorem get!_diff [LawfulBEq α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} [Inhabited (β k)] : (m₁.diff m₂).get! k = if m₂.contains k then default else m₁.get! k := by - simp_to_model [diff, get!, contains] using List.getValueCastD_filter_contains_map_fst_eq_false + simp_to_model [diff, get!, contains] using List.getValueCastD_filter_not_contains_map_fst theorem get!_diff_of_contains_eq_false_right [LawfulBEq α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} [Inhabited (β k)] (h : m₂.contains k = false) : (m₁.diff m₂).get! k = m₁.get! k := by revert h - simp_to_model [diff, contains, get!] using List.getValueCastD_filter_contains_map_fst_eq_false_of_containsKey_eq_false_right + simp_to_model [diff, contains, get!] using List.getValueCastD_filter_not_contains_map_fst_of_containsKey_eq_false_right theorem get!_diff_of_contains_right [LawfulBEq α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} [Inhabited (β k)] (h : m₂.contains k) : (m₁.diff m₂).get! k = default := by revert h - simp_to_model [diff, get!, contains] using List.getValueCastD_filter_contains_map_fst_eq_false_of_contains_map_fst_right + simp_to_model [diff, get!, contains] using List.getValueCastD_filter_not_contains_map_fst_of_contains_map_fst_right theorem get!_diff_of_contains_eq_false_left [LawfulBEq α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} [Inhabited (β k)] (h : m₁.contains k = false) : (m₁.diff m₂).get! k = default := by revert h - simp_to_model [diff, get!, contains] using List.getValueCastD_filter_contains_map_fst_eq_false_of_containsKey_eq_false_left + simp_to_model [diff, get!, contains] using List.getValueCastD_filter_not_contains_map_fst_of_containsKey_eq_false_left /- getKey? -/ theorem getKey?_diff [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} : (m₁.diff m₂).getKey? k = if m₂.contains k then none else m₁.getKey? k := by - simp_to_model [diff, contains, getKey?] using List.getKey?_filter_contains_map_fst_eq_false + simp_to_model [diff, contains, getKey?] using List.getKey?_filter_not_contains_map_fst theorem getKey?_diff_of_contains_eq_false_right [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} (h : m₂.contains k = false) : (m₁.diff m₂).getKey? k = m₁.getKey? k := by revert h - simp_to_model [contains, getKey?, diff] using List.getKey?_filter_contains_map_fst_eq_false_of_containsKey_eq_false_right + simp_to_model [contains, getKey?, diff] using List.getKey?_filter_not_contains_map_fst_of_containsKey_eq_false_right theorem getKey?_diff_of_contains_eq_false_left [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} (h : m₁.contains k = false) : (m₁.diff m₂).getKey? k = none := by revert h - simp_to_model [contains, getKey?, diff] using List.getKey?_filter_contains_map_fst_eq_false_of_containsKey_eq_false_left + simp_to_model [contains, getKey?, diff] using List.getKey?_filter_not_contains_map_fst_of_containsKey_eq_false_left theorem getKey?_diff_of_contains_right [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} (h : m₂.contains k) : (m₁.diff m₂).getKey? k = none := by revert h - simp_to_model [contains, getKey?, diff] using List.getKey?_filter_contains_map_fst_eq_false_of_containsKey_right + simp_to_model [contains, getKey?, diff] using List.getKey?_filter_not_contains_map_fst_of_containsKey_right /- getKey -/ theorem getKey_diff [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} {h_contains : (m₁.diff m₂).contains k} : (m₁.diff m₂).getKey k h_contains = m₁.getKey k ((contains_diff_iff h₁ h₂).1 h_contains).1 := by - simp_to_model [diff, contains, getKey] using List.getKey_filter_contains_map_fst_eq_false + simp_to_model [diff, contains, getKey] using List.getKey_filter_not_contains_map_fst /- getKeyD -/ theorem getKeyD_diff [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k fallback : α} : (m₁.diff m₂).getKeyD k fallback = if m₂.contains k then fallback else m₁.getKeyD k fallback := by - simp_to_model [diff, getKeyD, contains] using List.getKeyD_filter_contains_map_fst_eq_false + simp_to_model [diff, getKeyD, contains] using List.getKeyD_filter_not_contains_map_fst theorem getKeyD_diff_of_contains_eq_false_right [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k fallback : α} (h : m₂.contains k = false) : (m₁.diff m₂).getKeyD k fallback = m₁.getKeyD k fallback := by revert h - simp_to_model [contains, diff, getKeyD] using List.getKeyD_filter_contains_map_fst_eq_false_of_contains_eq_false_right + simp_to_model [contains, diff, getKeyD] using List.getKeyD_filter_not_contains_map_fst_of_contains_eq_false_right theorem getKeyD_diff_of_contains_right [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k fallback : α} (h : m₂.contains k) : (m₁.diff m₂).getKeyD k fallback = fallback := by revert h - simp_to_model [diff, getKeyD, contains] using List.getKeyD_filter_contains_map_fst_eq_false_of_contains_right + simp_to_model [diff, getKeyD, contains] using List.getKeyD_filter_not_contains_map_fst_of_contains_right theorem getKeyD_diff_of_contains_eq_false_left [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k fallback : α} (h : m₁.contains k = false) : (m₁.diff m₂).getKeyD k fallback = fallback := by revert h - simp_to_model [diff, getKeyD, contains] using getKeyD_filter_contains_map_fst_eq_false_of_contains_eq_false_left + simp_to_model [diff, getKeyD, contains] using getKeyD_filter_not_contains_map_fst_of_contains_eq_false_left /- getKey! -/ theorem getKey!_diff [EquivBEq α] [LawfulHashable α] [Inhabited α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} : (m₁.diff m₂).getKey! k = if m₂.contains k then default else m₁.getKey! k := by - simp_to_model [diff, getKey!, contains] using List.getKeyD_filter_contains_map_fst_eq_false + simp_to_model [diff, getKey!, contains] using List.getKeyD_filter_not_contains_map_fst theorem getKey!_diff_of_contains_eq_false_right [EquivBEq α] [LawfulHashable α] [Inhabited α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} (h : m₂.contains k = false) : (m₁.diff m₂).getKey! k = m₁.getKey! k := by revert h - simp_to_model [diff, getKey!, contains] using List.getKeyD_filter_contains_map_fst_eq_false_of_contains_eq_false_right + simp_to_model [diff, getKey!, contains] using List.getKeyD_filter_not_contains_map_fst_of_contains_eq_false_right theorem getKey!_diff_of_contains_right [EquivBEq α] [LawfulHashable α] [Inhabited α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} (h : m₂.contains k) : (m₁.diff m₂).getKey! k = default := by revert h - simp_to_model [diff, getKey!, contains] using List.getKeyD_filter_contains_map_fst_eq_false_of_contains_right + simp_to_model [diff, getKey!, contains] using List.getKeyD_filter_not_contains_map_fst_of_contains_right theorem getKey!_diff_of_contains_eq_false_left [EquivBEq α] [LawfulHashable α] [Inhabited α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} (h : m₁.contains k = false) : (m₁.diff m₂).getKey! k = default := by revert h - simp_to_model [diff, getKey!, contains] using List.getKeyD_filter_contains_map_fst_eq_false_of_contains_eq_false_left + simp_to_model [diff, getKey!, contains] using List.getKeyD_filter_not_contains_map_fst_of_contains_eq_false_left /- size -/ theorem size_diff_le_size_left [EquivBEq α] [LawfulHashable α] @@ -3537,12 +3537,12 @@ theorem size_diff_eq_size_left [EquivBEq α] [LawfulHashable α] (h : ∀ (a : α), m₁.contains a → m₂.contains a = false) : (m₁.diff m₂).1.size = m₁.1.size := by revert h - simp_to_model [diff, size, contains] using List.length_filter_contains_map_fst_eq_false_eq_length_left + simp_to_model [diff, size, contains] using List.length_filter_not_contains_map_fst_eq_length_left theorem size_diff_add_size_inter_eq_size_left [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) : (m₁.diff m₂).1.size + (m₁.inter m₂).1.size = m₁.1.size := by - simp_to_model [diff, inter, size] using List.size_filter_contains_map_fst_eq_false_add_size_inter_eq_size_left + simp_to_model [diff, inter, size] using List.size_filter_not_contains_map_fst_add_size_inter_eq_size_left /- isEmpty -/ @[simp] @@ -3550,11 +3550,11 @@ theorem isEmpty_diff_left [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) (h : m₁.1.isEmpty) : (m₁.diff m₂).1.isEmpty = true := by revert h - simp_to_model [isEmpty, diff] using List.isEmpty_filter_contains_map_fst_eq_false_left + simp_to_model [isEmpty, diff] using List.isEmpty_filter_not_contains_map_fst_left theorem isEmpty_diff_iff [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) : (m₁.diff m₂).1.isEmpty ↔ ∀ k, m₁.contains k → m₂.contains k := by - simp_to_model [diff, contains, isEmpty] using List.isEmpty_filter_contains_map_fst_eq_false_iff + simp_to_model [diff, contains, isEmpty] using List.isEmpty_filter_not_contains_map_fst_iff end Diff @@ -3565,80 +3565,80 @@ variable {β : Type v} {m₁ m₂ : Raw₀ α (fun _ => β)} /- get? -/ theorem get?_diff [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} : Const.get? (m₁.diff m₂) k = if m₂.contains k then none else Const.get? m₁ k := by - simp_to_model [diff, Const.get?, contains] using List.getValue?_filter_contains_map_fst_eq_false + simp_to_model [diff, Const.get?, contains] using List.getValue?_filter_not_contains_map_fst theorem get?_diff_of_contains_eq_false_right [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} (h : m₂.contains k = false) : Const.get? (m₁.diff m₂) k = Const.get? m₁ k := by revert h - simp_to_model [diff, contains, Const.get?] using List.getValue?_filter_contains_map_fst_eq_false_of_contains_eq_false_right + simp_to_model [diff, contains, Const.get?] using List.getValue?_filter_not_contains_map_fst_of_contains_eq_false_right theorem get?_diff_of_contains_eq_false_left [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} (h : m₁.contains k = false) : Const.get? (m₁.diff m₂) k = none := by revert h - simp_to_model [diff, Const.get?, contains] using List.getValue?_filter_contains_map_fst_eq_false_of_containsKey_eq_false_left + simp_to_model [diff, Const.get?, contains] using List.getValue?_filter_not_contains_map_fst_of_containsKey_eq_false_left theorem get?_diff_of_contains_right [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} (h : m₂.contains k) : Const.get? (m₁.diff m₂) k = none := by revert h - simp_to_model [diff, Const.get?, contains] using List.getValue?_filter_contains_map_fst_eq_false_of_contains_right + simp_to_model [diff, Const.get?, contains] using List.getValue?_filter_not_contains_map_fst_of_contains_right /- get -/ theorem get_diff [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} {h_contains : (m₁.diff m₂).contains k} : Const.get (m₁.diff m₂) k h_contains = Const.get m₁ k ((contains_diff_iff h₁ h₂).1 h_contains).1 := by - simp_to_model [diff, Const.get, contains] using List.getValue_filter_contains_eq_false + simp_to_model [diff, Const.get, contains] using List.getValue_filter_not_contains /- getD -/ theorem getD_diff [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} {fallback : β} : Const.getD (m₁.diff m₂) k fallback = if m₂.contains k then fallback else Const.getD m₁ k fallback := by - simp_to_model [diff, Const.getD, contains] using List.getValueD_filter_contains_map_fst_eq_false + simp_to_model [diff, Const.getD, contains] using List.getValueD_filter_not_contains_map_fst theorem getD_diff_of_contains_eq_false_right [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} {fallback : β} (h : m₂.contains k = false) : Const.getD (m₁.diff m₂) k fallback = Const.getD m₁ k fallback := by revert h - simp_to_model [diff, contains, Const.getD] using List.getValueD_filter_contains_map_fst_eq_false_of_containsKey_eq_false_right + simp_to_model [diff, contains, Const.getD] using List.getValueD_filter_not_contains_map_fst_of_containsKey_eq_false_right theorem getD_diff_of_contains_right [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} {fallback : β} (h : m₂.contains k) : Const.getD (m₁.diff m₂) k fallback = fallback := by revert h - simp_to_model [diff, Const.getD, contains] using List.getValueD_filter_contains_map_fst_eq_false_of_contains_map_fst_right + simp_to_model [diff, Const.getD, contains] using List.getValueD_filter_not_contains_map_fst_of_contains_map_fst_right theorem getD_diff_of_contains_eq_false_left [EquivBEq α] [LawfulHashable α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} {fallback : β} (h : m₁.contains k = false) : Const.getD (m₁.diff m₂) k fallback = fallback := by revert h - simp_to_model [diff, Const.getD, contains] using List.getValueD_filter_contains_map_fst_eq_false_of_containsKey_eq_false_left + simp_to_model [diff, Const.getD, contains] using List.getValueD_filter_not_contains_map_fst_of_containsKey_eq_false_left /- get! -/ theorem get!_diff [EquivBEq α] [LawfulHashable α] [Inhabited β] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} : Const.get! (m₁.diff m₂) k = if m₂.contains k then default else Const.get! m₁ k := by - simp_to_model [diff, Const.get!, contains] using List.getValueD_filter_contains_map_fst_eq_false + simp_to_model [diff, Const.get!, contains] using List.getValueD_filter_not_contains_map_fst theorem get!_diff_of_contains_eq_false_right [EquivBEq α] [LawfulHashable α] [Inhabited β] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} (h : m₂.contains k = false) : Const.get! (m₁.diff m₂) k = Const.get! m₁ k := by revert h - simp_to_model [diff, contains, Const.get!] using List.getValueD_filter_contains_map_fst_eq_false_of_containsKey_eq_false_right + simp_to_model [diff, contains, Const.get!] using List.getValueD_filter_not_contains_map_fst_of_containsKey_eq_false_right theorem get!_diff_of_contains_right [EquivBEq α] [LawfulHashable α] [Inhabited β] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} (h : m₂.contains k) : Const.get! (m₁.diff m₂) k = default := by revert h - simp_to_model [diff, Const.get!, contains] using List.getValueD_filter_contains_map_fst_eq_false_of_contains_map_fst_right + simp_to_model [diff, Const.get!, contains] using List.getValueD_filter_not_contains_map_fst_of_contains_map_fst_right theorem get!_diff_of_contains_eq_false_left [EquivBEq α] [LawfulHashable α] [Inhabited β] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} (h : m₁.contains k = false) : Const.get! (m₁.diff m₂) k = default := by revert h - simp_to_model [diff, Const.get!, contains] using List.getValueD_filter_contains_map_fst_eq_false_of_containsKey_eq_false_left + simp_to_model [diff, Const.get!, contains] using List.getValueD_filter_not_contains_map_fst_of_containsKey_eq_false_left end Const diff --git a/src/Std/Data/DHashMap/Internal/WF.lean b/src/Std/Data/DHashMap/Internal/WF.lean index 3d96c1c07d1f..53d65a602acb 100644 --- a/src/Std/Data/DHashMap/Internal/WF.lean +++ b/src/Std/Data/DHashMap/Internal/WF.lean @@ -1370,7 +1370,7 @@ theorem toListModel_diffₘ [BEq α] [Hashable α] [EquivBEq α] [LawfulHashable rw [containsKey_eq_contains_map_fst] · rw [eraseManyEntries_eq_eraseListₘ] apply Perm.trans (toListModel_eraseListₘ h₁) - · apply eraseList_perm_filter_contains_eq_false + · apply eraseList_perm_filter_not_contains · exact h₁.distinct theorem diff_eq_diffₘ [BEq α] [Hashable α] (m₁ m₂ : Raw₀ α β) : diff --git a/src/Std/Data/Internal/List/Associative.lean b/src/Std/Data/Internal/List/Associative.lean index 21f2a2283384..4402d138dd5b 100644 --- a/src/Std/Data/Internal/List/Associative.lean +++ b/src/Std/Data/Internal/List/Associative.lean @@ -5233,7 +5233,7 @@ theorem getValueCast?_filter_containsKey [BEq α] [LawfulBEq α] case h_2 => rfl exact dl₁ -theorem getValueCast?_filter_contains_eq_false [BEq α] [LawfulBEq α] +theorem getValueCast?_filter_not_contains [BEq α] [LawfulBEq α] {l₁ : List ((a : α) × β a)} {l₂ : List α} {k : α} (dl₁ : DistinctKeys l₁) : getValueCast? k (List.filter (fun p => !List.contains l₂ p.fst) l₁) = @@ -5280,73 +5280,73 @@ theorem getValueCast?_filter_contains_eq_false [BEq α] [LawfulBEq α] simp [heq] · exact dl₁ -theorem getValueCast?_filter_contains_map_fst_eq_false [BEq α] [LawfulBEq α] +theorem getValueCast?_filter_not_contains_map_fst [BEq α] [LawfulBEq α] {l₁ l₂ : List ((a : α) × β a)} {k : α} (dl₁ : DistinctKeys l₁) : getValueCast? k (List.filter (fun p => !List.contains (l₂.map Sigma.fst) p.fst) l₁) = if containsKey k l₂ = true then none else getValueCast? k l₁ := by - rw [getValueCast?_filter_contains_eq_false, ← containsKey_eq_contains_map_fst] + rw [getValueCast?_filter_not_contains, ← containsKey_eq_contains_map_fst] exact dl₁ -theorem getValueCast?_filter_contains_eq_false_of_contains_eq_false_right [BEq α] [LawfulBEq α] +theorem getValueCast?_filter_not_contains_of_contains_eq_false_right [BEq α] [LawfulBEq α] {l₁ : List ((a : α) × β a)} {l₂ : List α} {k : α} (dl₁ : DistinctKeys l₁) : List.contains l₂ k = false → getValueCast? k (List.filter (fun p => !List.contains l₂ p.fst) l₁) = getValueCast? k l₁ := by intro h₁ - rw [getValueCast?_filter_contains_eq_false, h₁] + rw [getValueCast?_filter_not_contains, h₁] all_goals simp [dl₁] -theorem getValueCast?_filter_contains_map_fst_eq_false_of_containsKey_eq_false_right [BEq α] [LawfulBEq α] +theorem getValueCast?_filter_not_contains_map_fst_of_containsKey_eq_false_right [BEq α] [LawfulBEq α] {l₁ l₂ : List ((a : α) × β a)} {k : α} (dl₁ : DistinctKeys l₁) : containsKey k l₂ = false → getValueCast? k (List.filter (fun p => !List.contains (l₂.map Sigma.fst) p.fst) l₁) = getValueCast? k l₁ := by intro h₁ rw [containsKey_eq_contains_map_fst] at h₁ - rw [getValueCast?_filter_contains_eq_false_of_contains_eq_false_right dl₁] + rw [getValueCast?_filter_not_contains_of_contains_eq_false_right dl₁] · exact h₁ -theorem getValueCast?_filter_contains_eq_false_of_contains_eq_false_left [BEq α] [LawfulBEq α] +theorem getValueCast?_filter_not_contains_of_contains_eq_false_left [BEq α] [LawfulBEq α] {l₁ : List ((a : α) × β a)} {l₂ : List α} {k : α} (dl₁ : DistinctKeys l₁) : containsKey k l₁ = false → getValueCast? k (List.filter (fun p => !List.contains l₂ p.fst) l₁) = none := by intro h₁ - rw [getValueCast?_filter_contains_eq_false] + rw [getValueCast?_filter_not_contains] · simp [getValueCast?_eq_none h₁] · exact dl₁ -theorem getValueCast?_filter_contains_map_fst_eq_false_of_containsKey_eq_false_left [BEq α] [LawfulBEq α] +theorem getValueCast?_filter_not_contains_map_fst_of_containsKey_eq_false_left [BEq α] [LawfulBEq α] {l₁ l₂ : List ((a : α) × β a)} {k : α} (dl₁ : DistinctKeys l₁) : containsKey k l₁ = false → getValueCast? k (List.filter (fun p => !List.contains (l₂.map Sigma.fst) p.fst) l₁) = none := by intro h₁ - rw [getValueCast?_filter_contains_eq_false_of_contains_eq_false_left dl₁] + rw [getValueCast?_filter_not_contains_of_contains_eq_false_left dl₁] · exact h₁ -theorem getValueCast?_filter_contains_eq_false_of_contains_right [BEq α] [LawfulBEq α] +theorem getValueCast?_filter_not_contains_of_contains_right [BEq α] [LawfulBEq α] {l₁ : List ((a : α) × β a)} {l₂ : List α} {k : α} (dl₁ : DistinctKeys l₁) : List.contains l₂ k → getValueCast? k (List.filter (fun p => !List.contains l₂ p.fst) l₁) = none := by intro h₁ - rw [getValueCast?_filter_contains_eq_false] + rw [getValueCast?_filter_not_contains] · simp only [ite_eq_left_iff, Bool.not_eq_true] intro h₂ rw [h₂] at h₁ contradiction · exact dl₁ -theorem getValueCast?_filter_contains_map_fst_eq_false_of_containsKey_right [BEq α] [LawfulBEq α] +theorem getValueCast?_filter_not_contains_map_fst_of_containsKey_right [BEq α] [LawfulBEq α] {l₁ l₂ : List ((a : α) × β a)} {k : α} (dl₁ : DistinctKeys l₁) : containsKey k l₂ → getValueCast? k (List.filter (fun p => !List.contains (l₂.map Sigma.fst) p.fst) l₁) = none := by intro h₁ rw [containsKey_eq_contains_map_fst] at h₁ - rw [getValueCast?_filter_contains_eq_false_of_contains_right dl₁] + rw [getValueCast?_filter_not_contains_of_contains_right dl₁] · exact h₁ theorem getKey?_filter_containsKey [BEq α] [EquivBEq α] @@ -5399,7 +5399,7 @@ def eraseList [BEq α] (l : List ((a : α) × β a)) (toErase : List α) : List | .nil => l | .cons k toErase => eraseList (eraseKey k l) toErase -theorem eraseList_perm_filter_contains_eq_false [BEq α] [EquivBEq α] (l toErase : List ((a : α) × β a)) (hl : DistinctKeys l): +theorem eraseList_perm_filter_not_contains [BEq α] [EquivBEq α] (l toErase : List ((a : α) × β a)) (hl : DistinctKeys l): List.Perm (eraseList l (toErase.map Sigma.fst)) (l.filter (fun p => !List.contains (toErase.map Sigma.fst) p.fst)) := by induction toErase generalizing l with | nil => @@ -5622,7 +5622,7 @@ theorem containsKey_filter_containsKey_iff [BEq α] [EquivBEq α] {l₁ l₂ : L . simp . exact hl₁ -theorem containsKey_filter_contains_eq_false [BEq α] [EquivBEq α] {l₁ : List ((a : α) × β a)} {l₂ : List α} {hl₁ : DistinctKeys l₁} {k : α} : +theorem containsKey_filter_not_contains [BEq α] [EquivBEq α] {l₁ : List ((a : α) × β a)} {l₂ : List α} {hl₁ : DistinctKeys l₁} {k : α} : (containsKey k (List.filter (fun p => !List.contains l₂ p.fst ) l₁)) = (containsKey k l₁ && !List.contains l₂ k) := by rw [containsKey_eq_isSome_getEntry?, containsKey_eq_isSome_getEntry?] @@ -5637,42 +5637,42 @@ theorem containsKey_filter_contains_eq_false [BEq α] [EquivBEq α] {l₁ : List apply List.beq_of_getEntry?_eq_some heq . exact hl₁ -theorem containsKey_filter_contains_map_fst_eq_false [BEq α] [EquivBEq α] {l₁ l₂ : List ((a : α) × β a)} {hl₁ : DistinctKeys l₁} {k : α} : +theorem containsKey_filter_not_contains_map_fst [BEq α] [EquivBEq α] {l₁ l₂ : List ((a : α) × β a)} {hl₁ : DistinctKeys l₁} {k : α} : (containsKey k (List.filter (fun p => !List.contains (l₂.map Sigma.fst) p.fst ) l₁)) = (containsKey k l₁ && !containsKey k l₂) := by - rw [containsKey_filter_contains_eq_false, ← containsKey_eq_contains_map_fst] + rw [containsKey_filter_not_contains, ← containsKey_eq_contains_map_fst] · exact hl₁ -theorem containsKey_filter_contains_eq_false_iff [BEq α] [EquivBEq α] {l₁ : List ((a : α) × β a)} {l₂ : List α} {hl₁ : DistinctKeys l₁} {k : α} : +theorem containsKey_filter_not_contains_iff [BEq α] [EquivBEq α] {l₁ : List ((a : α) × β a)} {l₂ : List α} {hl₁ : DistinctKeys l₁} {k : α} : (containsKey k (List.filter (fun p => !List.contains l₂ p.fst) l₁)) ↔ (containsKey k l₁ ∧ ¬ List.contains l₂ k) := by - rw [containsKey_filter_contains_eq_false] + rw [containsKey_filter_not_contains] . simp . exact hl₁ -theorem containsKey_filter_contains_map_fst_eq_false_iff [BEq α] [EquivBEq α] {l₁ l₂ : List ((a : α) × β a)} {hl₁ : DistinctKeys l₁} {k : α} : +theorem containsKey_filter_not_contains_map_fst_iff [BEq α] [EquivBEq α] {l₁ l₂ : List ((a : α) × β a)} {hl₁ : DistinctKeys l₁} {k : α} : (containsKey k (List.filter (fun p => !List.contains (l₂.map Sigma.fst) p.fst) l₁)) ↔ (containsKey k l₁ ∧ ¬ containsKey k l₂) := by - rw [containsKey_filter_contains_eq_false_iff, ← List.containsKey_eq_contains_map_fst] + rw [containsKey_filter_not_contains_iff, ← List.containsKey_eq_contains_map_fst] · exact hl₁ -theorem getValueCast_filter_contains_eq_false [BEq α] [LawfulBEq α] +theorem getValueCast_filter_not_contains [BEq α] [LawfulBEq α] {l₁ : List ((a : α) × β a)} {l₂ : List α} {k : α} (dl₁ : DistinctKeys l₁) {h₁ h₂} : getValueCast k (List.filter (fun p => !List.contains l₂ p.fst) l₁) h₁ = getValueCast k l₁ h₂ := by suffices some (getValueCast k (List.filter (fun p => !List.contains l₂ p.fst) l₁) h₁) = some (getValueCast k l₁ h₂) by injections simp only [← getValueCast?_eq_some_getValueCast] - rw [getValueCast?_filter_contains_eq_false] - rw [containsKey_filter_contains_eq_false_iff] at h₁ + rw [getValueCast?_filter_not_contains] + rw [containsKey_filter_not_contains_iff] at h₁ simp only [h₁.2, Bool.false_eq_true, ↓reduceIte] all_goals exact dl₁ -theorem getValueCast_filter_contains_map_fst_eq_false [BEq α] [LawfulBEq α] +theorem getValueCast_filter_not_contains_map_fst [BEq α] [LawfulBEq α] {l₁ l₂ : List ((a : α) × β a)} {k : α} (dl₁ : DistinctKeys l₁) {h₁ h₂} : getValueCast k (List.filter (fun p => !List.contains (l₂.map Sigma.fst) p.fst) l₁) h₁ = getValueCast k l₁ h₂ := by suffices some (getValueCast k (List.filter (fun p => !List.contains (l₂.map Sigma.fst) p.fst) l₁) h₁) = some (getValueCast k l₁ h₂) by injections - rw [getValueCast_filter_contains_eq_false] + rw [getValueCast_filter_not_contains] · exact dl₁ theorem getValueCast_filter_containsKey [BEq α] [LawfulBEq α] @@ -5717,17 +5717,17 @@ theorem containsKey_filter_containsKey_eq_false_of_containsKey_eq_false_left [BE . simp [h] . exact hl₁ -theorem containsKey_filter_contains_eq_false_eq_false_left [BEq α] [EquivBEq α] {l₁: List ((a : α) × β a)} {l₂ : List α} {hl₁ : DistinctKeys l₁} {k : α} : +theorem containsKey_filter_not_contains_eq_false_left [BEq α] [EquivBEq α] {l₁: List ((a : α) × β a)} {l₂ : List α} {hl₁ : DistinctKeys l₁} {k : α} : containsKey k l₁ = false → containsKey k (List.filter (fun p => !List.contains l₂ p.fst) l₁) = false := by intro h - rw [containsKey_filter_contains_eq_false] + rw [containsKey_filter_not_contains] . simp [h] . exact hl₁ -theorem containsKey_filter_contains_map_fst_eq_false_eq_false_left [BEq α] [EquivBEq α] {l₁ l₂ : List ((a : α) × β a)} {hl₁ : DistinctKeys l₁} {k : α} : +theorem containsKey_filter_not_contains_map_fst_eq_false_left [BEq α] [EquivBEq α] {l₁ l₂ : List ((a : α) × β a)} {hl₁ : DistinctKeys l₁} {k : α} : containsKey k l₁ = false → containsKey k (List.filter (fun p => !List.contains (l₂.map Sigma.fst) p.fst) l₁) = false := by intro h - rw [containsKey_filter_contains_eq_false_eq_false_left] + rw [containsKey_filter_not_contains_eq_false_left] · exact h · exact hl₁ @@ -5764,22 +5764,22 @@ theorem congr_filter_containsKey_of_perm {l₁ l₂ l₃ l₄ : List ((a : α) · exact h₁ · apply perm_filter_containsKey_of_perm h₂ hd -theorem containsKey_filter_contains_eq_false_of_contains_right [BEq α] [EquivBEq α] {l₁ : List ((a : α) × β a)} {l₂ : List α} {hl₁ : DistinctKeys l₁} {k : α} : +theorem containsKey_filter_not_contains_of_contains_right [BEq α] [EquivBEq α] {l₁ : List ((a : α) × β a)} {l₂ : List α} {hl₁ : DistinctKeys l₁} {k : α} : List.contains l₂ k → containsKey k (List.filter (fun p => !List.contains l₂ p.fst) l₁) = false := by intro h - rw [containsKey_filter_contains_eq_false] + rw [containsKey_filter_not_contains] . simp [h] . exact hl₁ -theorem containsKey_filter_contains_map_fst_eq_false_of_contains_map_fst_right [BEq α] [EquivBEq α] {l₁ l₂ : List ((a : α) × β a)} {hl₁ : DistinctKeys l₁} {k : α} : +theorem containsKey_filter_not_contains_map_fst_of_contains_map_fst_right [BEq α] [EquivBEq α] {l₁ l₂ : List ((a : α) × β a)} {hl₁ : DistinctKeys l₁} {k : α} : containsKey k l₂ → containsKey k (List.filter (fun p => !List.contains (l₂.map Sigma.fst) p.fst) l₁) = false := by intro h rw [containsKey_eq_contains_map_fst] at h - rw [containsKey_filter_contains_eq_false_of_contains_right] + rw [containsKey_filter_not_contains_of_contains_right] · simp [h] · exact hl₁ -theorem getValueCastD_filter_contains_eq_false [BEq α] [LawfulBEq α] +theorem getValueCastD_filter_not_contains [BEq α] [LawfulBEq α] {l₁ : List ((a : α) × β a)} {l₂ : List α} {k : α} {fallback : β k} (dl₁ : DistinctKeys l₁) : getValueCastD k (List.filter (fun p => !List.contains l₂ p.fst) l₁) fallback = @@ -5787,26 +5787,26 @@ theorem getValueCastD_filter_contains_eq_false [BEq α] [LawfulBEq α] split case isTrue h => apply getValueCastD_eq_fallback - apply containsKey_filter_contains_eq_false_of_contains_right + apply containsKey_filter_not_contains_of_contains_right · exact dl₁ · exact h case isFalse h => rw [getValueCastD_eq_getValueCast?, getValueCastD_eq_getValueCast?] congr 1 rw [Bool.not_eq_true] at h - rw [getValueCast?_filter_contains_eq_false, h] + rw [getValueCast?_filter_not_contains, h] simp [Bool.false_eq_true, ↓reduceIte] · exact dl₁ -theorem getValueCastD_filter_contains_map_fst_eq_false [BEq α] [LawfulBEq α] +theorem getValueCastD_filter_not_contains_map_fst [BEq α] [LawfulBEq α] {l₁ l₂ : List ((a : α) × β a)} {k : α} {fallback : β k} (dl₁ : DistinctKeys l₁) : getValueCastD k (List.filter (fun p => !List.contains (l₂.map Sigma.fst) p.fst) l₁) fallback = if containsKey k l₂ = true then fallback else getValueCastD k l₁ fallback := by - rw [getValueCastD_filter_contains_eq_false, containsKey_eq_contains_map_fst] + rw [getValueCastD_filter_not_contains, containsKey_eq_contains_map_fst] · exact dl₁ -theorem getKey?_filter_contains_eq_false [BEq α] [EquivBEq α] +theorem getKey?_filter_not_contains [BEq α] [EquivBEq α] {l₁ : List ((a : α) × β a)} {l₂ : List α} {k : α} (dl₁ : DistinctKeys l₁) : getKey? k (List.filter (fun p => !List.contains l₂ p.fst) l₁) = @@ -5840,93 +5840,93 @@ theorem getKey?_filter_contains_eq_false [BEq α] [EquivBEq α] simp only [heq] · exact dl₁ -theorem getKey?_filter_contains_map_fst_eq_false [BEq α] [EquivBEq α] +theorem getKey?_filter_not_contains_map_fst [BEq α] [EquivBEq α] {l₁ l₂ : List ((a : α) × β a)} {k : α} (dl₁ : DistinctKeys l₁) : getKey? k (List.filter (fun p => !List.contains (l₂.map Sigma.fst) p.fst) l₁) = if containsKey k l₂ = true then none else getKey? k l₁ := by - rw [getKey?_filter_contains_eq_false, ← List.containsKey_eq_contains_map_fst] + rw [getKey?_filter_not_contains, ← List.containsKey_eq_contains_map_fst] exact dl₁ -theorem getKey?_filter_contains_eq_false_of_contains_eq_false_right [BEq α] [EquivBEq α] +theorem getKey?_filter_not_contains_of_contains_eq_false_right [BEq α] [EquivBEq α] {l₁ : List ((a : α) × β a)} {l₂ : List α} {k : α} (dl₁ : DistinctKeys l₁) : List.contains l₂ k = false → getKey? k (List.filter (fun p => !List.contains l₂ p.fst) l₁) = getKey? k l₁ := by intro h - rw [getKey?_filter_contains_eq_false, h] + rw [getKey?_filter_not_contains, h] · simp only [Bool.false_eq_true, ↓reduceIte] · exact dl₁ -theorem getKey?_filter_contains_map_fst_eq_false_of_containsKey_eq_false_right [BEq α] [EquivBEq α] +theorem getKey?_filter_not_contains_map_fst_of_containsKey_eq_false_right [BEq α] [EquivBEq α] {l₁ l₂ : List ((a : α) × β a)} {k : α} (dl₁ : DistinctKeys l₁) : containsKey k l₂ = false → getKey? k (List.filter (fun p => !List.contains (l₂.map Sigma.fst) p.fst) l₁) = getKey? k l₁ := by intro h - rw [getKey?_filter_contains_eq_false_of_contains_eq_false_right] + rw [getKey?_filter_not_contains_of_contains_eq_false_right] · exact dl₁ · rwa [List.containsKey_eq_contains_map_fst] at h -theorem getKey?_filter_contains_eq_false_of_containsKey_eq_false_left [BEq α] [EquivBEq α] +theorem getKey?_filter_not_contains_of_containsKey_eq_false_left [BEq α] [EquivBEq α] {l₁ : List ((a : α) × β a)} {l₂ : List α} {k : α} (dl₁ : DistinctKeys l₁) : containsKey k l₁ = false → getKey? k (List.filter (fun p => !List.contains l₂ p.fst) l₁) = none := by intro h - rw [getKey?_filter_contains_eq_false] + rw [getKey?_filter_not_contains] · split · rfl · rw [getKey?_eq_none h] · exact dl₁ -theorem getKey?_filter_contains_map_fst_eq_false_of_containsKey_eq_false_left [BEq α] [EquivBEq α] +theorem getKey?_filter_not_contains_map_fst_of_containsKey_eq_false_left [BEq α] [EquivBEq α] {l₁ l₂ : List ((a : α) × β a)} {k : α} (dl₁ : DistinctKeys l₁) : containsKey k l₁ = false → getKey? k (List.filter (fun p => !List.contains (l₂.map Sigma.fst) p.fst) l₁) = none := by intro h - rw [getKey?_filter_contains_eq_false_of_containsKey_eq_false_left] + rw [getKey?_filter_not_contains_of_containsKey_eq_false_left] · exact dl₁ · exact h -theorem getKey?_filter_contains_eq_false_of_containsKey_right [BEq α] [EquivBEq α] +theorem getKey?_filter_not_contains_of_containsKey_right [BEq α] [EquivBEq α] {l₁ : List ((a : α) × β a)} {l₂ : List α} {k : α} (dl₁ : DistinctKeys l₁) : List.contains l₂ k → getKey? k (List.filter (fun p => !List.contains l₂ p.fst) l₁) = none := by intro h - rw [getKey?_filter_contains_eq_false, h] + rw [getKey?_filter_not_contains, h] · simp only [↓reduceIte] · exact dl₁ -theorem getKey?_filter_contains_map_fst_eq_false_of_containsKey_right [BEq α] [EquivBEq α] +theorem getKey?_filter_not_contains_map_fst_of_containsKey_right [BEq α] [EquivBEq α] {l₁ l₂ : List ((a : α) × β a)} {k : α} (dl₁ : DistinctKeys l₁) : containsKey k l₂ → getKey? k (List.filter (fun p => !List.contains (l₂.map Sigma.fst) p.fst) l₁) = none := by intro h - rw [getKey?_filter_contains_eq_false_of_containsKey_right] + rw [getKey?_filter_not_contains_of_containsKey_right] · exact dl₁ · rwa [List.containsKey_eq_contains_map_fst] at h -theorem getKey_filter_contains_eq_false [BEq α] [EquivBEq α] +theorem getKey_filter_not_contains [BEq α] [EquivBEq α] {l₁ : List ((a : α) × β a)} {l₂ : List α} {k : α} (dl₁ : DistinctKeys l₁) {h₁ h₂} : getKey k (List.filter (fun p => !List.contains l₂ p.fst) l₁) h₁ = getKey k l₁ h₂ := by suffices some (getKey k (List.filter (fun p => !List.contains l₂ p.fst) l₁) h₁) = some (getKey k l₁ h₂) by injections simp only [← getKey?_eq_some_getKey] - · apply getKey?_filter_contains_eq_false_of_contains_eq_false_right + · apply getKey?_filter_not_contains_of_contains_eq_false_right · exact dl₁ - · rw [containsKey_filter_contains_eq_false_iff] at h₁ + · rw [containsKey_filter_not_contains_iff] at h₁ · replace h₁ := h₁.2 rwa [Bool.not_eq_true] at h₁ · exact dl₁ -theorem getKey_filter_contains_map_fst_eq_false [BEq α] [EquivBEq α] +theorem getKey_filter_not_contains_map_fst [BEq α] [EquivBEq α] {l₁ l₂ : List ((a : α) × β a)} {k : α} (dl₁ : DistinctKeys l₁) {h₁ h₂} : getKey k (List.filter (fun p => !List.contains (l₂.map Sigma.fst) p.fst) l₁) h₁ = getKey k l₁ h₂ := by - apply getKey_filter_contains_eq_false dl₁ + apply getKey_filter_not_contains dl₁ -theorem getKeyD_filter_contains_eq_false [BEq α] [EquivBEq α] +theorem getKeyD_filter_not_contains [BEq α] [EquivBEq α] {l₁ : List ((a : α) × β a)} {l₂ : List α} {k : α} {fallback : α} (dl₁ : DistinctKeys l₁) : getKeyD k (List.filter (fun p => !List.contains l₂ p.fst) l₁) fallback = @@ -5934,24 +5934,24 @@ theorem getKeyD_filter_contains_eq_false [BEq α] [EquivBEq α] split case isTrue h => apply getKeyD_eq_fallback - apply containsKey_filter_contains_eq_false_of_contains_right + apply containsKey_filter_not_contains_of_contains_right · exact dl₁ · exact h case isFalse h => rw [getKeyD_eq_getKey?, getKeyD_eq_getKey?] congr 1 - apply getKey?_filter_contains_eq_false_of_contains_eq_false_right dl₁ + apply getKey?_filter_not_contains_of_contains_eq_false_right dl₁ simp only [h] -theorem getKeyD_filter_contains_map_fst_eq_false [BEq α] [EquivBEq α] +theorem getKeyD_filter_not_contains_map_fst [BEq α] [EquivBEq α] {l₁ l₂ : List ((a : α) × β a)} {k : α} {fallback : α} (dl₁ : DistinctKeys l₁) : getKeyD k (List.filter (fun p => !List.contains (l₂.map Sigma.fst) p.fst) l₁) fallback = if containsKey k l₂ then fallback else getKeyD k l₁ fallback := by - rw [getKeyD_filter_contains_eq_false, containsKey_eq_contains_map_fst] + rw [getKeyD_filter_not_contains, containsKey_eq_contains_map_fst] · exact dl₁ -theorem getKeyD_filter_contains_eq_false_of_contains_eq_false_right [BEq α] [EquivBEq α] +theorem getKeyD_filter_not_contains_of_contains_eq_false_right [BEq α] [EquivBEq α] {l₁ : List ((a : α) × β a)} {l₂ : List α} {k : α} {fallback : α} (dl₁ : DistinctKeys l₁) : List.contains l₂ k = false → getKeyD k (List.filter (fun p => !List.contains l₂ p.fst) l₁) fallback = @@ -5959,71 +5959,71 @@ theorem getKeyD_filter_contains_eq_false_of_contains_eq_false_right [BEq α] [Eq intro h rw [getKeyD_eq_getKey?, getKeyD_eq_getKey?] congr 1 - apply getKey?_filter_contains_eq_false_of_contains_eq_false_right dl₁ h + apply getKey?_filter_not_contains_of_contains_eq_false_right dl₁ h -theorem getKeyD_filter_contains_map_fst_eq_false_of_contains_eq_false_right [BEq α] [EquivBEq α] +theorem getKeyD_filter_not_contains_map_fst_of_contains_eq_false_right [BEq α] [EquivBEq α] {l₁ l₂ : List ((a : α) × β a)} {k : α} {fallback : α} (dl₁ : DistinctKeys l₁) : containsKey k l₂ = false → getKeyD k (List.filter (fun p => !List.contains (l₂.map Sigma.fst) p.fst) l₁) fallback = getKeyD k l₁ fallback := by intro h - rw [getKeyD_filter_contains_eq_false_of_contains_eq_false_right dl₁] + rw [getKeyD_filter_not_contains_of_contains_eq_false_right dl₁] rwa [containsKey_eq_contains_map_fst] at h -theorem getKeyD_filter_contains_eq_false_of_contains_right [BEq α] [EquivBEq α] +theorem getKeyD_filter_not_contains_of_contains_right [BEq α] [EquivBEq α] {l₁ : List ((a : α) × β a)} {l₂ : List α} {k : α} {fallback : α} (dl₁ : DistinctKeys l₁) : List.contains l₂ k → getKeyD k (List.filter (fun p => !List.contains l₂ p.fst) l₁) fallback = fallback := by intro h apply getKeyD_eq_fallback - rw [containsKey_filter_contains_eq_false_of_contains_right h] + rw [containsKey_filter_not_contains_of_contains_right h] · exact dl₁ -theorem getKeyD_filter_contains_map_fst_eq_false_of_contains_right [BEq α] [EquivBEq α] +theorem getKeyD_filter_not_contains_map_fst_of_contains_right [BEq α] [EquivBEq α] {l₁ l₂ : List ((a : α) × β a)} {k : α} {fallback : α} (dl₁ : DistinctKeys l₁) : containsKey k l₂ → getKeyD k (List.filter (fun p => !List.contains (l₂.map Sigma.fst) p.fst) l₁) fallback = fallback := by rw [containsKey_eq_contains_map_fst] - apply getKeyD_filter_contains_eq_false_of_contains_right + apply getKeyD_filter_not_contains_of_contains_right exact dl₁ -theorem getKeyD_filter_contains_eq_false_of_contains_eq_false_left [BEq α] [EquivBEq α] +theorem getKeyD_filter_not_contains_of_contains_eq_false_left [BEq α] [EquivBEq α] {l₁ : List ((a : α) × β a)} {l₂ : List α} {k : α} {fallback} (dl₁ : DistinctKeys l₁) : containsKey k l₁ = false → getKeyD k (List.filter (fun p => !List.contains l₂ p.fst) l₁) fallback = fallback := by intro h apply getKeyD_eq_fallback - apply containsKey_filter_contains_eq_false_eq_false_left + apply containsKey_filter_not_contains_eq_false_left · exact dl₁ · exact h -theorem getKeyD_filter_contains_map_fst_eq_false_of_contains_eq_false_left [BEq α] [EquivBEq α] +theorem getKeyD_filter_not_contains_map_fst_of_contains_eq_false_left [BEq α] [EquivBEq α] {l₁ l₂ : List ((a : α) × β a)} {k : α} {fallback} (dl₁ : DistinctKeys l₁) : containsKey k l₁ = false → getKeyD k (List.filter (fun p => !List.contains (l₂.map Sigma.fst) p.fst) l₁) fallback = fallback := by intro h - apply getKeyD_filter_contains_eq_false_of_contains_eq_false_left + apply getKeyD_filter_not_contains_of_contains_eq_false_left · exact dl₁ · exact h -theorem getValueCastD_filter_contains_eq_false_of_containsKey_eq_false_left {β : α → Type v} [BEq α] [LawfulBEq α] +theorem getValueCastD_filter_not_contains_of_containsKey_eq_false_left {β : α → Type v} [BEq α] [LawfulBEq α] {l₁ : List ((a : α) × β a)} {l₂ : List α} {k : α} {fallback : β k} (dl₁ : DistinctKeys l₁) : containsKey k l₁ = false → getValueCastD k (List.filter (fun p => !List.contains l₂ p.fst) l₁) fallback = fallback := by intro h apply getValueCastD_eq_fallback - apply List.containsKey_filter_contains_eq_false_eq_false_left + apply List.containsKey_filter_not_contains_eq_false_left · exact dl₁ · exact h -theorem getValueCastD_filter_contains_map_fst_eq_false_of_containsKey_eq_false_left {β : α → Type v} [BEq α] [LawfulBEq α] +theorem getValueCastD_filter_not_contains_map_fst_of_containsKey_eq_false_left {β : α → Type v} [BEq α] [LawfulBEq α] {l₁ l₂ : List ((a : α) × β a)} {k : α} {fallback : β k} (dl₁ : DistinctKeys l₁) : containsKey k l₁ = false → getValueCastD k (List.filter (fun p => !List.contains (l₂.map Sigma.fst) p.fst) l₁) fallback = fallback := by - apply getValueCastD_filter_contains_eq_false_of_containsKey_eq_false_left + apply getValueCastD_filter_not_contains_of_containsKey_eq_false_left · exact dl₁ -theorem length_filter_contains_eq_false_eq_length_left [BEq α] [EquivBEq α] +theorem length_filter_not_contains_eq_length_left [BEq α] [EquivBEq α] {l₁ : List ((a : α) × β a)} {l₂ : List α} (hl₁ : DistinctKeys l₁) (w : ∀ (a : α), containsKey a l₁ → List.contains l₂ a = false) : (l₁.filter fun p => !List.contains l₂ p.fst).length = l₁.length := by @@ -6045,15 +6045,15 @@ theorem length_filter_contains_eq_false_eq_length_left [BEq α] [EquivBEq α] apply w rw [containsKey_cons, BEq.rfl, Bool.true_or] -theorem length_filter_contains_map_fst_eq_false_eq_length_left [BEq α] [EquivBEq α] +theorem length_filter_not_contains_map_fst_eq_length_left [BEq α] [EquivBEq α] {l₁ l₂ : List ((a : α) × β a)} (hl₁ : DistinctKeys l₁) (w : ∀ (a : α), containsKey a l₁ → containsKey a l₂ = false) : (l₁.filter fun p => !List.contains (l₂.map Sigma.fst) p.fst).length = l₁.length := by - apply length_filter_contains_eq_false_eq_length_left hl₁ + apply length_filter_not_contains_eq_length_left hl₁ intro a mem rw [← List.containsKey_eq_contains_map_fst, w a mem] -theorem size_filter_contains_eq_false_add_size_inter_eq_size_left [BEq α] [EquivBEq α] +theorem size_filter_not_contains_add_size_inter_eq_size_left [BEq α] [EquivBEq α] {l₁ : List ((a : α) × β a)} {l₂ : List α} (dl₁ : DistinctKeys l₁) : (List.filter (fun p => !List.contains l₂ p.fst) l₁).length + (List.filter (fun p => List.contains l₂ p.fst) l₁).length = l₁.length := by @@ -6073,7 +6073,7 @@ theorem size_filter_contains_eq_false_add_size_inter_eq_size_left [BEq α] [Equi simp only [hc, ↓reduceIte, List.length_cons] omega -theorem size_filter_contains_map_fst_eq_false_add_size_inter_eq_size_left [BEq α] [EquivBEq α] +theorem size_filter_not_contains_map_fst_add_size_inter_eq_size_left [BEq α] [EquivBEq α] {l₁ l₂ : List ((a : α) × β a)} (dl₁ : DistinctKeys l₁) : (List.filter (fun p => !List.contains (l₂.map Sigma.fst) p.fst) l₁).length + (List.filter (fun p => containsKey p.fst l₂) l₁).length = l₁.length := by @@ -6084,22 +6084,22 @@ theorem size_filter_contains_map_fst_eq_false_add_size_inter_eq_size_left [BEq lhs ext p rw [containsKey_eq_contains_map_fst] - apply size_filter_contains_eq_false_add_size_inter_eq_size_left dl₁ + apply size_filter_not_contains_add_size_inter_eq_size_left dl₁ -theorem isEmpty_filter_contains_eq_false_left [BEq α] [EquivBEq α] +theorem isEmpty_filter_not_contains_left [BEq α] [EquivBEq α] {l₁ : List ((a : α) × β a)} {l₂ : List α}: l₁.isEmpty → (List.filter (fun p => !List.contains l₂ p.fst) l₁).isEmpty := by intro hyp simp at hyp simp [hyp] -theorem isEmpty_filter_contains_map_fst_eq_false_left [BEq α] [EquivBEq α] +theorem isEmpty_filter_not_contains_map_fst_left [BEq α] [EquivBEq α] {l₁ l₂ : List ((a : α) × β a)} : l₁.isEmpty → (List.filter (fun p => !List.contains (l₂.map Sigma.fst) p.fst) l₁).isEmpty := by - apply isEmpty_filter_contains_eq_false_left + apply isEmpty_filter_not_contains_left -theorem isEmpty_filter_contains_eq_false_iff [BEq α] [EquivBEq α] +theorem isEmpty_filter_not_contains_iff [BEq α] [EquivBEq α] {l₁ : List ((a : α) × β a)} {l₂ : List α} (dl₁ : DistinctKeys l₁) : (List.filter (fun p => !List.contains l₂ p.fst) l₁).isEmpty = true ↔ ∀ (k : α), containsKey k l₁ = true → List.contains l₂ k := by @@ -6137,7 +6137,7 @@ theorem isEmpty_filter_contains_eq_false_iff [BEq α] [EquivBEq α] · exact hyp · exact heq2 -theorem isEmpty_filter_contains_map_fst_eq_false_iff [BEq α] [EquivBEq α] +theorem isEmpty_filter_not_contains_map_fst_iff [BEq α] [EquivBEq α] {l₁ l₂ : List ((a : α) × β a)} (dl₁ : DistinctKeys l₁) : (List.filter (fun p => !List.contains (l₂.map Sigma.fst) p.fst) l₁).isEmpty = true ↔ ∀ (k : α), containsKey k l₁ = true → containsKey k l₂ := by @@ -6146,9 +6146,9 @@ theorem isEmpty_filter_contains_map_fst_eq_false_iff [BEq α] [EquivBEq α] ext k rhs rw [containsKey_eq_contains_map_fst] - apply isEmpty_filter_contains_eq_false_iff dl₁ + apply isEmpty_filter_not_contains_iff dl₁ -theorem getValue?_filter_contains_eq_false {β : Type v} [BEq α] [EquivBEq α] +theorem getValue?_filter_not_contains {β : Type v} [BEq α] [EquivBEq α] {l₁ : List ((_ : α) × β)} {l₂ : List α} {k : α} (dl₁ : DistinctKeys l₁) : getValue? k (List.filter (fun p => !List.contains l₂ p.fst) l₁) = @@ -6157,7 +6157,7 @@ theorem getValue?_filter_contains_eq_false {β : Type v} [BEq α] [EquivBEq α] split case isTrue h => simp only [Option.map_eq_none_iff, getEntry?_eq_none] - apply containsKey_filter_contains_eq_false_of_contains_right + apply containsKey_filter_not_contains_of_contains_right · exact dl₁ · exact h case isFalse h => @@ -6175,91 +6175,91 @@ theorem getValue?_filter_contains_eq_false {β : Type v} [BEq α] [EquivBEq α] simp [heq] · exact dl₁ -theorem getValue?_filter_contains_map_fst_eq_false {β : Type v} [BEq α] [EquivBEq α] +theorem getValue?_filter_not_contains_map_fst {β : Type v} [BEq α] [EquivBEq α] {l₁ l₂ : List ((_ : α) × β)} {k : α} (dl₁ : DistinctKeys l₁) : getValue? k (List.filter (fun p => !List.contains (l₂.map Sigma.fst) p.fst) l₁) = if containsKey k l₂ then none else getValue? k l₁ := by - rw [getValue?_filter_contains_eq_false, ← List.containsKey_eq_contains_map_fst] + rw [getValue?_filter_not_contains, ← List.containsKey_eq_contains_map_fst] exact dl₁ -theorem getValue?_filter_contains_eq_false_of_contains_eq_false_right {β : Type v} [BEq α] [EquivBEq α] +theorem getValue?_filter_not_contains_of_contains_eq_false_right {β : Type v} [BEq α] [EquivBEq α] {l₁ : List ((_ : α) × β)} {l₂ : List α} {k : α} (dl₁ : DistinctKeys l₁) : List.contains l₂ k = false → getValue? k (List.filter (fun p => !List.contains l₂ p.fst) l₁) = getValue? k l₁ := by intro h - rw [getValue?_filter_contains_eq_false, h] + rw [getValue?_filter_not_contains, h] · simp only [Bool.false_eq_true, ↓reduceIte] · exact dl₁ -theorem getValue?_filter_contains_map_fst_eq_false_of_contains_eq_false_right {β : Type v} [BEq α] [EquivBEq α] +theorem getValue?_filter_not_contains_map_fst_of_contains_eq_false_right {β : Type v} [BEq α] [EquivBEq α] {l₁ l₂ : List ((_ : α) × β)} {k : α} (dl₁ : DistinctKeys l₁) : containsKey k l₂ = false → getValue? k (List.filter (fun p => !List.contains (l₂.map Sigma.fst) p.fst) l₁) = getValue? k l₁ := by intro h - apply getValue?_filter_contains_eq_false_of_contains_eq_false_right dl₁ + apply getValue?_filter_not_contains_of_contains_eq_false_right dl₁ rwa [List.containsKey_eq_contains_map_fst] at h -theorem getValue?_filter_contains_eq_false_of_containsKey_eq_false_left {β : Type v} [BEq α] [EquivBEq α] +theorem getValue?_filter_not_contains_of_containsKey_eq_false_left {β : Type v} [BEq α] [EquivBEq α] {l₁ : List ((_ : α) × β)} {l₂ : List α} {k : α} (dl₁ : DistinctKeys l₁) : containsKey k l₁ = false → getValue? k (List.filter (fun p => !List.contains l₂ p.fst) l₁) = none := by intro h - rw [getValue?_filter_contains_eq_false] + rw [getValue?_filter_not_contains] · split · rfl · rw [@getValue?_eq_none α β _ l₁ k] exact h · exact dl₁ -theorem getValue?_filter_contains_map_fst_eq_false_of_containsKey_eq_false_left {β : Type v} [BEq α] [EquivBEq α] +theorem getValue?_filter_not_contains_map_fst_of_containsKey_eq_false_left {β : Type v} [BEq α] [EquivBEq α] {l₁ l₂ : List ((_ : α) × β)} {k : α} (dl₁ : DistinctKeys l₁) : containsKey k l₁ = false → getValue? k (List.filter (fun p => !List.contains (l₂.map Sigma.fst) p.fst) l₁) = none := by - apply getValue?_filter_contains_eq_false_of_containsKey_eq_false_left dl₁ + apply getValue?_filter_not_contains_of_containsKey_eq_false_left dl₁ -theorem getValue?_filter_contains_eq_false_of_contains_right {β : Type v} [BEq α] [EquivBEq α] +theorem getValue?_filter_not_contains_of_contains_right {β : Type v} [BEq α] [EquivBEq α] {l₁ : List ((_ : α) × β)} {l₂ : List α} {k : α} (dl₁ : DistinctKeys l₁) : List.contains l₂ k → getValue? k (List.filter (fun p => !List.contains l₂ p.fst) l₁) = none := by intro h - rw [getValue?_filter_contains_eq_false, h] + rw [getValue?_filter_not_contains, h] · simp only [↓reduceIte] · exact dl₁ -theorem getValue?_filter_contains_map_fst_eq_false_of_contains_right {β : Type v} [BEq α] [EquivBEq α] +theorem getValue?_filter_not_contains_map_fst_of_contains_right {β : Type v} [BEq α] [EquivBEq α] {l₁ l₂ : List ((_ : α) × β)} {k : α} (dl₁ : DistinctKeys l₁) : containsKey k l₂ → getValue? k (List.filter (fun p => !List.contains (l₂.map Sigma.fst) p.fst) l₁) = none := by intro h - rw [getValue?_filter_contains_eq_false_of_contains_right dl₁] + rw [getValue?_filter_not_contains_of_contains_right dl₁] rwa [List.containsKey_eq_contains_map_fst] at h -theorem getValue_filter_contains_eq_false {β : Type v} [BEq α] [EquivBEq α] +theorem getValue_filter_not_contains {β : Type v} [BEq α] [EquivBEq α] {l₁ : List ((_ : α) × β)} {l₂ : List α} {k : α} (dl₁ : DistinctKeys l₁) {h₁ h₂} : getValue k (List.filter (fun p => !List.contains l₂ p.fst) l₁) h₁ = getValue k l₁ h₂ := by suffices some (getValue k (List.filter (fun p => !List.contains l₂ p.fst) l₁) h₁) = some (getValue k l₁ h₂) by injections simp only [← getValue?_eq_some_getValue] - apply getValue?_filter_contains_eq_false_of_contains_eq_false_right - · rw [containsKey_filter_contains_eq_false_iff] at h₁ + apply getValue?_filter_not_contains_of_contains_eq_false_right + · rw [containsKey_filter_not_contains_iff] at h₁ · exact dl₁ · exact dl₁ - · rw [containsKey_filter_contains_eq_false_iff] at h₁ + · rw [containsKey_filter_not_contains_iff] at h₁ · replace h₁ := h₁.2 simp [h₁] · exact dl₁ -theorem getValue_filter_contains_map_fst_eq_false {β : Type v} [BEq α] [EquivBEq α] +theorem getValue_filter_not_contains_map_fst {β : Type v} [BEq α] [EquivBEq α] {l₁ l₂ : List ((_ : α) × β)} {k : α} (dl₁ : DistinctKeys l₁) {h₁ h₂} : getValue k (List.filter (fun p => !List.contains (l₂.map Sigma.fst) p.fst) l₁) h₁ = getValue k l₁ h₂ := by - apply getValue_filter_contains_eq_false dl₁ + apply getValue_filter_not_contains dl₁ -theorem getValueD_filter_contains_eq_false {β : Type v} [BEq α] [EquivBEq α] +theorem getValueD_filter_not_contains {β : Type v} [BEq α] [EquivBEq α] {l₁ : List ((_ : α) × β)} {l₂ : List α} {k : α} {fallback : β} (dl₁ : DistinctKeys l₁) : getValueD k (List.filter (fun p => !List.contains l₂ p.fst) l₁) fallback = @@ -6267,24 +6267,24 @@ theorem getValueD_filter_contains_eq_false {β : Type v} [BEq α] [EquivBEq α] split case isTrue h => apply getValueD_eq_fallback - apply containsKey_filter_contains_eq_false_of_contains_right + apply containsKey_filter_not_contains_of_contains_right · exact dl₁ · exact h case isFalse h => rw [getValueD_eq_getValue?, getValueD_eq_getValue?] congr 1 - apply getValue?_filter_contains_eq_false_of_contains_eq_false_right dl₁ + apply getValue?_filter_not_contains_of_contains_eq_false_right dl₁ simp [h] -theorem getValueD_filter_contains_map_fst_eq_false {β : Type v} [BEq α] [EquivBEq α] +theorem getValueD_filter_not_contains_map_fst {β : Type v} [BEq α] [EquivBEq α] {l₁ l₂ : List ((_ : α) × β)} {k : α} {fallback : β} (dl₁ : DistinctKeys l₁) : getValueD k (List.filter (fun p => !List.contains (l₂.map Sigma.fst) p.fst) l₁) fallback = if containsKey k l₂ then fallback else getValueD k l₁ fallback := by rw [containsKey_eq_contains_map_fst] - apply getValueD_filter_contains_eq_false dl₁ + apply getValueD_filter_not_contains dl₁ -theorem getValueD_filter_contains_eq_false_of_containsKey_eq_false_right {β : Type v} [BEq α] [EquivBEq α] +theorem getValueD_filter_not_contains_of_containsKey_eq_false_right {β : Type v} [BEq α] [EquivBEq α] {l₁ : List ((_ : α) × β)} {l₂ : List α} {k : α} {fallback : β} (dl₁ : DistinctKeys l₁) : List.contains l₂ k = false → getValueD k (List.filter (fun p => !List.contains l₂ p.fst) l₁) fallback = @@ -6292,70 +6292,70 @@ theorem getValueD_filter_contains_eq_false_of_containsKey_eq_false_right {β : T intro h rw [getValueD_eq_getValue?, getValueD_eq_getValue?] congr 1 - apply getValue?_filter_contains_eq_false_of_contains_eq_false_right dl₁ h + apply getValue?_filter_not_contains_of_contains_eq_false_right dl₁ h -theorem getValueD_filter_contains_map_fst_eq_false_of_containsKey_eq_false_right {β : Type v} [BEq α] [EquivBEq α] +theorem getValueD_filter_not_contains_map_fst_of_containsKey_eq_false_right {β : Type v} [BEq α] [EquivBEq α] {l₁ l₂ : List ((_ : α) × β)} {k : α} {fallback : β} (dl₁ : DistinctKeys l₁) : containsKey k l₂ = false → getValueD k (List.filter (fun p => !List.contains (l₂.map Sigma.fst) p.fst) l₁) fallback = getValueD k l₁ fallback := by intro h - apply getValueD_filter_contains_eq_false_of_containsKey_eq_false_right dl₁ + apply getValueD_filter_not_contains_of_containsKey_eq_false_right dl₁ rwa [containsKey_eq_contains_map_fst] at h -theorem getValueD_filter_contains_eq_false_of_contains_right {β : Type v} [BEq α] [EquivBEq α] +theorem getValueD_filter_not_contains_of_contains_right {β : Type v} [BEq α] [EquivBEq α] {l₁ : List ((_ : α) × β)} {l₂ : List α} {k : α} {fallback : β} (dl₁ : DistinctKeys l₁) : List.contains l₂ k → getValueD k (List.filter (fun p => !List.contains l₂ p.fst) l₁) fallback = fallback := by intro h apply getValueD_eq_fallback - apply List.containsKey_filter_contains_eq_false_of_contains_right + apply List.containsKey_filter_not_contains_of_contains_right · exact dl₁ · exact h -theorem getValueD_filter_contains_map_fst_eq_false_of_contains_map_fst_right {β : Type v} [BEq α] [EquivBEq α] +theorem getValueD_filter_not_contains_map_fst_of_contains_map_fst_right {β : Type v} [BEq α] [EquivBEq α] {l₁ l₂ : List ((_ : α) × β)} {k : α} {fallback : β} (dl₁ : DistinctKeys l₁) : containsKey k l₂ → getValueD k (List.filter (fun p => !List.contains (l₂.map Sigma.fst) p.fst) l₁) fallback = fallback := by intro hyp rw [containsKey_eq_contains_map_fst] at hyp - apply getValueD_filter_contains_eq_false_of_contains_right + apply getValueD_filter_not_contains_of_contains_right · exact dl₁ · exact hyp -theorem getValueD_filter_contains_eq_false_of_containsKey_eq_false_left {β : Type v} [BEq α] [EquivBEq α] +theorem getValueD_filter_not_contains_of_containsKey_eq_false_left {β : Type v} [BEq α] [EquivBEq α] {l₁ : List ((_ : α) × β)} {l₂ : List α} {k : α} {fallback : β} (dl₁ : DistinctKeys l₁) : containsKey k l₁ = false → getValueD k (List.filter (fun p => !List.contains l₂ p.fst) l₁) fallback = fallback := by intro h apply getValueD_eq_fallback - apply containsKey_filter_contains_eq_false_eq_false_left + apply containsKey_filter_not_contains_eq_false_left · exact dl₁ · exact h -theorem getValueD_filter_contains_map_fst_eq_false_of_containsKey_eq_false_left {β : Type v} [BEq α] [EquivBEq α] +theorem getValueD_filter_not_contains_map_fst_of_containsKey_eq_false_left {β : Type v} [BEq α] [EquivBEq α] {l₁ l₂ : List ((_ : α) × β)} {k : α} {fallback : β} (dl₁ : DistinctKeys l₁) : containsKey k l₁ = false → getValueD k (List.filter (fun p => !List.contains (l₂.map Sigma.fst) p.fst) l₁) fallback = fallback := by - apply getValueD_filter_contains_eq_false_of_containsKey_eq_false_left dl₁ + apply getValueD_filter_not_contains_of_containsKey_eq_false_left dl₁ -theorem getValueCastD_filter_contains_eq_false_of_contains_right {β : α → Type v} [BEq α] [LawfulBEq α] +theorem getValueCastD_filter_not_contains_of_contains_right {β : α → Type v} [BEq α] [LawfulBEq α] {l₁ : List ((a : α) × β a)} {l₂ : List α} {k : α} {fallback : β k} (dl₁ : DistinctKeys l₁) : List.contains l₂ k → getValueCastD k (List.filter (fun p => !List.contains l₂ p.fst) l₁) fallback = fallback := by intro h apply getValueCastD_eq_fallback - apply List.containsKey_filter_contains_eq_false_of_contains_right + apply List.containsKey_filter_not_contains_of_contains_right · exact dl₁ · exact h -theorem getValueCastD_filter_contains_map_fst_eq_false_of_contains_map_fst_right {β : α → Type v} [BEq α] [LawfulBEq α] +theorem getValueCastD_filter_not_contains_map_fst_of_contains_map_fst_right {β : α → Type v} [BEq α] [LawfulBEq α] {l₁ l₂ : List ((a : α) × β a)} {k : α} {fallback : β k} (dl₁ : DistinctKeys l₁) : containsKey k l₂ → getValueCastD k (List.filter (fun p => !List.contains (l₂.map Sigma.fst) p.fst) l₁) fallback = fallback := by intro hyp rw [containsKey_eq_contains_map_fst] at hyp - apply getValueCastD_filter_contains_eq_false_of_contains_right + apply getValueCastD_filter_not_contains_of_contains_right · exact dl₁ · exact hyp @@ -6463,7 +6463,7 @@ theorem getValueCastD_filter_containsKey_of_containsKey_eq_false_right [BEq α] . exact dl₁ . exact h -theorem getValueCastD_filter_contains_eq_false_of_contains_eq_false_right [BEq α] [LawfulBEq α] +theorem getValueCastD_filter_not_contains_of_contains_eq_false_right [BEq α] [LawfulBEq α] {l₁ : List ((a : α) × β a)} {l₂ : List α} {k : α} {fallback} (dl₁ : DistinctKeys l₁) : List.contains l₂ k = false → @@ -6472,14 +6472,14 @@ theorem getValueCastD_filter_contains_eq_false_of_contains_eq_false_right [BEq intro h rw [getValueCastD_eq_getValueCast?, getValueCastD_eq_getValueCast?] congr 1 - rw [getValueCast?_filter_contains_eq_false dl₁] + rw [getValueCast?_filter_not_contains dl₁] split · rename_i hyp rw [hyp] at h contradiction · rfl -theorem getValueCastD_filter_contains_map_fst_eq_false_of_containsKey_eq_false_right [BEq α] [LawfulBEq α] +theorem getValueCastD_filter_not_contains_map_fst_of_containsKey_eq_false_right [BEq α] [LawfulBEq α] {l₁ l₂ : List ((a : α) × β a)} {k : α} {fallback} (dl₁ : DistinctKeys l₁) : containsKey k l₂ = false → @@ -6487,7 +6487,7 @@ theorem getValueCastD_filter_contains_map_fst_eq_false_of_containsKey_eq_false_r getValueCastD k l₁ fallback := by intro hyp rw [List.containsKey_eq_contains_map_fst] at hyp - apply getValueCastD_filter_contains_eq_false_of_contains_eq_false_right dl₁ hyp + apply getValueCastD_filter_not_contains_of_contains_eq_false_right dl₁ hyp theorem getKeyD_filter_containsKey [BEq α] [EquivBEq α] {l₁ l₂ : List ((a : α) × β a)} {k : α} {fallback : α} From 8f760c2cd17bbb1cd0720154de5ecabb811ccd5c Mon Sep 17 00:00:00 2001 From: Wojciech Rozowski Date: Thu, 27 Nov 2025 12:15:55 +0000 Subject: [PATCH 86/88] add congr lemmas --- src/Std/Data/DHashMap/Internal/RawLemmas.lean | 26 ++++++++++++- src/Std/Data/DHashMap/Lemmas.lean | 16 ++++++++ src/Std/Data/DHashMap/RawLemmas.lean | 31 +++++++++++++++ src/Std/Data/HashMap/Lemmas.lean | 16 ++++++++ src/Std/Data/HashMap/RawLemmas.lean | 16 ++++++++ src/Std/Data/HashSet/Lemmas.lean | 16 ++++++++ src/Std/Data/HashSet/RawLemmas.lean | 16 ++++++++ src/Std/Data/Internal/List/Associative.lean | 38 +++++++++++++++++++ 8 files changed, 174 insertions(+), 1 deletion(-) diff --git a/src/Std/Data/DHashMap/Internal/RawLemmas.lean b/src/Std/Data/DHashMap/Internal/RawLemmas.lean index 45150a476b37..9fd7f20c7f2f 100644 --- a/src/Std/Data/DHashMap/Internal/RawLemmas.lean +++ b/src/Std/Data/DHashMap/Internal/RawLemmas.lean @@ -3041,7 +3041,7 @@ theorem Equiv.inter_right {m₃ : Raw₀ α β} [EquivBEq α] [LawfulHashable α all_goals wf_trivial theorem Equiv.inter_congr {m₃ m₄ : Raw₀ α β} [EquivBEq α] [LawfulHashable α] - (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) (h₃ : m₃.val.WF) (h₄ : m₄.val.WF) (equiv₁ : m₁.1.Equiv m₃.1) (equiv₂ : m₂.1.Equiv m₄.1) : + (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) (h₃ : m₃.val.WF) (h₄ : m₄.val.WF) (equiv₁ : m₁.1.Equiv m₃.1) (equiv₂ : m₂.1.Equiv m₄.1) : (m₁.inter m₂).1.Equiv (m₃.inter m₄).1 := by revert equiv₁ equiv₂ simp_to_model [Equiv, inter] @@ -3376,6 +3376,30 @@ theorem contains_diff_eq_false_of_contains_right [EquivBEq α] [LawfulHashable revert h simp_to_model [diff, contains] using List.containsKey_filter_not_contains_map_fst_of_contains_map_fst_right +theorem Equiv.diff_left {m₃ : Raw₀ α β} [EquivBEq α] [LawfulHashable α] + (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) (h₃ : m₃.val.WF) (equiv : m₁.1.Equiv m₂.1) : + (m₁.diff m₃).1.Equiv (m₂.diff m₃).1 := by + revert equiv + simp_to_model [Equiv, diff] using List.Perm.filter + +theorem Equiv.diff_right {m₃ : Raw₀ α β} [EquivBEq α] [LawfulHashable α] + (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) (h₃ : m₃.val.WF) (equiv : m₂.1.Equiv m₃.1) : + (m₁.diff m₂).1.Equiv (m₁.diff m₃).1 := by + revert equiv + simp_to_model [Equiv, diff] + intro equiv + apply perm_filter_not_contains_map_fst_of_perm equiv + all_goals wf_trivial + +theorem Equiv.diff_congr {m₃ m₄ : Raw₀ α β} [EquivBEq α] [LawfulHashable α] + (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) (h₃ : m₃.val.WF) (h₄ : m₄.val.WF) (equiv₁ : m₁.1.Equiv m₃.1) (equiv₂ : m₂.1.Equiv m₄.1) : + (m₁.diff m₂).1.Equiv (m₃.diff m₄).1 := by + revert equiv₁ equiv₂ + simp_to_model [Equiv, diff] + intro equiv₁ equiv₂ + apply List.congr_filter_not_contains_map_fst_of_perm + all_goals wf_trivial + /- get? -/ theorem get?_diff [LawfulBEq α] (h₁ : m₁.val.WF) (h₂ : m₂.val.WF) {k : α} : (m₁.diff m₂).get? k = if m₂.contains k then none else m₁.get? k := by diff --git a/src/Std/Data/DHashMap/Lemmas.lean b/src/Std/Data/DHashMap/Lemmas.lean index e79a5f91dbce..955e72327441 100644 --- a/src/Std/Data/DHashMap/Lemmas.lean +++ b/src/Std/Data/DHashMap/Lemmas.lean @@ -2519,6 +2519,22 @@ theorem not_mem_diff_of_mem_right [EquivBEq α] [LawfulHashable α] {k : α} rw [← contains_eq_false_iff_not_mem] exact @Raw₀.contains_diff_eq_false_of_contains_right _ _ _ _ ⟨m₁.1, m₁.2.size_buckets_pos⟩ ⟨m₂.1, m₂.2.size_buckets_pos⟩ _ _ m₁.2 m₂.2 k mem +/- Equiv -/ +theorem Equiv.diff_left {m₃ : DHashMap α β} [EquivBEq α] [LawfulHashable α] + (equiv : m₁ ~m m₂) : + (m₁ \ m₃) ~m (m₂ \ m₃) := + ⟨@Raw₀.Equiv.diff_left α β _ _ ⟨m₁.1, m₁.2.size_buckets_pos⟩ ⟨m₂.1, m₂.2.size_buckets_pos⟩ ⟨m₃.1, m₃.2.size_buckets_pos⟩ _ _ m₁.2 m₂.2 m₃.2 equiv.1⟩ + +theorem Equiv.diff_right {m₃ : DHashMap α β} [EquivBEq α] [LawfulHashable α] + (equiv : m₂ ~m m₃) : + (m₁ \ m₂) ~m (m₁ \ m₃) := + ⟨@Raw₀.Equiv.diff_right α β _ _ ⟨m₁.1, m₁.2.size_buckets_pos⟩ ⟨m₂.1, m₂.2.size_buckets_pos⟩ ⟨m₃.1, m₃.2.size_buckets_pos⟩ _ _ m₁.2 m₂.2 m₃.2 equiv.1⟩ + +theorem Equiv.diff_congr {m₃ m₄ : DHashMap α β} [EquivBEq α] [LawfulHashable α] + (equiv₁ : m₁ ~m m₃) (equiv₂ : m₂ ~m m₄) : + (m₁ \ m₂) ~m (m₃ \ m₄) := + ⟨@Raw₀.Equiv.diff_congr α β _ _ ⟨m₁.1, m₁.2.size_buckets_pos⟩ ⟨m₂.1, m₂.2.size_buckets_pos⟩ ⟨m₃.1, m₃.2.size_buckets_pos⟩ ⟨m₄.1, m₄.2.size_buckets_pos⟩ _ _ m₁.2 m₂.2 m₃.2 m₄.2 equiv₁.1 equiv₂.1⟩ + /- get? -/ theorem get?_diff [LawfulBEq α] {k : α} : (m₁ \ m₂).get? k = if k ∈ m₂ then none else m₁.get? k := diff --git a/src/Std/Data/DHashMap/RawLemmas.lean b/src/Std/Data/DHashMap/RawLemmas.lean index ae7b13a2c530..014cec081d98 100644 --- a/src/Std/Data/DHashMap/RawLemmas.lean +++ b/src/Std/Data/DHashMap/RawLemmas.lean @@ -2948,6 +2948,37 @@ theorem not_mem_diff_of_mem_right [EquivBEq α] [LawfulHashable α] (h₁ : m₁ simp only [SDiff.sdiff] simp_to_raw using Raw₀.contains_diff_eq_false_of_contains_right +/- Equiv -/ +theorem Equiv.diff_left {m₃ : Raw α β} [EquivBEq α] [LawfulHashable α] (h₁ : m₁.WF) (h₂ : m₂.WF) (h₃ : m₃.WF) + (equiv : m₁ ~m m₂) : + (m₁ \ m₃) ~m (m₂ \ m₃) := by + revert equiv + simp only [SDiff.sdiff] + simp_to_raw + intro hyp + apply Raw₀.Equiv.diff_left + all_goals wf_trivial + +theorem Equiv.diff_right {m₃ : Raw α β} [EquivBEq α] [LawfulHashable α] (h₁ : m₁.WF) (h₂ : m₂.WF) (h₃ : m₃.WF) + (equiv : m₂ ~m m₃) : + (m₁ \ m₂) ~m (m₁ \ m₃) := by + revert equiv + simp only [SDiff.sdiff] + simp_to_raw + intro hyp + apply Raw₀.Equiv.diff_right + all_goals wf_trivial + +theorem Equiv.diff_congr {m₃ m₄ : Raw α β} [EquivBEq α] [LawfulHashable α] (h₁ : m₁.WF) (h₂ : m₂.WF) (h₃ : m₃.WF) (h₄ : m₄.WF) + (equiv₁ : m₁ ~m m₃) (equiv₂ : m₂ ~m m₄) : + (m₁ \ m₂) ~m (m₃ \ m₄) := by + revert equiv₁ equiv₂ + simp only [SDiff.sdiff] + simp_to_raw + intro equiv₁ equiv₂ + apply Raw₀.Equiv.diff_congr + all_goals wf_trivial + /- get? -/ theorem get?_diff [LawfulBEq α] (h₁ : m₁.WF) (h₂ : m₂.WF) {k : α} : (m₁ \ m₂).get? k = if k ∈ m₂ then none else m₁.get? k := by diff --git a/src/Std/Data/HashMap/Lemmas.lean b/src/Std/Data/HashMap/Lemmas.lean index ca6a3d34105a..7446672161fd 100644 --- a/src/Std/Data/HashMap/Lemmas.lean +++ b/src/Std/Data/HashMap/Lemmas.lean @@ -1843,6 +1843,22 @@ theorem get?_diff_of_mem_right [EquivBEq α] [LawfulHashable α] (m₁ \ m₂).get? k = none := @DHashMap.Const.get?_diff_of_mem_right _ _ _ _ m₁.inner m₂.inner _ _ k mem +/- Equiv -/ +theorem Equiv.diff_left {m₃ : HashMap α β} [EquivBEq α] [LawfulHashable α] + (equiv : m₁ ~m m₂) : + (m₁ \ m₃) ~m (m₂ \ m₃) := + ⟨DHashMap.Equiv.diff_left equiv.1⟩ + +theorem Equiv.diff_right {m₃ : HashMap α β} [EquivBEq α] [LawfulHashable α] + (equiv : m₂ ~m m₃) : + (m₁ \ m₂) ~m (m₁ \ m₃) := + ⟨DHashMap.Equiv.diff_right equiv.1⟩ + +theorem Equiv.diff_congr {m₃ m₄ : HashMap α β} [EquivBEq α] [LawfulHashable α] + (equiv₁ : m₁ ~m m₃) (equiv₂ : m₂ ~m m₄) : + (m₁ \ m₂) ~m (m₃ \ m₄) := + ⟨DHashMap.Equiv.diff_congr equiv₁.1 equiv₂.1⟩ + /- get -/ @[simp] theorem get_diff [EquivBEq α] [LawfulHashable α] {k : α} {h_mem : k ∈ m₁ \ m₂} : diff --git a/src/Std/Data/HashMap/RawLemmas.lean b/src/Std/Data/HashMap/RawLemmas.lean index 4e2b8fa38521..1a259c4ae58f 100644 --- a/src/Std/Data/HashMap/RawLemmas.lean +++ b/src/Std/Data/HashMap/RawLemmas.lean @@ -1814,6 +1814,22 @@ theorem not_mem_diff_of_mem_right [EquivBEq α] [LawfulHashable α] k ∉ m₁ \ m₂ := @DHashMap.Raw.not_mem_diff_of_mem_right _ _ _ _ m₁.inner m₂.inner _ _ h₁.out h₂.out k mem +/- Equiv -/ +theorem Equiv.diff_left {m₃ : Raw α β} [EquivBEq α] [LawfulHashable α] (h₁ : m₁.WF) (h₂ : m₂.WF) (h₃ : m₃.WF) + (equiv : m₁ ~m m₂) : + (m₁ \ m₃) ~m (m₂ \ m₃) := + ⟨@DHashMap.Raw.Equiv.diff_left _ _ _ _ m₁.inner m₂.inner m₃.inner _ _ h₁.out h₂.out h₃.out equiv.1⟩ + +theorem Equiv.diff_right {m₃ : Raw α β} [EquivBEq α] [LawfulHashable α] (h₁ : m₁.WF) (h₂ : m₂.WF) (h₃ : m₃.WF) + (equiv : m₂ ~m m₃) : + (m₁ \ m₂) ~m (m₁ \ m₃) := + ⟨@DHashMap.Raw.Equiv.diff_right _ _ _ _ m₁.inner m₂.inner m₃.inner _ _ h₁.out h₂.out h₃.out equiv.1⟩ + +theorem Equiv.diff_congr {m₃ m₄ : Raw α β} [EquivBEq α] [LawfulHashable α] (h₁ : m₁.WF) (h₂ : m₂.WF) (h₃ : m₃.WF) (h₄ : m₄.WF) + (equiv₁ : m₁ ~m m₃) (equiv₂ : m₂ ~m m₄) : + (m₁ \ m₂) ~m (m₃ \ m₄) := + ⟨@DHashMap.Raw.Equiv.diff_congr _ _ _ _ m₁.inner m₂.inner m₃.inner m₄.inner _ _ h₁.out h₂.out h₃.out h₄.out equiv₁.1 equiv₂.1⟩ + /- get? -/ theorem get?_diff [EquivBEq α] [LawfulHashable α] (h₁ : m₁.WF) (h₂ : m₂.WF) {k : α} : get? (m₁ \ m₂) k = diff --git a/src/Std/Data/HashSet/Lemmas.lean b/src/Std/Data/HashSet/Lemmas.lean index fa69c872d9e5..554babd22c9c 100644 --- a/src/Std/Data/HashSet/Lemmas.lean +++ b/src/Std/Data/HashSet/Lemmas.lean @@ -1114,6 +1114,22 @@ theorem not_mem_diff_of_mem_right [EquivBEq α] [LawfulHashable α] {k : α} k ∉ m₁ \ m₂ := @HashMap.not_mem_diff_of_mem_right _ _ _ _ m₁.inner m₂.inner _ _ k mem +/- Equiv -/ +theorem Equiv.diff_left {m₃ : HashSet α} [EquivBEq α] [LawfulHashable α] + (equiv : m₁ ~m m₂) : + (m₁ \ m₃) ~m (m₂ \ m₃) := + ⟨HashMap.Equiv.diff_left equiv.1⟩ + +theorem Equiv.diff_right {m₃ : HashSet α} [EquivBEq α] [LawfulHashable α] + (equiv : m₂ ~m m₃) : + (m₁ \ m₂) ~m (m₁ \ m₃) := + ⟨HashMap.Equiv.diff_right equiv.1⟩ + +theorem Equiv.diff_congr {m₃ m₄ : HashSet α} [EquivBEq α] [LawfulHashable α] + (equiv₁ : m₁ ~m m₃) (equiv₂ : m₂ ~m m₄) : + (m₁ \ m₂) ~m (m₃ \ m₄) := + ⟨HashMap.Equiv.diff_congr equiv₁.1 equiv₂.1⟩ + /- get? -/ theorem get?_diff [EquivBEq α] [LawfulHashable α] {k : α} : (m₁ \ m₂).get? k = diff --git a/src/Std/Data/HashSet/RawLemmas.lean b/src/Std/Data/HashSet/RawLemmas.lean index 21858134082f..1ca8451b5818 100644 --- a/src/Std/Data/HashSet/RawLemmas.lean +++ b/src/Std/Data/HashSet/RawLemmas.lean @@ -1184,6 +1184,22 @@ theorem not_mem_diff_of_mem_right [EquivBEq α] [LawfulHashable α] k ∉ m₁ \ m₂ := @HashMap.Raw.not_mem_diff_of_mem_right _ _ _ _ m₁.inner m₂.inner _ _ h₁.out h₂.out k mem +/- Equiv -/ +theorem Equiv.diff_left {m₃ : Raw α} [EquivBEq α] [LawfulHashable α] (h₁ : m₁.WF) (h₂ : m₂.WF) (h₃ : m₃.WF) + (equiv : m₁ ~m m₂) : + (m₁ \ m₃) ~m (m₂ \ m₃) := + ⟨@HashMap.Raw.Equiv.diff_left _ _ _ _ m₁.inner m₂.inner m₃.inner _ _ h₁.out h₂.out h₃.out equiv.1⟩ + +theorem Equiv.diff_right {m₃ : Raw α} [EquivBEq α] [LawfulHashable α] (h₁ : m₁.WF) (h₂ : m₂.WF) (h₃ : m₃.WF) + (equiv : m₂ ~m m₃) : + (m₁ \ m₂) ~m (m₁ \ m₃) := + ⟨@HashMap.Raw.Equiv.diff_right _ _ _ _ m₁.inner m₂.inner m₃.inner _ _ h₁.out h₂.out h₃.out equiv.1⟩ + +theorem Equiv.diff_congr {m₃ m₄ : Raw α} [EquivBEq α] [LawfulHashable α] (h₁ : m₁.WF) (h₂ : m₂.WF) (h₃ : m₃.WF) (h₄ : m₄.WF) + (equiv₁ : m₁ ~m m₃) (equiv₂ : m₂ ~m m₄) : + (m₁ \ m₂) ~m (m₃ \ m₄) := + ⟨@HashMap.Raw.Equiv.diff_congr _ _ _ _ m₁.inner m₂.inner m₃.inner m₄.inner _ _ h₁.out h₂.out h₃.out h₄.out equiv₁.1 equiv₂.1⟩ + /- get? -/ theorem get?_diff [EquivBEq α] [LawfulHashable α] (h₁ : m₁.WF) (h₂ : m₂.WF) diff --git a/src/Std/Data/Internal/List/Associative.lean b/src/Std/Data/Internal/List/Associative.lean index d4a248dc87a8..eb81d1dc2f2c 100644 --- a/src/Std/Data/Internal/List/Associative.lean +++ b/src/Std/Data/Internal/List/Associative.lean @@ -5761,6 +5761,34 @@ theorem perm_filter_containsKey_of_perm {l₁ l₂ l₃ : List ((a : α) × β a · simp only [List.perm_cons, ih] · simp [ih] +theorem perm_filter_not_contains_map_fst_of_perm {l₁ l₂ l₃ : List ((a : α) × β a)} + [BEq α] [EquivBEq α] + (h : l₂.Perm l₃) + (wf₁ : DistinctKeys l₁) : + (l₁.filter (fun p => !List.contains (l₂.map Sigma.fst) p.1)).Perm (l₁.filter (fun p => !List.contains (l₃.map Sigma.fst) p.1)) := by + conv => + lhs + lhs + ext p + congr + rw [← List.containsKey_eq_contains_map_fst] + conv => + rhs + lhs + ext p + congr + rw [← List.containsKey_eq_contains_map_fst] + induction l₁ + case nil => simp + case cons hd tl ih => + rw [List.distinctKeys_cons_iff] at wf₁ + rw [List.filter_cons, List.filter_cons] + rw [List.containsKey_of_perm h] + specialize ih wf₁.1 + split + · simp only [List.perm_cons, ih] + · simp [ih] + theorem congr_filter_containsKey_of_perm {l₁ l₂ l₃ l₄ : List ((a : α) × β a)} [BEq α] [EquivBEq α] (h₁ : l₁.Perm l₃) (h₂ : l₂.Perm l₄) @@ -5771,6 +5799,16 @@ theorem congr_filter_containsKey_of_perm {l₁ l₂ l₃ l₄ : List ((a : α) · exact h₁ · apply perm_filter_containsKey_of_perm h₂ hd +theorem congr_filter_not_contains_map_fst_of_perm {l₁ l₂ l₃ l₄ : List ((a : α) × β a)} + [BEq α] [EquivBEq α] + (h₁ : l₁.Perm l₃) (h₂ : l₂.Perm l₄) + (hd : DistinctKeys l₃) : + (l₁.filter (fun p => !List.contains (l₂.map Sigma.fst) p.1)).Perm (l₃.filter (fun p => !List.contains (l₄.map Sigma.fst) p.1)) := by + apply Perm.trans + · apply List.Perm.filter + · exact h₁ + · apply perm_filter_not_contains_map_fst_of_perm h₂ hd + theorem containsKey_filter_not_contains_of_contains_right [BEq α] [EquivBEq α] {l₁ : List ((a : α) × β a)} {l₂ : List α} {hl₁ : DistinctKeys l₁} {k : α} : List.contains l₂ k → containsKey k (List.filter (fun p => !List.contains l₂ p.fst) l₁) = false := by intro h From 87a1da9fbc92ec0b3d67fc208f5c8b0d054ca60c Mon Sep 17 00:00:00 2001 From: Wojciech Rozowski Date: Thu, 27 Nov 2025 12:32:56 +0000 Subject: [PATCH 87/88] Add Ext(D)HashMap lemmas --- src/Std/Data/ExtDHashMap/Basic.lean | 11 + src/Std/Data/ExtDHashMap/Lemmas.lean | 343 ++++++++++++++++++++++++++- src/Std/Data/ExtHashMap/Basic.lean | 5 + src/Std/Data/ExtHashMap/Lemmas.lean | 201 ++++++++++++++++ 4 files changed, 547 insertions(+), 13 deletions(-) diff --git a/src/Std/Data/ExtDHashMap/Basic.lean b/src/Std/Data/ExtDHashMap/Basic.lean index 07f1484b1441..bf39770e6715 100644 --- a/src/Std/Data/ExtDHashMap/Basic.lean +++ b/src/Std/Data/ExtDHashMap/Basic.lean @@ -373,6 +373,17 @@ def inter [EquivBEq α] [LawfulHashable α] (m₁ m₂ : ExtDHashMap α β) : Ex instance [EquivBEq α] [LawfulHashable α] : Inter (ExtDHashMap α β) := ⟨inter⟩ +@[inline, inherit_doc DHashMap.diff] +def diff [EquivBEq α] [LawfulHashable α] (m₁ m₂ : ExtDHashMap α β) : ExtDHashMap α β := lift₂ (fun x y : DHashMap α β => mk (x.diff y)) + (fun a b c d equiv₁ equiv₂ => by + simp only [DHashMap.diff_eq, mk'.injEq] + apply Quotient.sound + apply DHashMap.Equiv.diff_congr + · exact equiv₁ + · exact equiv₂) m₁ m₂ + +instance [EquivBEq α] [LawfulHashable α] : SDiff (ExtDHashMap α β) := ⟨diff⟩ + @[inline, inherit_doc DHashMap.Const.unitOfArray] def Const.unitOfArray [BEq α] [Hashable α] (l : Array α) : ExtDHashMap α (fun _ => Unit) := diff --git a/src/Std/Data/ExtDHashMap/Lemmas.lean b/src/Std/Data/ExtDHashMap/Lemmas.lean index 46b894f59090..36631a2b2e6b 100644 --- a/src/Std/Data/ExtDHashMap/Lemmas.lean +++ b/src/Std/Data/ExtDHashMap/Lemmas.lean @@ -2669,25 +2669,25 @@ variable {β : Type v} {m₁ m₂ : ExtDHashMap α (fun _ => β)} /- get? -/ theorem get?_inter [EquivBEq α] [LawfulHashable α] {k : α} : - Const.get? (m₁.inter m₂) k = + Const.get? (m₁ ∩ m₂) k = if k ∈ m₂ then Const.get? m₁ k else none := m₁.inductionOn₂ m₂ fun _ _ => DHashMap.Const.get?_inter theorem get?_inter_of_mem_right [EquivBEq α] [LawfulHashable α] {k : α} (mem : k ∈ m₂) : - Const.get? (m₁.inter m₂) k = Const.get? m₁ k := by + Const.get? (m₁ ∩ m₂) k = Const.get? m₁ k := by revert mem exact m₁.inductionOn₂ m₂ fun _ _ h => DHashMap.Const.get?_inter_of_mem_right h theorem get?_inter_of_not_mem_left [EquivBEq α] [LawfulHashable α] {k : α} (not_mem : k ∉ m₁) : - Const.get? (m₁.inter m₂) k = none := by + Const.get? (m₁ ∩ m₂) k = none := by revert not_mem exact m₁.inductionOn₂ m₂ fun _ _ h => DHashMap.Const.get?_inter_of_not_mem_left h theorem get?_inter_of_not_mem_right [EquivBEq α] [LawfulHashable α] {k : α} (not_mem : k ∉ m₂) : - Const.get? (m₁.inter m₂) k = none := by + Const.get? (m₁ ∩ m₂) k = none := by revert not_mem exact m₁.inductionOn₂ m₂ fun _ _ h => DHashMap.Const.get?_inter_of_not_mem_right h @@ -2695,7 +2695,7 @@ theorem get?_inter_of_not_mem_right [EquivBEq α] [LawfulHashable α] @[simp] theorem get_inter [EquivBEq α] [LawfulHashable α] {k : α} {h_mem : k ∈ m₁ ∩ m₂} : - Const.get (m₁.inter m₂) k h_mem = + Const.get (m₁ ∩ m₂) k h_mem = Const.get m₁ k (mem_inter_iff.1 h_mem).1 := by induction m₁ case mk a => @@ -2705,54 +2705,371 @@ theorem get_inter [EquivBEq α] [LawfulHashable α] /- getD -/ theorem getD_inter [EquivBEq α] [LawfulHashable α] {k : α} {fallback : β} : - Const.getD (m₁.inter m₂) k fallback = + Const.getD (m₁ ∩ m₂) k fallback = if k ∈ m₂ then Const.getD m₁ k fallback else fallback := m₁.inductionOn₂ m₂ fun _ _ => DHashMap.Const.getD_inter theorem getD_inter_of_mem_right [EquivBEq α] [LawfulHashable α] {k : α} {fallback : β} (mem : k ∈ m₂) : - Const.getD (m₁.inter m₂) k fallback = Const.getD m₁ k fallback := by + Const.getD (m₁ ∩ m₂) k fallback = Const.getD m₁ k fallback := by revert mem exact m₁.inductionOn₂ m₂ fun _ _ h => DHashMap.Const.getD_inter_of_mem_right h theorem getD_inter_of_not_mem_right [EquivBEq α] [LawfulHashable α] {k : α} {fallback : β} (not_mem : k ∉ m₂) : - Const.getD (m₁.inter m₂) k fallback = fallback := by + Const.getD (m₁ ∩ m₂) k fallback = fallback := by revert not_mem exact m₁.inductionOn₂ m₂ fun _ _ h => DHashMap.Const.getD_inter_of_not_mem_right h theorem getD_inter_of_not_mem_left [EquivBEq α] [LawfulHashable α] {k : α} {fallback : β} (not_mem : k ∉ m₁) : - Const.getD (m₁.inter m₂) k fallback = fallback := by + Const.getD (m₁ ∩ m₂) k fallback = fallback := by revert not_mem exact m₁.inductionOn₂ m₂ fun _ _ h => DHashMap.Const.getD_inter_of_not_mem_left h /- get! -/ theorem get!_inter [EquivBEq α] [LawfulHashable α] [Inhabited β] {k : α} : - Const.get! (m₁.inter m₂) k = + Const.get! (m₁ ∩ m₂) k = if k ∈ m₂ then Const.get! m₁ k else default := m₁.inductionOn₂ m₂ fun _ _ => DHashMap.Const.get!_inter theorem get!_inter_of_mem_right [EquivBEq α] [LawfulHashable α] [Inhabited β] {k : α} (mem : k ∈ m₂) : - Const.get! (m₁.inter m₂) k = Const.get! m₁ k := by + Const.get! (m₁ ∩ m₂) k = Const.get! m₁ k := by revert mem exact m₁.inductionOn₂ m₂ fun _ _ h => DHashMap.Const.get!_inter_of_mem_right h theorem get!_inter_of_not_mem_right [EquivBEq α] [LawfulHashable α] [Inhabited β] {k : α} (not_mem : k ∉ m₂) : - Const.get! (m₁.inter m₂) k = default := by + Const.get! (m₁ ∩ m₂) k = default := by revert not_mem exact m₁.inductionOn₂ m₂ fun _ _ h => DHashMap.Const.get!_inter_of_not_mem_right h theorem get!_inter_of_not_mem_left [EquivBEq α] [LawfulHashable α] [Inhabited β] {k : α} (not_mem : k ∉ m₁) : - Const.get! (m₁.inter m₂) k = default := by + Const.get! (m₁ ∩ m₂) k = default := by revert not_mem exact m₁.inductionOn₂ m₂ fun _ _ h => DHashMap.Const.get!_inter_of_not_mem_left h end Const +section Diff + +variable {m₁ m₂ : ExtDHashMap α β} + +@[simp] +theorem diff_eq [EquivBEq α] [LawfulHashable α] : m₁.diff m₂ = m₁ \ m₂ := by + simp only [SDiff.sdiff] + +/- contains -/ +@[simp] +theorem contains_diff [EquivBEq α] [LawfulHashable α] + {k : α} : + (m₁ \ m₂).contains k = (m₁.contains k && !m₂.contains k) := + m₁.inductionOn₂ m₂ fun _ _ => DHashMap.contains_diff + +/- mem -/ +@[simp] +theorem mem_diff_iff [EquivBEq α] [LawfulHashable α] {k : α} : + k ∈ m₁ \ m₂ ↔ k ∈ m₁ ∧ k ∉ m₂ := + m₁.inductionOn₂ m₂ fun _ _ => DHashMap.mem_diff_iff + +theorem not_mem_diff_of_not_mem_left [EquivBEq α] [LawfulHashable α] {k : α} + (not_mem : k ∉ m₁) : + k ∉ m₁ \ m₂ := by + revert not_mem + exact m₁.inductionOn₂ m₂ fun _ _ => DHashMap.not_mem_diff_of_not_mem_left + +theorem not_mem_diff_of_mem_right [EquivBEq α] [LawfulHashable α] {k : α} + (mem : k ∈ m₂) : + k ∉ m₁ \ m₂ := by + revert mem + exact m₁.inductionOn₂ m₂ fun _ _ => DHashMap.not_mem_diff_of_mem_right + +/- get? -/ +theorem get?_diff [LawfulBEq α] {k : α} : + (m₁ \ m₂).get? k = if k ∈ m₂ then none else m₁.get? k := + m₁.inductionOn₂ m₂ fun _ _ => DHashMap.get?_diff + +theorem get?_diff_of_not_mem_right [LawfulBEq α] + {k : α} (not_mem : k ∉ m₂) : + (m₁ \ m₂).get? k = m₁.get? k := by + revert not_mem + exact m₁.inductionOn₂ m₂ fun _ _ h => DHashMap.get?_diff_of_not_mem_right h + +theorem get?_diff_of_not_mem_left [LawfulBEq α] + {k : α} (not_mem : k ∉ m₁) : + (m₁ \ m₂).get? k = none := by + revert not_mem + exact m₁.inductionOn₂ m₂ fun _ _ h => DHashMap.get?_diff_of_not_mem_left h + +theorem get?_diff_of_mem_right [LawfulBEq α] + {k : α} (mem : k ∈ m₂) : + (m₁ \ m₂).get? k = none := by + revert mem + exact m₁.inductionOn₂ m₂ fun _ _ h => DHashMap.get?_diff_of_mem_right h + +/- get -/ +@[simp] +theorem get_diff [LawfulBEq α] + {k : α} {h_mem : k ∈ m₁ \ m₂} : + (m₁ \ m₂).get k h_mem = + m₁.get k (mem_diff_iff.1 h_mem).1 := by + induction m₁ + case mk a => + induction m₂ + case mk b => + apply DHashMap.get_diff + +/- getD -/ +theorem getD_diff [LawfulBEq α] {k : α} {fallback : β k} : + (m₁ \ m₂).getD k fallback = + if k ∈ m₂ then fallback else m₁.getD k fallback := + m₁.inductionOn₂ m₂ fun _ _ => DHashMap.getD_diff + +theorem getD_diff_of_not_mem_right [LawfulBEq α] + {k : α} {fallback : β k} (not_mem : k ∉ m₂) : + (m₁ \ m₂).getD k fallback = m₁.getD k fallback := by + revert not_mem + exact m₁.inductionOn₂ m₂ fun _ _ h => DHashMap.getD_diff_of_not_mem_right h + +theorem getD_diff_of_mem_right [LawfulBEq α] + {k : α} {fallback : β k} (mem : k ∈ m₂) : + (m₁ \ m₂).getD k fallback = fallback := by + revert mem + exact m₁.inductionOn₂ m₂ fun _ _ h => DHashMap.getD_diff_of_mem_right h + +theorem getD_diff_of_not_mem_left [LawfulBEq α] + {k : α} {fallback : β k} (not_mem : k ∉ m₁) : + (m₁ \ m₂).getD k fallback = fallback := by + revert not_mem + exact m₁.inductionOn₂ m₂ fun _ _ h => DHashMap.getD_diff_of_not_mem_left h + +/- get! -/ +theorem get!_diff [LawfulBEq α] {k : α} [Inhabited (β k)] : + (m₁ \ m₂).get! k = + if k ∈ m₂ then default else m₁.get! k := + m₁.inductionOn₂ m₂ fun _ _ => DHashMap.get!_diff + +theorem get!_diff_of_not_mem_right [LawfulBEq α] + {k : α} [Inhabited (β k)] (not_mem : k ∉ m₂) : + (m₁ \ m₂).get! k = m₁.get! k := by + revert not_mem + exact m₁.inductionOn₂ m₂ fun _ _ h => DHashMap.get!_diff_of_not_mem_right h + +theorem get!_diff_of_mem_right [LawfulBEq α] + {k : α} [Inhabited (β k)] (mem : k ∈ m₂) : + (m₁ \ m₂).get! k = default := by + revert mem + exact m₁.inductionOn₂ m₂ fun _ _ h => DHashMap.get!_diff_of_mem_right h + +theorem get!_diff_of_not_mem_left [LawfulBEq α] + {k : α} [Inhabited (β k)] (not_mem : k ∉ m₁) : + (m₁ \ m₂).get! k = default := by + revert not_mem + exact m₁.inductionOn₂ m₂ fun _ _ h => DHashMap.get!_diff_of_not_mem_left h + +/- getKey? -/ +theorem getKey?_diff [EquivBEq α] [LawfulHashable α] {k : α} : + (m₁ \ m₂).getKey? k = + if k ∈ m₂ then none else m₁.getKey? k := + m₁.inductionOn₂ m₂ fun _ _ => DHashMap.getKey?_diff + +theorem getKey?_diff_of_not_mem_right [EquivBEq α] [LawfulHashable α] + {k : α} (not_mem : k ∉ m₂) : + (m₁ \ m₂).getKey? k = m₁.getKey? k := by + revert not_mem + exact m₁.inductionOn₂ m₂ fun _ _ h => DHashMap.getKey?_diff_of_not_mem_right h + +theorem getKey?_diff_of_not_mem_left [EquivBEq α] [LawfulHashable α] + {k : α} (not_mem : k ∉ m₁) : + (m₁ \ m₂).getKey? k = none := by + revert not_mem + exact m₁.inductionOn₂ m₂ fun _ _ h => DHashMap.getKey?_diff_of_not_mem_left h + +theorem getKey?_diff_of_mem_right [EquivBEq α] [LawfulHashable α] + {k : α} (mem : k ∈ m₂) : + (m₁ \ m₂).getKey? k = none := by + revert mem + exact m₁.inductionOn₂ m₂ fun _ _ h => DHashMap.getKey?_diff_of_mem_right h + +/- getKey -/ +@[simp] +theorem getKey_diff [EquivBEq α] [LawfulHashable α] + {k : α} {h_mem : k ∈ m₁ \ m₂} : + (m₁ \ m₂).getKey k h_mem = + m₁.getKey k (mem_diff_iff.1 h_mem).1 := by + induction m₁ + case mk a => + induction m₂ + case mk b => + apply DHashMap.getKey_diff + +/- getKeyD -/ +theorem getKeyD_diff [EquivBEq α] [LawfulHashable α] {k fallback : α} : + (m₁ \ m₂).getKeyD k fallback = + if k ∈ m₂ then fallback else m₁.getKeyD k fallback := + m₁.inductionOn₂ m₂ fun _ _ => DHashMap.getKeyD_diff + +theorem getKeyD_diff_of_not_mem_right [EquivBEq α] [LawfulHashable α] + {k fallback : α} (not_mem : k ∉ m₂) : + (m₁ \ m₂).getKeyD k fallback = m₁.getKeyD k fallback := by + revert not_mem + exact m₁.inductionOn₂ m₂ fun _ _ h => DHashMap.getKeyD_diff_of_not_mem_right h + +theorem getKeyD_diff_of_mem_right [EquivBEq α] [LawfulHashable α] + {k fallback : α} (mem : k ∈ m₂) : + (m₁ \ m₂).getKeyD k fallback = fallback := by + revert mem + exact m₁.inductionOn₂ m₂ fun _ _ h => DHashMap.getKeyD_diff_of_mem_right h + +theorem getKeyD_diff_of_not_mem_left [EquivBEq α] [LawfulHashable α] + {k fallback : α} (not_mem : k ∉ m₁) : + (m₁ \ m₂).getKeyD k fallback = fallback := by + revert not_mem + exact m₁.inductionOn₂ m₂ fun _ _ h => DHashMap.getKeyD_diff_of_not_mem_left h + +/- getKey! -/ +theorem getKey!_diff [EquivBEq α] [LawfulHashable α] [Inhabited α] {k : α} : + (m₁ \ m₂).getKey! k = + if k ∈ m₂ then default else m₁.getKey! k := + m₁.inductionOn₂ m₂ fun _ _ => DHashMap.getKey!_diff + +theorem getKey!_diff_of_not_mem_right [EquivBEq α] [LawfulHashable α] [Inhabited α] + {k : α} (not_mem : k ∉ m₂) : + (m₁ \ m₂).getKey! k = m₁.getKey! k := by + revert not_mem + exact m₁.inductionOn₂ m₂ fun _ _ h => DHashMap.getKey!_diff_of_not_mem_right h + +theorem getKey!_diff_of_mem_right [EquivBEq α] [LawfulHashable α] [Inhabited α] + {k : α} (mem : k ∈ m₂) : + (m₁ \ m₂).getKey! k = default := by + revert mem + exact m₁.inductionOn₂ m₂ fun _ _ h => DHashMap.getKey!_diff_of_mem_right h + +theorem getKey!_diff_of_not_mem_left [EquivBEq α] [LawfulHashable α] [Inhabited α] + {k : α} (not_mem : k ∉ m₁) : + (m₁ \ m₂).getKey! k = default := by + revert not_mem + exact m₁.inductionOn₂ m₂ fun _ _ h => DHashMap.getKey!_diff_of_not_mem_left h + +/- size -/ +theorem size_diff_le_size_left [EquivBEq α] [LawfulHashable α] : + (m₁ \ m₂).size ≤ m₁.size := + m₁.inductionOn₂ m₂ fun _ _ => DHashMap.size_diff_le_size_left + +theorem size_diff_eq_size_left [EquivBEq α] [LawfulHashable α] + (h : ∀ (a : α), a ∈ m₁ → a ∉ m₂) : + (m₁ \ m₂).size = m₁.size := + m₁.inductionOn₂ m₂ (fun _ _ => DHashMap.size_diff_eq_size_left) h + +theorem size_diff_add_size_inter_eq_size_left [EquivBEq α] [LawfulHashable α] : + (m₁ \ m₂).size + (m₁ ∩ m₂).size = m₁.size := + m₁.inductionOn₂ m₂ fun _ _ => DHashMap.size_diff_add_size_inter_eq_size_left + +/- isEmpty -/ +@[simp] +theorem isEmpty_diff_left [EquivBEq α] [LawfulHashable α] (h : m₁.isEmpty) : + (m₁ \ m₂).isEmpty = true := by + revert h + exact m₁.inductionOn₂ m₂ fun _ _ h => DHashMap.isEmpty_diff_left h + +theorem isEmpty_diff_iff [EquivBEq α] [LawfulHashable α] : + (m₁ \ m₂).isEmpty ↔ ∀ k, k ∈ m₁ → k ∈ m₂ := + m₁.inductionOn₂ m₂ fun _ _ => DHashMap.isEmpty_diff_iff + +end Diff + +namespace Const + +variable {β : Type v} {m₁ m₂ : ExtDHashMap α (fun _ => β)} + +/- get? -/ +theorem get?_diff [EquivBEq α] [LawfulHashable α] {k : α} : + Const.get? (m₁ \ m₂) k = + if k ∈ m₂ then none else Const.get? m₁ k := + m₁.inductionOn₂ m₂ fun _ _ => DHashMap.Const.get?_diff + +theorem get?_diff_of_not_mem_right [EquivBEq α] [LawfulHashable α] + {k : α} (not_mem : k ∉ m₂) : + Const.get? (m₁ \ m₂) k = Const.get? m₁ k := by + revert not_mem + exact m₁.inductionOn₂ m₂ fun _ _ h => DHashMap.Const.get?_diff_of_not_mem_right h + +theorem get?_diff_of_not_mem_left [EquivBEq α] [LawfulHashable α] + {k : α} (not_mem : k ∉ m₁) : + Const.get? (m₁ \ m₂) k = none := by + revert not_mem + exact m₁.inductionOn₂ m₂ fun _ _ h => DHashMap.Const.get?_diff_of_not_mem_left h + +theorem get?_diff_of_mem_right [EquivBEq α] [LawfulHashable α] + {k : α} (mem : k ∈ m₂) : + Const.get? (m₁ \ m₂) k = none := by + revert mem + exact m₁.inductionOn₂ m₂ fun _ _ h => DHashMap.Const.get?_diff_of_mem_right h + +/- get -/ +@[simp] +theorem get_diff [EquivBEq α] [LawfulHashable α] + {k : α} {h_mem : k ∈ m₁ \ m₂} : + Const.get (m₁ \ m₂) k h_mem = + Const.get m₁ k (mem_diff_iff.1 h_mem).1 := by + induction m₁ + case mk a => + induction m₂ + case mk b => + apply DHashMap.Const.get_diff + +/- getD -/ +theorem getD_diff [EquivBEq α] [LawfulHashable α] {k : α} {fallback : β} : + Const.getD (m₁ \ m₂) k fallback = + if k ∈ m₂ then fallback else Const.getD m₁ k fallback := + m₁.inductionOn₂ m₂ fun _ _ => DHashMap.Const.getD_diff + +theorem getD_diff_of_not_mem_right [EquivBEq α] [LawfulHashable α] + {k : α} {fallback : β} (not_mem : k ∉ m₂) : + Const.getD (m₁ \ m₂) k fallback = Const.getD m₁ k fallback := by + revert not_mem + exact m₁.inductionOn₂ m₂ fun _ _ h => DHashMap.Const.getD_diff_of_not_mem_right h + +theorem getD_diff_of_mem_right [EquivBEq α] [LawfulHashable α] + {k : α} {fallback : β} (mem : k ∈ m₂) : + Const.getD (m₁ \ m₂) k fallback = fallback := by + revert mem + exact m₁.inductionOn₂ m₂ fun _ _ h => DHashMap.Const.getD_diff_of_mem_right h + +theorem getD_diff_of_not_mem_left [EquivBEq α] [LawfulHashable α] + {k : α} {fallback : β} (not_mem : k ∉ m₁) : + Const.getD (m₁ \ m₂) k fallback = fallback := by + revert not_mem + exact m₁.inductionOn₂ m₂ fun _ _ h => DHashMap.Const.getD_diff_of_not_mem_left h + +/- get! -/ +theorem get!_diff [EquivBEq α] [LawfulHashable α] [Inhabited β] {k : α} : + Const.get! (m₁ \ m₂) k = + if k ∈ m₂ then default else Const.get! m₁ k := + m₁.inductionOn₂ m₂ fun _ _ => DHashMap.Const.get!_diff + +theorem get!_diff_of_not_mem_right [EquivBEq α] [LawfulHashable α] [Inhabited β] + {k : α} (not_mem : k ∉ m₂) : + Const.get! (m₁ \ m₂) k = Const.get! m₁ k := by + revert not_mem + exact m₁.inductionOn₂ m₂ fun _ _ h => DHashMap.Const.get!_diff_of_not_mem_right h + +theorem get!_diff_of_mem_right [EquivBEq α] [LawfulHashable α] [Inhabited β] + {k : α} (mem : k ∈ m₂) : + Const.get! (m₁ \ m₂) k = default := by + revert mem + exact m₁.inductionOn₂ m₂ fun _ _ h => DHashMap.Const.get!_diff_of_mem_right h + +theorem get!_diff_of_not_mem_left [EquivBEq α] [LawfulHashable α] [Inhabited β] + {k : α} (not_mem : k ∉ m₁) : + Const.get! (m₁ \ m₂) k = default := by + revert not_mem + exact m₁.inductionOn₂ m₂ fun _ _ h => DHashMap.Const.get!_diff_of_not_mem_left h + +end Const + variable {m : ExtDHashMap α β} section Alter diff --git a/src/Std/Data/ExtHashMap/Basic.lean b/src/Std/Data/ExtHashMap/Basic.lean index 1695280b0659..7bb48e8f1e0b 100644 --- a/src/Std/Data/ExtHashMap/Basic.lean +++ b/src/Std/Data/ExtHashMap/Basic.lean @@ -254,6 +254,11 @@ def inter [EquivBEq α] [LawfulHashable α] (m₁ m₂ : ExtHashMap α β) : Ext instance [EquivBEq α] [LawfulHashable α] : Inter (ExtHashMap α β) := ⟨inter⟩ +@[inline, inherit_doc ExtDHashMap.inter] +def diff [EquivBEq α] [LawfulHashable α] (m₁ m₂ : ExtHashMap α β) : ExtHashMap α β := ⟨ExtDHashMap.diff m₁.inner m₂.inner⟩ + +instance [EquivBEq α] [LawfulHashable α] : SDiff (ExtHashMap α β) := ⟨diff⟩ + @[inline, inherit_doc ExtDHashMap.Const.unitOfArray] def unitOfArray [BEq α] [Hashable α] (l : Array α) : ExtHashMap α Unit := diff --git a/src/Std/Data/ExtHashMap/Lemmas.lean b/src/Std/Data/ExtHashMap/Lemmas.lean index c1d237aadf09..54431353cd5d 100644 --- a/src/Std/Data/ExtHashMap/Lemmas.lean +++ b/src/Std/Data/ExtHashMap/Lemmas.lean @@ -1761,6 +1761,207 @@ theorem isEmpty_inter_iff [EquivBEq α] [LawfulHashable α] : end Inter +section Diff + +variable (m₁ m₂ : ExtHashMap α β) + +variable {m₁ m₂} + +@[simp] +theorem diff_eq [EquivBEq α] [LawfulHashable α] : m₁.diff m₂ = m₁ \ m₂ := by + simp only [SDiff.sdiff] + +/- contains -/ +@[simp] +theorem contains_diff [EquivBEq α] [LawfulHashable α] + {k : α} : + (m₁ \ m₂).contains k = (m₁.contains k && !m₂.contains k) := + ExtDHashMap.contains_diff + +/- mem -/ +@[simp] +theorem mem_diff_iff [EquivBEq α] [LawfulHashable α] {k : α} : + k ∈ m₁ \ m₂ ↔ k ∈ m₁ ∧ k ∉ m₂ := + ExtDHashMap.mem_diff_iff + +theorem not_mem_diff_of_not_mem_left [EquivBEq α] [LawfulHashable α] {k : α} + (not_mem : k ∉ m₁) : + k ∉ m₁ \ m₂ := + ExtDHashMap.not_mem_diff_of_not_mem_left not_mem + +theorem not_mem_diff_of_mem_right [EquivBEq α] [LawfulHashable α] {k : α} + (mem : k ∈ m₂) : + k ∉ m₁ \ m₂ := + ExtDHashMap.not_mem_diff_of_mem_right mem + +/- get? -/ +theorem get?_diff [EquivBEq α] [LawfulHashable α] {k : α} : + (m₁ \ m₂).get? k = + if k ∈ m₂ then none else m₁.get? k := + ExtDHashMap.Const.get?_diff + +theorem get?_diff_of_not_mem_right [EquivBEq α] [LawfulHashable α] + {k : α} (not_mem : k ∉ m₂) : + (m₁ \ m₂).get? k = m₁.get? k := + ExtDHashMap.Const.get?_diff_of_not_mem_right not_mem + +theorem get?_diff_of_not_mem_left [EquivBEq α] [LawfulHashable α] + {k : α} (not_mem : k ∉ m₁) : + (m₁ \ m₂).get? k = none := + ExtDHashMap.Const.get?_diff_of_not_mem_left not_mem + +theorem get?_diff_of_mem_right [EquivBEq α] [LawfulHashable α] + {k : α} (mem : k ∈ m₂) : + (m₁ \ m₂).get? k = none := + ExtDHashMap.Const.get?_diff_of_mem_right mem + +/- get -/ +@[simp] +theorem get_diff [EquivBEq α] [LawfulHashable α] + {k : α} {h_mem : k ∈ m₁ \ m₂} : + (m₁ \ m₂).get k h_mem = + m₁.get k (mem_diff_iff.1 h_mem).1 := + ExtDHashMap.Const.get_diff + +/- getD -/ +theorem getD_diff [EquivBEq α] [LawfulHashable α] {k : α} {fallback : β} : + (m₁ \ m₂).getD k fallback = + if k ∈ m₂ then fallback else m₁.getD k fallback := + ExtDHashMap.Const.getD_diff + +theorem getD_diff_of_not_mem_right [EquivBEq α] [LawfulHashable α] + {k : α} {fallback : β} (not_mem : k ∉ m₂) : + (m₁ \ m₂).getD k fallback = m₁.getD k fallback := + ExtDHashMap.Const.getD_diff_of_not_mem_right not_mem + +theorem getD_diff_of_mem_right [EquivBEq α] [LawfulHashable α] + {k : α} {fallback : β} (mem : k ∈ m₂) : + (m₁ \ m₂).getD k fallback = fallback := + ExtDHashMap.Const.getD_diff_of_mem_right mem + +theorem getD_diff_of_not_mem_left [EquivBEq α] [LawfulHashable α] + {k : α} {fallback : β} (not_mem : k ∉ m₁) : + (m₁ \ m₂).getD k fallback = fallback := + ExtDHashMap.Const.getD_diff_of_not_mem_left not_mem + +/- get! -/ +theorem get!_diff [EquivBEq α] [LawfulHashable α] {k : α} [Inhabited β] : + (m₁ \ m₂).get! k = + if k ∈ m₂ then default else m₁.get! k := + ExtDHashMap.Const.get!_diff + +theorem get!_diff_of_not_mem_right [EquivBEq α] [LawfulHashable α] + {k : α} [Inhabited β] (not_mem : k ∉ m₂) : + (m₁ \ m₂).get! k = m₁.get! k := + ExtDHashMap.Const.get!_diff_of_not_mem_right not_mem + +theorem get!_diff_of_mem_right [EquivBEq α] [LawfulHashable α] + {k : α} [Inhabited β] (mem : k ∈ m₂) : + (m₁ \ m₂).get! k = default := + ExtDHashMap.Const.get!_diff_of_mem_right mem + +theorem get!_diff_of_not_mem_left [EquivBEq α] [LawfulHashable α] + {k : α} [Inhabited β] (not_mem : k ∉ m₁) : + (m₁ \ m₂).get! k = default := + ExtDHashMap.Const.get!_diff_of_not_mem_left not_mem + +/- getKey? -/ +theorem getKey?_diff [EquivBEq α] [LawfulHashable α] {k : α} : + (m₁ \ m₂).getKey? k = + if k ∈ m₂ then none else m₁.getKey? k := + ExtDHashMap.getKey?_diff + +theorem getKey?_diff_of_not_mem_right [EquivBEq α] [LawfulHashable α] + {k : α} (not_mem : k ∉ m₂) : + (m₁ \ m₂).getKey? k = m₁.getKey? k := + ExtDHashMap.getKey?_diff_of_not_mem_right not_mem + +theorem getKey?_diff_of_not_mem_left [EquivBEq α] [LawfulHashable α] + {k : α} (not_mem : k ∉ m₁) : + (m₁ \ m₂).getKey? k = none := + ExtDHashMap.getKey?_diff_of_not_mem_left not_mem + +theorem getKey?_diff_of_mem_right [EquivBEq α] [LawfulHashable α] + {k : α} (mem : k ∈ m₂) : + (m₁ \ m₂).getKey? k = none := + ExtDHashMap.getKey?_diff_of_mem_right mem + +/- getKey -/ +@[simp] +theorem getKey_diff [EquivBEq α] [LawfulHashable α] + {k : α} {h_mem : k ∈ m₁ \ m₂} : + (m₁ \ m₂).getKey k h_mem = + m₁.getKey k (mem_diff_iff.1 h_mem).1 := + ExtDHashMap.getKey_diff + +/- getKeyD -/ +theorem getKeyD_diff [EquivBEq α] [LawfulHashable α] {k fallback : α} : + (m₁ \ m₂).getKeyD k fallback = + if k ∈ m₂ then fallback else m₁.getKeyD k fallback := + ExtDHashMap.getKeyD_diff + +theorem getKeyD_diff_of_not_mem_right [EquivBEq α] [LawfulHashable α] + {k fallback : α} (not_mem : k ∉ m₂) : + (m₁ \ m₂).getKeyD k fallback = m₁.getKeyD k fallback := + ExtDHashMap.getKeyD_diff_of_not_mem_right not_mem + +theorem getKeyD_diff_of_mem_right [EquivBEq α] [LawfulHashable α] + {k fallback : α} (mem : k ∈ m₂) : + (m₁ \ m₂).getKeyD k fallback = fallback := + ExtDHashMap.getKeyD_diff_of_mem_right mem + +theorem getKeyD_diff_of_not_mem_left [EquivBEq α] [LawfulHashable α] + {k fallback : α} (not_mem : k ∉ m₁) : + (m₁ \ m₂).getKeyD k fallback = fallback := + ExtDHashMap.getKeyD_diff_of_not_mem_left not_mem + +/- getKey! -/ +theorem getKey!_diff [EquivBEq α] [LawfulHashable α] [Inhabited α] {k : α} : + (m₁ \ m₂).getKey! k = + if k ∈ m₂ then default else m₁.getKey! k := + ExtDHashMap.getKey!_diff + +theorem getKey!_diff_of_not_mem_right [EquivBEq α] [LawfulHashable α] [Inhabited α] + {k : α} (not_mem : k ∉ m₂) : + (m₁ \ m₂).getKey! k = m₁.getKey! k := + ExtDHashMap.getKey!_diff_of_not_mem_right not_mem + +theorem getKey!_diff_of_mem_right [EquivBEq α] [LawfulHashable α] [Inhabited α] + {k : α} (mem : k ∈ m₂) : + (m₁ \ m₂).getKey! k = default := + ExtDHashMap.getKey!_diff_of_mem_right mem + +theorem getKey!_diff_of_not_mem_left [EquivBEq α] [LawfulHashable α] [Inhabited α] + {k : α} (not_mem : k ∉ m₁) : + (m₁ \ m₂).getKey! k = default := + ExtDHashMap.getKey!_diff_of_not_mem_left not_mem + +/- size -/ +theorem size_diff_le_size_left [EquivBEq α] [LawfulHashable α] : + (m₁ \ m₂).size ≤ m₁.size := + ExtDHashMap.size_diff_le_size_left + +theorem size_diff_eq_size_left [EquivBEq α] [LawfulHashable α] + (h : ∀ (a : α), a ∈ m₁ → a ∉ m₂) : + (m₁ \ m₂).size = m₁.size := + ExtDHashMap.size_diff_eq_size_left h + +theorem size_diff_add_size_inter_eq_size_left [EquivBEq α] [LawfulHashable α] : + (m₁ \ m₂).size + (m₁ ∩ m₂).size = m₁.size := + ExtDHashMap.size_diff_add_size_inter_eq_size_left + +/- isEmpty -/ +@[simp] +theorem isEmpty_diff_left [EquivBEq α] [LawfulHashable α] (h : m₁.isEmpty) : + (m₁ \ m₂).isEmpty = true := + ExtDHashMap.isEmpty_diff_left h + +theorem isEmpty_diff_iff [EquivBEq α] [LawfulHashable α] : + (m₁ \ m₂).isEmpty ↔ ∀ k, k ∈ m₁ → k ∈ m₂ := + ExtDHashMap.isEmpty_diff_iff + +end Diff + section Alter variable {m : ExtHashMap α β} From 76282638b9687b8bb9653c08e5a1b24070d4ef4a Mon Sep 17 00:00:00 2001 From: Wojciech Rozowski Date: Thu, 27 Nov 2025 12:41:26 +0000 Subject: [PATCH 88/88] Add the remaining ExtHashSet lemmas --- src/Std/Data/ExtHashMap/Basic.lean | 2 +- src/Std/Data/ExtHashSet/Basic.lean | 10 +++ src/Std/Data/ExtHashSet/Lemmas.lean | 127 ++++++++++++++++++++++++++++ 3 files changed, 138 insertions(+), 1 deletion(-) diff --git a/src/Std/Data/ExtHashMap/Basic.lean b/src/Std/Data/ExtHashMap/Basic.lean index 7bb48e8f1e0b..dabe67d7ead7 100644 --- a/src/Std/Data/ExtHashMap/Basic.lean +++ b/src/Std/Data/ExtHashMap/Basic.lean @@ -254,7 +254,7 @@ def inter [EquivBEq α] [LawfulHashable α] (m₁ m₂ : ExtHashMap α β) : Ext instance [EquivBEq α] [LawfulHashable α] : Inter (ExtHashMap α β) := ⟨inter⟩ -@[inline, inherit_doc ExtDHashMap.inter] +@[inline, inherit_doc ExtDHashMap.diff] def diff [EquivBEq α] [LawfulHashable α] (m₁ m₂ : ExtHashMap α β) : ExtHashMap α β := ⟨ExtDHashMap.diff m₁.inner m₂.inner⟩ instance [EquivBEq α] [LawfulHashable α] : SDiff (ExtHashMap α β) := ⟨diff⟩ diff --git a/src/Std/Data/ExtHashSet/Basic.lean b/src/Std/Data/ExtHashSet/Basic.lean index 43590d675247..7d3888462890 100644 --- a/src/Std/Data/ExtHashSet/Basic.lean +++ b/src/Std/Data/ExtHashSet/Basic.lean @@ -214,6 +214,16 @@ def inter [EquivBEq α] [LawfulHashable α] (m₁ m₂ : ExtHashSet α) : ExtHas instance [EquivBEq α] [LawfulHashable α] : Inter (ExtHashSet α) := ⟨inter⟩ +/-- +Computes the difference of the given hash sets. + +This function always iterates through the smaller set. +-/ +@[inline] +def diff [EquivBEq α] [LawfulHashable α] (m₁ m₂ : ExtHashSet α) : ExtHashSet α := ⟨ExtHashMap.diff m₁.inner m₂.inner⟩ + +instance [EquivBEq α] [LawfulHashable α] : SDiff (ExtHashSet α) := ⟨diff⟩ + /-- 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/ExtHashSet/Lemmas.lean b/src/Std/Data/ExtHashSet/Lemmas.lean index ddf512a06193..d3716e23fa62 100644 --- a/src/Std/Data/ExtHashSet/Lemmas.lean +++ b/src/Std/Data/ExtHashSet/Lemmas.lean @@ -1017,4 +1017,131 @@ theorem isEmpty_inter_iff [EquivBEq α] [LawfulHashable α] : end Inter +section Diff + +variable {m₁ m₂ : ExtHashSet α} + +@[simp] +theorem diff_eq [EquivBEq α] [LawfulHashable α] : m₁.diff m₂ = m₁ \ m₂ := by + simp only [SDiff.sdiff] + +/- contains -/ +@[simp] +theorem contains_diff [EquivBEq α] [LawfulHashable α] {k : α} : + (m₁ \ m₂).contains k = (m₁.contains k && !m₂.contains k) := + ExtHashMap.contains_diff + +/- mem -/ +@[simp] +theorem mem_diff_iff [EquivBEq α] [LawfulHashable α] {k : α} : + k ∈ m₁ \ m₂ ↔ k ∈ m₁ ∧ k ∉ m₂ := + ExtHashMap.mem_diff_iff + +theorem not_mem_diff_of_not_mem_left [EquivBEq α] [LawfulHashable α] {k : α} + (not_mem : k ∉ m₁) : + k ∉ m₁ \ m₂ := + ExtHashMap.not_mem_diff_of_not_mem_left not_mem + +theorem not_mem_diff_of_mem_right [EquivBEq α] [LawfulHashable α] {k : α} + (mem : k ∈ m₂) : + k ∉ m₁ \ m₂ := + ExtHashMap.not_mem_diff_of_mem_right mem + +/- get? -/ +theorem get?_diff [EquivBEq α] [LawfulHashable α] {k : α} : + (m₁ \ m₂).get? k = + if k ∈ m₂ then none else m₁.get? k := + ExtHashMap.getKey?_diff + +theorem get?_diff_of_not_mem_right [EquivBEq α] [LawfulHashable α] + {k : α} (not_mem : k ∉ m₂) : + (m₁ \ m₂).get? k = m₁.get? k := + ExtHashMap.getKey?_diff_of_not_mem_right not_mem + +theorem get?_diff_of_not_mem_left [EquivBEq α] [LawfulHashable α] + {k : α} (not_mem : k ∉ m₁) : + (m₁ \ m₂).get? k = none := + ExtHashMap.getKey?_diff_of_not_mem_left not_mem + +theorem get?_diff_of_mem_right [EquivBEq α] [LawfulHashable α] + {k : α} (mem : k ∈ m₂) : + (m₁ \ m₂).get? k = none := + ExtHashMap.getKey?_diff_of_mem_right mem + +/- get -/ +@[simp] +theorem get_diff [EquivBEq α] [LawfulHashable α] + {k : α} {h_mem : k ∈ m₁ \ m₂} : + (m₁ \ m₂).get k h_mem = + m₁.get k (mem_diff_iff.1 h_mem).1 := + ExtHashMap.getKey_diff + +/- getD -/ +theorem getD_diff [EquivBEq α] [LawfulHashable α] {k fallback : α} : + (m₁ \ m₂).getD k fallback = + if k ∈ m₂ then fallback else m₁.getD k fallback := + ExtHashMap.getKeyD_diff + +theorem getD_diff_of_not_mem_right [EquivBEq α] [LawfulHashable α] + {k fallback : α} (not_mem : k ∉ m₂) : + (m₁ \ m₂).getD k fallback = m₁.getD k fallback := + ExtHashMap.getKeyD_diff_of_not_mem_right not_mem + +theorem getD_diff_of_mem_right [EquivBEq α] [LawfulHashable α] + {k fallback : α} (mem : k ∈ m₂) : + (m₁ \ m₂).getD k fallback = fallback := + ExtHashMap.getKeyD_diff_of_mem_right mem + +theorem getD_diff_of_not_mem_left [EquivBEq α] [LawfulHashable α] + {k fallback : α} (not_mem : k ∉ m₁) : + (m₁ \ m₂).getD k fallback = fallback := + ExtHashMap.getKeyD_diff_of_not_mem_left not_mem + +/- get! -/ +theorem get!_diff [EquivBEq α] [LawfulHashable α] [Inhabited α] {k : α} : + (m₁ \ m₂).get! k = + if k ∈ m₂ then default else m₁.get! k := + ExtHashMap.getKey!_diff + +theorem get!_diff_of_not_mem_right [EquivBEq α] [LawfulHashable α] [Inhabited α] + {k : α} (not_mem : k ∉ m₂) : + (m₁ \ m₂).get! k = m₁.get! k := + ExtHashMap.getKey!_diff_of_not_mem_right not_mem + +theorem get!_diff_of_mem_right [EquivBEq α] [LawfulHashable α] [Inhabited α] + {k : α} (mem : k ∈ m₂) : + (m₁ \ m₂).get! k = default := + ExtHashMap.getKey!_diff_of_mem_right mem + +theorem get!_diff_of_not_mem_left [EquivBEq α] [LawfulHashable α] [Inhabited α] + {k : α} (not_mem : k ∉ m₁) : + (m₁ \ m₂).get! k = default := + ExtHashMap.getKey!_diff_of_not_mem_left not_mem + +/- size -/ +theorem size_diff_le_size_left [EquivBEq α] [LawfulHashable α] : + (m₁ \ m₂).size ≤ m₁.size := + ExtHashMap.size_diff_le_size_left + +theorem size_diff_eq_size_left [EquivBEq α] [LawfulHashable α] + (h : ∀ (a : α), a ∈ m₁ → a ∉ m₂) : + (m₁ \ m₂).size = m₁.size := + ExtHashMap.size_diff_eq_size_left h + +theorem size_diff_add_size_inter_eq_size_left [EquivBEq α] [LawfulHashable α] : + (m₁ \ m₂).size + (m₁ ∩ m₂).size = m₁.size := + ExtHashMap.size_diff_add_size_inter_eq_size_left + +/- isEmpty -/ +@[simp] +theorem isEmpty_diff_left [EquivBEq α] [LawfulHashable α] (h : m₁.isEmpty) : + (m₁ \ m₂).isEmpty = true := + ExtHashMap.isEmpty_diff_left h + +theorem isEmpty_diff_iff [EquivBEq α] [LawfulHashable α] : + (m₁ \ m₂).isEmpty ↔ ∀ k, k ∈ m₁ → k ∈ m₂ := + ExtHashMap.isEmpty_diff_iff + +end Diff + end Std.ExtHashSet