Skip to content

Commit

Permalink
chore: add missing deprecation dates (#5884)
Browse files Browse the repository at this point in the history
  • Loading branch information
kim-em authored Oct 30, 2024
1 parent 09802e8 commit 38c3948
Show file tree
Hide file tree
Showing 6 changed files with 18 additions and 18 deletions.
2 changes: 1 addition & 1 deletion src/Init/Data/BitVec/Lemmas.lean
Original file line number Diff line number Diff line change
Expand Up @@ -1792,7 +1792,7 @@ theorem setWidth_succ (x : BitVec w) :
· simp_all
· omega

@[deprecated "Use the reverse direction of `cons_msb_setWidth`"]
@[deprecated "Use the reverse direction of `cons_msb_setWidth`" (since := "2024-09-23")]
theorem eq_msb_cons_setWidth (x : BitVec (w+1)) : x = (cons x.msb (x.setWidth w)) := by
simp

Expand Down
4 changes: 2 additions & 2 deletions src/Init/PropLemmas.lean
Original file line number Diff line number Diff line change
Expand Up @@ -643,11 +643,11 @@ theorem decide_ite (u : Prop) [du : Decidable u] (p q : Prop)
(@ite _ p h q (decide p)) = (decide p && q) := by
split <;> simp_all

@[deprecated ite_then_decide_self]
@[deprecated ite_then_decide_self (since := "2024-08-29")]
theorem ite_true_decide_same (p : Prop) [Decidable p] (b : Bool) :
(if p then decide p else b) = (decide p || b) := ite_then_decide_self p b

@[deprecated ite_false_decide_same]
@[deprecated ite_false_decide_same (since := "2024-08-29")]
theorem ite_false_decide_same (p : Prop) [Decidable p] (b : Bool) :
(if p then b else decide p) = (decide p && b) := ite_else_decide_self p b

Expand Down
12 changes: 6 additions & 6 deletions src/Lean/Data/HashMap.lean
Original file line number Diff line number Diff line change
Expand Up @@ -271,11 +271,11 @@ def ofListWith (l : List (α × β)) (f : β → β → β) : HashMap α β :=
| none => m.insert p.fst p.snd
| some v => m.insert p.fst $ f v p.snd)

attribute [deprecated Std.HashMap] HashMap
attribute [deprecated Std.HashMap.Raw] HashMapImp
attribute [deprecated Std.HashMap.Raw.empty] mkHashMapImp
attribute [deprecated Std.HashMap.empty] mkHashMap
attribute [deprecated Std.HashMap.empty] HashMap.empty
attribute [deprecated Std.HashMap.ofList] HashMap.ofList
attribute [deprecated Std.HashMap (since := "2024-08-08")] HashMap
attribute [deprecated Std.HashMap.Raw (since := "2024-08-08")] HashMapImp
attribute [deprecated Std.HashMap.Raw.empty (since := "2024-08-08")] mkHashMapImp
attribute [deprecated Std.HashMap.empty (since := "2024-08-08")] mkHashMap
attribute [deprecated Std.HashMap.empty (since := "2024-08-08")] HashMap.empty
attribute [deprecated Std.HashMap.ofList (since := "2024-08-08")] HashMap.ofList

end Lean.HashMap
10 changes: 5 additions & 5 deletions src/Lean/Data/HashSet.lean
Original file line number Diff line number Diff line change
Expand Up @@ -219,8 +219,8 @@ def merge {α : Type u} [BEq α] [Hashable α] (s t : HashSet α) : HashSet α :
t.fold (init := s) fun s a => s.insert a
-- We don't use `insertMany` here because it gives weird universes.

attribute [deprecated Std.HashSet] HashSet
attribute [deprecated Std.HashSet.Raw] HashSetImp
attribute [deprecated Std.HashSet.Raw.empty] mkHashSetImp
attribute [deprecated Std.HashSet.empty] mkHashSet
attribute [deprecated Std.HashSet.empty] HashSet.empty
attribute [deprecated Std.HashSet (since := "2024-08-08")] HashSet
attribute [deprecated Std.HashSet.Raw (since := "2024-08-08")] HashSetImp
attribute [deprecated Std.HashSet.Raw.empty (since := "2024-08-08")] mkHashSetImp
attribute [deprecated Std.HashSet.empty (since := "2024-08-08")] mkHashSet
attribute [deprecated Std.HashSet.empty (since := "2024-08-08")] HashSet.empty
2 changes: 1 addition & 1 deletion src/Lean/Meta/Tactic/Apply.lean
Original file line number Diff line number Diff line change
Expand Up @@ -254,7 +254,7 @@ Apply `And.intro` as much as possible to goal `mvarId`.
abbrev splitAnd (mvarId : MVarId) : MetaM (List MVarId) :=
splitAndCore mvarId

@[deprecated splitAnd] -- 2024-03-17
@[deprecated splitAnd (since := "2024-03-17")]
def _root_.Lean.Meta.splitAnd (mvarId : MVarId) : MetaM (List MVarId) :=
mvarId.splitAnd

Expand Down
6 changes: 3 additions & 3 deletions src/Std/Data/HashMap/Basic.lean
Original file line number Diff line number Diff line change
Expand Up @@ -118,7 +118,7 @@ Tries to retrieve the mapping for the given key, returning `none` if no such map
@[inline] def get? (m : HashMap α β) (a : α) : Option β :=
DHashMap.Const.get? m.inner a

@[deprecated get? "Use `m[a]?` or `m.get? a` instead", inherit_doc get?]
@[deprecated get? "Use `m[a]?` or `m.get? a` instead" (since := "2024-08-07"), inherit_doc get?]
def find? (m : HashMap α β) (a : α) : Option β :=
m.get? a

Expand All @@ -145,7 +145,7 @@ Retrieves the mapping for the given key. Ensures that such a mapping exists by r
(fallback : β) : β :=
DHashMap.Const.getD m.inner a fallback

@[deprecated getD, inherit_doc getD]
@[deprecated getD (since := "2024-08-07"), inherit_doc getD]
def findD (m : HashMap α β) (a : α) (fallback : β) : β :=
m.getD a fallback

Expand All @@ -157,7 +157,7 @@ Tries to retrieve the mapping for the given key, panicking if no such mapping is
@[inline] def get! [Inhabited β] (m : HashMap α β) (a : α) : β :=
DHashMap.Const.get! m.inner a

@[deprecated get! "Use `m[a]!` or `m.get! a` instead", inherit_doc get!]
@[deprecated get! "Use `m[a]!` or `m.get! a` instead" (since := "2024-08-07"), inherit_doc get!]
def find! [Inhabited β] (m : HashMap α β) (a : α) : Option β :=
m.get! a

Expand Down

0 comments on commit 38c3948

Please sign in to comment.