Skip to content

Commit

Permalink
feat: three tiny homeomorph lemmas (#9166)
Browse files Browse the repository at this point in the history
  • Loading branch information
PatrickMassot committed Jan 15, 2024
1 parent d8dff07 commit b1ff86d
Show file tree
Hide file tree
Showing 2 changed files with 12 additions and 0 deletions.
4 changes: 4 additions & 0 deletions Mathlib/Topology/Algebra/Group/Basic.lean
Expand Up @@ -334,6 +334,10 @@ protected def Homeomorph.inv (G : Type*) [TopologicalSpace G] [InvolutiveInv G]
#align homeomorph.inv Homeomorph.inv
#align homeomorph.neg Homeomorph.neg

@[to_additive (attr := simp)]
lemma Homeomorph.coe_inv {G : Type*} [TopologicalSpace G] [InvolutiveInv G] [ContinuousInv G] :
⇑(Homeomorph.inv G) = Inv.inv := rfl

@[to_additive]
theorem isOpenMap_inv : IsOpenMap (Inv.inv : G → G) :=
(Homeomorph.inv _).isOpenMap
Expand Down
8 changes: 8 additions & 0 deletions Mathlib/Topology/Homeomorph.lean
Expand Up @@ -234,6 +234,9 @@ theorem preimage_image (h : X ≃ₜ Y) (s : Set X) : h ⁻¹' (h '' s) = s :=
h.toEquiv.preimage_image s
#align homeomorph.preimage_image Homeomorph.preimage_image

lemma image_compl (h : X ≃ₜ Y) (s : Set X) : h '' (sᶜ) = (h '' s)ᶜ :=
h.toEquiv.image_compl s

protected theorem inducing (h : X ≃ₜ Y) : Inducing h :=
inducing_of_inducing_compose h.continuous h.symm.continuous <| by
simp only [symm_comp_self, inducing_id]
Expand Down Expand Up @@ -427,6 +430,11 @@ theorem map_nhds_eq (h : X ≃ₜ Y) (x : X) : map h (𝓝 x) = 𝓝 (h x) :=
h.embedding.map_nhds_of_mem _ (by simp)
#align homeomorph.map_nhds_eq Homeomorph.map_nhds_eq

@[simp]
theorem map_punctured_nhds_eq (h : X ≃ₜ Y) (x : X) : map h (𝓝[≠] x) = 𝓝[≠] (h x) := by
convert h.embedding.map_nhdsWithin_eq ({x}ᶜ) x
rw [h.image_compl, Set.image_singleton]

theorem symm_map_nhds_eq (h : X ≃ₜ Y) (x : X) : map h.symm (𝓝 (h x)) = 𝓝 x := by
rw [h.symm.map_nhds_eq, h.symm_apply_apply]
#align homeomorph.symm_map_nhds_eq Homeomorph.symm_map_nhds_eq
Expand Down

0 comments on commit b1ff86d

Please sign in to comment.