Skip to content

Commit

Permalink
chore: adapt stdlib to new variable behavior
Browse files Browse the repository at this point in the history
  • Loading branch information
Kha authored and kim-em committed Aug 12, 2024
1 parent a44cb65 commit 0edf1ba
Show file tree
Hide file tree
Showing 7 changed files with 14 additions and 6 deletions.
8 changes: 4 additions & 4 deletions src/Init/Data/List/Erase.lean
Original file line number Diff line number Diff line change
Expand Up @@ -354,7 +354,7 @@ theorem erase_eq_iff [LawfulBEq α] {a : α} {l : List α} :
rw [erase_of_not_mem]
simp_all

theorem Nodup.erase_eq_filter [BEq α] [LawfulBEq α] {l} (d : Nodup l) (a : α) : l.erase a = l.filter (· != a) := by
theorem Nodup.erase_eq_filter [LawfulBEq α] {l} (d : Nodup l) (a : α) : l.erase a = l.filter (· != a) := by
induction d with
| nil => rfl
| cons m _n ih =>
Expand All @@ -367,13 +367,13 @@ theorem Nodup.erase_eq_filter [BEq α] [LawfulBEq α] {l} (d : Nodup l) (a : α)
simpa [@eq_comm α] using m
· simp [beq_false_of_ne h, ih, h]

theorem Nodup.mem_erase_iff [BEq α] [LawfulBEq α] {a : α} (d : Nodup l) : a ∈ l.erase b ↔ a ≠ b ∧ a ∈ l := by
theorem Nodup.mem_erase_iff [LawfulBEq α] {a : α} (d : Nodup l) : a ∈ l.erase b ↔ a ≠ b ∧ a ∈ l := by
rw [Nodup.erase_eq_filter d, mem_filter, and_comm, bne_iff_ne]

theorem Nodup.not_mem_erase [BEq α] [LawfulBEq α] {a : α} (h : Nodup l) : a ∉ l.erase a := fun H => by
theorem Nodup.not_mem_erase [LawfulBEq α] {a : α} (h : Nodup l) : a ∉ l.erase a := fun H => by
simpa using ((Nodup.mem_erase_iff h).mp H).left

theorem Nodup.erase [BEq α] [LawfulBEq α] (a : α) : Nodup l → Nodup (l.erase a) :=
theorem Nodup.erase [LawfulBEq α] (a : α) : Nodup l → Nodup (l.erase a) :=
Nodup.sublist <| erase_sublist _ _

end erase
Expand Down
1 change: 1 addition & 0 deletions src/Init/WF.lean
Original file line number Diff line number Diff line change
Expand Up @@ -68,6 +68,7 @@ noncomputable def recursion {C : α → Sort v} (a : α) (h : ∀ x, (∀ y, r y
induction (apply hwf a) with
| intro x₁ _ ih => exact h x₁ ih

include hwf in
theorem induction {C : α → Prop} (a : α) (h : ∀ x, (∀ y, r y x → C y) → C x) : C a :=
recursion hwf a h

Expand Down
6 changes: 5 additions & 1 deletion src/Std/Data/DHashMap/Internal/RawLemmas.lean
Original file line number Diff line number Diff line change
Expand Up @@ -20,7 +20,7 @@ open Std.DHashMap.Internal.List

universe u v

variable {α : Type u} {β : α → Type v} [BEq α] [Hashable α]
variable {α : Type u} {β : α → Type v}

namespace Std.DHashMap.Internal

Expand All @@ -41,6 +41,8 @@ theorem Raw.buckets_emptyc {i : Nat} {h} :
(∅ : Raw α β).buckets[i]'h = AssocList.nil :=
buckets_empty

variable [BEq α] [Hashable α]

@[simp]
theorem buckets_empty {c} {i : Nat} {h} :
(empty c : DHashMap α β).1.buckets[i]'h = AssocList.nil := by
Expand All @@ -55,7 +57,9 @@ end empty

namespace Raw₀

variable [BEq α] [Hashable α]
variable (m : Raw₀ α β) (h : m.1.WF)
set_option deprecated.oldSectionVars true

/-- Internal implementation detail of the hash map -/
scoped macro "wf_trivial" : tactic => `(tactic|
Expand Down
1 change: 1 addition & 0 deletions src/Std/Data/DHashMap/RawLemmas.lean
Original file line number Diff line number Diff line change
Expand Up @@ -75,6 +75,7 @@ namespace Raw
open Internal.Raw₀ Internal.Raw

variable {m : Raw α β} (h : m.WF)
set_option deprecated.oldSectionVars true

@[simp]
theorem isEmpty_empty {c} : (empty c : Raw α β).isEmpty := by
Expand Down
1 change: 1 addition & 0 deletions src/Std/Data/HashMap/RawLemmas.lean
Original file line number Diff line number Diff line change
Expand Up @@ -27,6 +27,7 @@ namespace Std.HashMap
namespace Raw

variable {m : Raw α β} (h : m.WF)
set_option deprecated.oldSectionVars true

private theorem ext {m m' : Raw α β} : m.inner = m'.inner → m = m' := by
cases m; cases m'; rintro rfl; rfl
Expand Down
2 changes: 1 addition & 1 deletion src/Std/Data/HashSet/Lemmas.lean
Original file line number Diff line number Diff line change
Expand Up @@ -20,7 +20,7 @@ set_option autoImplicit false

universe u v

variable {α : Type u} {_ : BEq α} {_ : Hashable α} [EquivBEq α] [LawfulHashable α]
variable {α : Type u} {_ : BEq α} {_ : Hashable α}

namespace Std.HashSet

Expand Down
1 change: 1 addition & 0 deletions src/Std/Data/HashSet/RawLemmas.lean
Original file line number Diff line number Diff line change
Expand Up @@ -27,6 +27,7 @@ namespace Std.HashSet
namespace Raw

variable {m : Raw α} (h : m.WF)
set_option deprecated.oldSectionVars true

private theorem ext {m m' : Raw α} : m.inner = m'.inner → m = m' := by
cases m; cases m'; rintro rfl; rfl
Expand Down

0 comments on commit 0edf1ba

Please sign in to comment.