Skip to content

Commit

Permalink
feat(topology/constructions): add subtype.dense_iff (#14632)
Browse files Browse the repository at this point in the history
Also add `inducing.dense_iff`.
  • Loading branch information
urkud committed Jun 9, 2022
1 parent 48f557d commit e0f3ea3
Show file tree
Hide file tree
Showing 2 changed files with 7 additions and 0 deletions.
3 changes: 3 additions & 0 deletions src/topology/constructions.lean
Expand Up @@ -764,6 +764,9 @@ lemma continuous_at_subtype_coe {p : α → Prop} {a : subtype p} :
continuous_at (coe : subtype p → α) a :=
continuous_iff_continuous_at.mp continuous_subtype_coe _

lemma subtype.dense_iff {s : set α} {t : set s} : dense t ↔ s ⊆ closure (coe '' t) :=
by { rw [inducing_coe.dense_iff, set_coe.forall], refl }

lemma map_nhds_subtype_coe_eq {a : α} (ha : p a) (h : {a | p a} ∈ 𝓝 a) :
map (coe : subtype p → α) (𝓝 ⟨a, ha⟩) = 𝓝 a :=
map_nhds_induced_of_mem $ by simpa only [subtype.coe_mk, subtype.range_coe] using h
Expand Down
4 changes: 4 additions & 0 deletions src/topology/maps.lean
Expand Up @@ -129,6 +129,10 @@ lemma inducing.is_open_iff {f : α → β} (hf : inducing f) {s : set α} :
is_open s ↔ ∃ t, is_open t ∧ f ⁻¹' t = s :=
by rw [hf.induced, is_open_induced_iff]

lemma inducing.dense_iff {f : α → β} (hf : inducing f) {s : set α} :
dense s ↔ ∀ x, f x ∈ closure (f '' s) :=
by simp only [dense, hf.closure_eq_preimage_closure_image, mem_preimage]

end inducing

section embedding
Expand Down

0 comments on commit e0f3ea3

Please sign in to comment.