File tree Expand file tree Collapse file tree 9 files changed +160
-0
lines changed
Expand file tree Collapse file tree 9 files changed +160
-0
lines changed Original file line number Diff line number Diff line change @@ -75,6 +75,22 @@ theorem size_emptyWithCapacity {c} : (emptyWithCapacity c : Raw₀ α β).1.size
7575theorem isEmpty_eq_size_eq_zero : m.1 .isEmpty = (m.1 .size == 0 ) := by
7676 simp [Raw.isEmpty]
7777
78+ @[simp]
79+ theorem toList_emptyWithCapacity {c} : (emptyWithCapacity c : Raw₀ α β).1 .toList = [] := by
80+ rw [Raw.toList_eq_toListModel, toListModel_buckets_emptyWithCapacity]
81+
82+ @[simp]
83+ theorem Const.toList_emptyWithCapacity {c} {β : Type v} : Raw.Const.toList (emptyWithCapacity c : Raw₀ α (fun _ => β)).1 = [] := by
84+ rw [Raw.Const.toList_eq_toListModel_map, toListModel_buckets_emptyWithCapacity, List.map_nil]
85+
86+ @[simp]
87+ theorem keys_emptyWithCapacity {c} : (emptyWithCapacity c : Raw₀ α β).1 .keys = [] := by
88+ rw [Raw.keys_eq_keys_toListModel, toListModel_buckets_emptyWithCapacity, List.keys_nil]
89+
90+ @[simp]
91+ theorem Const.values_emptyWithCapacity {c} {β : Type v} : (emptyWithCapacity c : Raw₀ α (fun _ => β)).1 .values = [] := by
92+ rw [Raw.values_eq_values_toListModel, toListModel_buckets_emptyWithCapacity, List.values_nil]
93+
7894variable [BEq α] [Hashable α]
7995
8096/-- Internal implementation detail of the hash map -/
Original file line number Diff line number Diff line change @@ -131,6 +131,11 @@ theorem foldRev_cons_key {l : Raw α β} {acc : List α} :
131131 List.keys (toListModel l.buckets) ++ acc := by
132132 rw [foldRev_cons_apply, keys_eq_map]
133133
134+ theorem foldRev_cons_value {β : Type v} {l : Raw α (fun _ => β)} {acc : List β} :
135+ Raw.Internal.foldRev (fun acc _ v => v :: acc) acc l =
136+ List.values (toListModel l.buckets) ++ acc := by
137+ rw [foldRev_cons_apply, values_eq_map]
138+
134139theorem fold_push {l : Raw α β} {acc : Array ((a : α) × β a)} :
135140 Raw.fold (fun acc k v => acc.push ⟨k, v⟩) acc l = acc ++ (toListModel l.buckets).toArray := by
136141 simp [fold_push_apply]
@@ -207,6 +212,10 @@ theorem keys_eq_keys_toListModel {m : Raw α β} :
207212 m.keys = List.keys (toListModel m.buckets) := by
208213 simp [Raw.keys, foldRev_cons_key, keys_eq_map]
209214
215+ theorem values_eq_values_toListModel {β : Type v} {m : Raw α (fun _ => β)} :
216+ m.values = List.values (toListModel m.buckets) := by
217+ simp [Raw.values, foldRev_cons_value, values_eq_map]
218+
210219theorem keysArray_eq_toArray_keys_toListModel {m : Raw α β} :
211220 m.keysArray = (List.keys (toListModel m.buckets)).toArray := by
212221 simp [Raw.keysArray, fold_push_key]
Original file line number Diff line number Diff line change @@ -148,6 +148,38 @@ theorem size_empty : (∅ : DHashMap α β).size = 0 :=
148148
149149theorem isEmpty_eq_size_eq_zero : m.isEmpty = (m.size == 0 ) := (rfl)
150150
151+ @[simp]
152+ theorem toList_emptyWithCapacity {c} : (emptyWithCapacity c : DHashMap α β).toList = [] :=
153+ Raw₀.toList_emptyWithCapacity
154+
155+ @[simp]
156+ theorem toList_empty : (∅ : DHashMap α β).toList = [] :=
157+ toList_emptyWithCapacity
158+
159+ @[simp]
160+ theorem Const.toList_emptyWithCapacity {c} {β : Type v}: Const.toList (emptyWithCapacity c : DHashMap α (fun _ => β)) = [] :=
161+ Raw₀.Const.toList_emptyWithCapacity
162+
163+ @[simp]
164+ theorem Const.toList_empty {β : Type v} : Const.toList (∅ : DHashMap α (fun _ => β)) = [] :=
165+ Const.toList_emptyWithCapacity
166+
167+ @[simp]
168+ theorem keys_emptyWithCapacity {c} : (emptyWithCapacity c : DHashMap α β).keys = [] :=
169+ Raw₀.keys_emptyWithCapacity
170+
171+ @[simp]
172+ theorem keys_empty : (∅ : DHashMap α β).keys = [] :=
173+ keys_emptyWithCapacity
174+
175+ @[simp]
176+ theorem Const.values_emptyWithCapacity {c} {β : Type v} : (emptyWithCapacity c : DHashMap α (fun _ => β)).values = [] :=
177+ Raw₀.Const.values_emptyWithCapacity
178+
179+ @[simp]
180+ theorem Const.values_empty {β : Type v} : (∅ : DHashMap α (fun _ => β)).values = [] :=
181+ Const.values_emptyWithCapacity
182+
151183@[grind =] theorem size_insert [EquivBEq α] [LawfulHashable α] {k : α} {v : β k} :
152184 (m.insert k v).size = if k ∈ m then m.size else m.size + 1 :=
153185 Raw₀.size_insert ⟨m.1 , _⟩ m.2
Original file line number Diff line number Diff line change @@ -81,6 +81,38 @@ theorem size_empty : (∅ : Raw α β).size = 0 :=
8181theorem isEmpty_eq_size_eq_zero : m.isEmpty = (m.size == 0 ) := by
8282 simp [isEmpty]
8383
84+ @[simp]
85+ theorem toList_emptyWithCapacity {c} : (emptyWithCapacity c : Raw α β).toList = [] := by
86+ simp_to_raw using Raw₀.toList_emptyWithCapacity
87+
88+ @[simp]
89+ theorem Const.toList_emptyWithCapacity {β : Type v} {c} : Const.toList (emptyWithCapacity c : Raw α (fun _ => β)) = [] := by
90+ simp_to_raw using Raw₀.Const.toList_emptyWithCapacity
91+
92+ @[simp]
93+ theorem toList_empty : (∅ : Raw α β).toList = [] :=
94+ toList_emptyWithCapacity
95+
96+ @[simp]
97+ theorem Const.toList_empty {β : Type v} : Const.toList (∅ : Raw α (fun _ => β)) = [] :=
98+ Const.toList_emptyWithCapacity
99+
100+ @[simp]
101+ theorem keys_emptyWithCapacity {c} : (emptyWithCapacity c : Raw α β).keys = [] := by
102+ simp_to_raw using Raw₀.keys_emptyWithCapacity
103+
104+ @[simp]
105+ theorem keys_empty : (∅ : Raw α β).keys = [] :=
106+ keys_emptyWithCapacity
107+
108+ @[simp]
109+ theorem Const.values_emptyWithCapacity {c} {β : Type v} : (emptyWithCapacity c : Raw α (fun _ => β)).values = [] := by
110+ simp_to_raw using Raw₀.Const.values_emptyWithCapacity
111+
112+ @[simp]
113+ theorem Const.values_empty {β : Type v} : (∅ : Raw α (fun _ => β)).values = [] :=
114+ Const.values_emptyWithCapacity
115+
84116variable [BEq α] [Hashable α]
85117
86118@[simp, grind =]
Original file line number Diff line number Diff line change @@ -147,6 +147,30 @@ theorem size_empty : (∅ : HashMap α β).size = 0 :=
147147theorem isEmpty_eq_size_eq_zero : m.isEmpty = (m.size == 0 ) :=
148148 DHashMap.isEmpty_eq_size_eq_zero
149149
150+ @[simp]
151+ theorem toList_emptyWithCapacity {c} : (emptyWithCapacity c : HashMap α β).toList = [] :=
152+ DHashMap.Const.toList_emptyWithCapacity
153+
154+ @[simp]
155+ theorem toList_empty : (∅ : HashMap α β).toList = [] :=
156+ toList_emptyWithCapacity
157+
158+ @[simp]
159+ theorem keys_emptyWithCapacity {c} : (emptyWithCapacity c : HashMap α β).keys = [] :=
160+ DHashMap.keys_emptyWithCapacity
161+
162+ @[simp]
163+ theorem keys_empty : (∅ : HashMap α β).keys = [] :=
164+ keys_emptyWithCapacity
165+
166+ @[simp]
167+ theorem values_emptyWithCapacity {c} {β : Type v} : (emptyWithCapacity c : HashMap α β).values = [] :=
168+ DHashMap.Const.values_emptyWithCapacity
169+
170+ @[simp]
171+ theorem values_empty {β : Type v} : (∅ : HashMap α β).values = [] :=
172+ values_emptyWithCapacity
173+
150174@[grind =] theorem size_insert [EquivBEq α] [LawfulHashable α] {k : α} {v : β} :
151175 (m.insert k v).size = if k ∈ m then m.size else m.size + 1 :=
152176 DHashMap.size_insert
Original file line number Diff line number Diff line change @@ -44,6 +44,30 @@ theorem size_empty : (∅ : Raw α β).size = 0 :=
4444theorem isEmpty_eq_size_eq_zero : m.isEmpty = (m.size == 0 ) :=
4545 DHashMap.Raw.isEmpty_eq_size_eq_zero
4646
47+ @[simp]
48+ theorem toList_emptyWithCapacity {c} : (emptyWithCapacity c : Raw α β).toList = [] :=
49+ DHashMap.Raw.Const.toList_emptyWithCapacity
50+
51+ @[simp]
52+ theorem toList_empty : (∅ : Raw α β).toList = [] :=
53+ toList_emptyWithCapacity
54+
55+ @[simp]
56+ theorem keys_emptyWithCapacity {c} : (emptyWithCapacity c : Raw α β).keys = [] :=
57+ DHashMap.Raw.keys_emptyWithCapacity
58+
59+ @[simp]
60+ theorem keys_empty : (∅ : Raw α β).keys = [] :=
61+ keys_emptyWithCapacity
62+
63+ @[simp]
64+ theorem values_emptyWithCapacity {c} {β : Type v} : (emptyWithCapacity c : Raw α β).values = [] :=
65+ DHashMap.Raw.Const.values_emptyWithCapacity
66+
67+ @[simp]
68+ theorem values_empty {β : Type v} : (∅ : Raw α β).values = [] :=
69+ values_emptyWithCapacity
70+
4771private theorem ext {m m' : Raw α β} : m.inner = m'.inner → m = m' := by
4872 cases m; cases m'; rintro rfl; rfl
4973
Original file line number Diff line number Diff line change @@ -151,6 +151,14 @@ theorem size_empty : (∅ : HashSet α).size = 0 :=
151151theorem isEmpty_eq_size_eq_zero : m.isEmpty = (m.size == 0 ) :=
152152 HashMap.isEmpty_eq_size_eq_zero
153153
154+ @[simp]
155+ theorem toList_emptyWithCapacity {c} : (emptyWithCapacity c : HashSet α).toList = [] :=
156+ HashMap.keys_emptyWithCapacity
157+
158+ @[simp]
159+ theorem toList_empty : (∅ : HashSet α).toList = [] :=
160+ toList_emptyWithCapacity
161+
154162@[grind =] theorem size_insert [EquivBEq α] [LawfulHashable α] {k : α} :
155163 (m.insert k).size = if k ∈ m then m.size else m.size + 1 :=
156164 HashMap.size_insertIfNew
Original file line number Diff line number Diff line change @@ -46,6 +46,14 @@ theorem size_empty : (∅ : Raw α).size = 0 :=
4646theorem isEmpty_eq_size_eq_zero : m.isEmpty = (m.size == 0 ) :=
4747 HashMap.Raw.isEmpty_eq_size_eq_zero
4848
49+ @[simp]
50+ theorem toList_emptyWithCapacity {c} : (emptyWithCapacity c : Raw α).toList = [] :=
51+ HashMap.Raw.keys_emptyWithCapacity
52+
53+ @[simp]
54+ theorem toList_empty : (∅ : Raw α).toList = [] :=
55+ toList_emptyWithCapacity
56+
4957variable [BEq α] [Hashable α]
5058
5159@[simp, grind =]
Original file line number Diff line number Diff line change @@ -162,6 +162,12 @@ theorem keys_eq_map {l : List ((a : α) × β a)} :
162162 | nil => rfl
163163 | cons => simp [keys, *]
164164
165+ theorem values_eq_map {β : Type v} {l : List ((_ : α) × β)} :
166+ values l = l.map (·.2 ) := by
167+ induction l with
168+ | nil => rfl
169+ | cons => simp [values, *]
170+
165171theorem getEntry ?_eq_some_iff [BEq α] [EquivBEq α] {l : List ((a : α) × β a)} {e} {k}
166172 (hd : DistinctKeys l) :
167173 getEntry? k l = some e ↔ k == e.1 ∧ e ∈ l := by
@@ -393,6 +399,7 @@ theorem containsKey_eq_contains_map_fst [BEq α] [PartialEquivBEq α] {l : List
393399 rw [BEq.comm]
394400
395401@[simp] theorem keys_nil : keys ([] : List ((a : α) × β a)) = [] := (rfl)
402+ @[simp] theorem values_nil {β : Type v} : values ([] : List ((_ : α) × β )) = [] := (rfl)
396403@[simp] theorem keys_cons {l : List ((a : α) × β a)} {k : α} {v : β k} :
397404 keys (⟨k, v⟩ :: l) = k :: keys l := (rfl)
398405
You can’t perform that action at this time.
0 commit comments