@@ -46,8 +46,8 @@ private meta def baseNames : Array Name :=
4646 ``get ?_eq, ``get_eq , ``get !_eq, ``getD_eq ,
4747 ``Const .get?_eq, ``Const .get_eq, ``Const .getD_eq, ``Const .get!_eq,
4848 ``getKey ?_eq, ``getKey_eq , ``getKey !_eq, ``getKeyD_eq ,
49- ``insertMany_eq , ``Const .insertMany_eq, ``Const .insertManyIfNewUnit_eq,
50- ``ofList_eq , ``Const .ofList_eq, ``Const .unitOfList_eq,
49+ ``insertMany_eq , ``Const .insertMany_eq, ``Const .insertManyIfNewUnit_eq, ``ofArray_eq , ``Const .ofArray_eq,
50+ ``ofList_eq , ``Const .ofList_eq, ``Const .unitOfList_eq, ``Const .unitOfArray_eq,
5151 ``alter_eq , ``Const .alter_eq, ``modify_eq , ``Const .modify_eq, ``union_eq , ``inter_eq ]
5252
5353/-- Internal implementation detail of the hash map -/
@@ -3340,6 +3340,12 @@ variable [BEq α] [Hashable α]
33403340
33413341open Internal.Raw Internal.Raw₀
33423342
3343+ @[simp, grind =]
3344+ theorem ofArray_eq_ofList (a : Array ((a : α) × (β a))) :
3345+ ofArray a = ofList a.toList := by
3346+ simp_to_raw
3347+ rw [Raw₀.insertMany_array_eq_insertMany_toList]
3348+
33433349@[simp, grind =]
33443350theorem ofList_nil :
33453351 ofList ([] : List ((a : α) × (β a))) = ∅ := by
@@ -3492,6 +3498,12 @@ namespace Const
34923498
34933499variable {β : Type v}
34943500
3501+ @[simp, grind =]
3502+ theorem ofArray_eq_ofList (a : Array (α × β)) :
3503+ ofArray a = ofList a.toList := by
3504+ simp_to_raw
3505+ rw [Const.insertMany_array_eq_insertMany_toList]
3506+
34953507@[simp, grind =]
34963508theorem ofList_nil :
34973509 ofList ([] : List (α × β)) = ∅ := by
@@ -3640,6 +3652,12 @@ theorem isEmpty_ofList [EquivBEq α] [LawfulHashable α]
36403652 (ofList l).isEmpty = l.isEmpty := by
36413653 simp_to_raw using Raw₀.Const.isEmpty_insertMany_emptyWithCapacity_list
36423654
3655+ @[simp, grind =]
3656+ theorem unitOfArray_eq_unitOfList (a : Array α) :
3657+ unitOfArray a = unitOfList a.toList := by
3658+ simp_to_raw
3659+ rw [Raw₀.Const.insertManyIfNewUnit_array_eq_insertManyIfNewUnit_toList]
3660+
36433661@[simp]
36443662theorem unitOfList_nil :
36453663 unitOfList ([] : List α) = ∅ := by
0 commit comments