Skip to content

Commit

Permalink
feat(topology/continuous_on): generalise lemma slightly (#16840)
Browse files Browse the repository at this point in the history
Generalise `frequently_nhds_within_iff` to allow `within` any set, not just complement of singleton



Co-authored-by: Bhavik Mehta <bm489@cam.ac.uk>
  • Loading branch information
b-mehta and b-mehta committed Oct 8, 2022
1 parent 827f0b9 commit 90fab28
Showing 1 changed file with 3 additions and 3 deletions.
6 changes: 3 additions & 3 deletions src/topology/continuous_on.lean
Expand Up @@ -44,9 +44,9 @@ lemma eventually_nhds_within_iff {a : α} {s : set α} {p : α → Prop} :
(∀ᶠ x in 𝓝[s] a, p x) ↔ ∀ᶠ x in 𝓝 a, x ∈ s → p x :=
eventually_inf_principal

lemma frequently_nhds_within_iff {z : α} {p : α → Prop} :
(∃ᶠ x in 𝓝[] z, p x) ↔ (∃ᶠ x in 𝓝 z, p x ∧ x ≠ z) :=
iff.not (by simp [eventually_nhds_within_iff, not_imp_not])
lemma frequently_nhds_within_iff {z : α} {s : set α} {p : α → Prop} :
(∃ᶠ x in 𝓝[s] z, p x) ↔ (∃ᶠ x in 𝓝 z, p x ∧ x ∈ s) :=
iff.not (by simp [eventually_nhds_within_iff, not_and'])

lemma mem_closure_ne_iff_frequently_within {z : α} {s : set α} :
z ∈ closure (s \ {z}) ↔ ∃ᶠ x in 𝓝[≠] z, x ∈ s :=
Expand Down

0 comments on commit 90fab28

Please sign in to comment.