Documentation

Mathlib.Topology.OpenPartialHomeomorph.Basic

Partial homeomorphisms: basic theory #

Main definitions #

The identity on the whole space as an open partial homeomorphism.

Equations
Instances For
    theorem OpenPartialHomeomorph.isOpen_image_of_subset_source {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] (e : OpenPartialHomeomorph X Y) {s : Set X} (hs : IsOpen s) (hse : s e.source) :
    IsOpen (e '' s)

    An open partial homeomorphism is an open map on its source: the image of an open subset of the source is open.

    The image of the restriction of an open set to the source is open.

    The inverse of an open partial homeomorphism e is an open map on e.target.

    A PartialEquiv with continuous open forward map and open source is a OpenPartialHomeomorph.

    Equations
    Instances For
      @[simp]
      theorem OpenPartialHomeomorph.coe_ofContinuousOpenRestrict {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] (e : PartialEquiv X Y) (hc : ContinuousOn (↑e) e.source) (ho : IsOpenMap (e.source.restrict e)) (hs : IsOpen e.source) :
      (ofContinuousOpenRestrict e hc ho hs) = e
      @[simp]

      A PartialEquiv with continuous open forward map and open source is a OpenPartialHomeomorph.

      Equations
      Instances For
        @[simp]
        theorem OpenPartialHomeomorph.coe_ofContinuousOpen {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] (e : PartialEquiv X Y) (hc : ContinuousOn (↑e) e.source) (ho : IsOpenMap e) (hs : IsOpen e.source) :
        (ofContinuousOpen e hc ho hs) = e
        @[simp]
        theorem OpenPartialHomeomorph.coe_ofContinuousOpen_symm {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] (e : PartialEquiv X Y) (hc : ContinuousOn (↑e) e.source) (ho : IsOpenMap e) (hs : IsOpen e.source) :
        (ofContinuousOpen e hc ho hs).symm = e.symm
        def OpenPartialHomeomorph.homeomorphOfImageSubsetSource {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] (e : OpenPartialHomeomorph X Y) {s : Set X} {t : Set Y} (hs : s e.source) (ht : e '' s = t) :
        s ≃ₜ t

        The homeomorphism obtained by restricting an OpenPartialHomeomorph to a subset of the source.

        Equations
        • One or more equations did not get rendered due to their size.
        Instances For
          @[simp]
          theorem OpenPartialHomeomorph.homeomorphOfImageSubsetSource_symm_apply {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] (e : OpenPartialHomeomorph X Y) {s : Set X} {t : Set Y} (hs : s e.source) (ht : e '' s = t) (a✝ : t) :
          (e.homeomorphOfImageSubsetSource hs ht).symm a✝ = Set.MapsTo.restrict (↑e.symm) t s a✝
          @[simp]
          theorem OpenPartialHomeomorph.homeomorphOfImageSubsetSource_apply {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] (e : OpenPartialHomeomorph X Y) {s : Set X} {t : Set Y} (hs : s e.source) (ht : e '' s = t) (a✝ : s) :
          (e.homeomorphOfImageSubsetSource hs ht) a✝ = Set.MapsTo.restrict (↑e) s t a✝

          An open partial homeomorphism defines a homeomorphism between its source and target.

          Equations
          Instances For

            If an open partial homeomorphism has source and target equal to univ, then it induces a homeomorphism between the whole spaces, expressed in this definition.

            Equations
            Instances For

              An open partial homeomorphism whose source is all of X defines an open embedding of X into Y. The converse is also true; see IsOpenEmbedding.toOpenPartialHomeomorph.

              Open embeddings #

              An open embedding of X into Y, with X nonempty, defines an open partial homeomorphism whose source is all of X. The converse is also true; see OpenPartialHomeomorph.to_isOpenEmbedding.

              Equations
              Instances For
                @[deprecated Topology.IsOpenEmbedding.toOpenPartialHomeomorph (since := "2025-08-29")]

                Alias of Topology.IsOpenEmbedding.toOpenPartialHomeomorph.


                An open embedding of X into Y, with X nonempty, defines an open partial homeomorphism whose source is all of X. The converse is also true; see OpenPartialHomeomorph.to_isOpenEmbedding.

                Equations
                Instances For
                  @[deprecated Topology.IsOpenEmbedding.toOpenPartialHomeomorph_left_inv (since := "2025-08-29")]
                  theorem Topology.IsOpenEmbedding.toPartialHomeomorph_left_inv {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] (f : XY) (h : IsOpenEmbedding f) [Nonempty X] {x : X} :
                  (toOpenPartialHomeomorph f h).symm (f x) = x

                  Alias of Topology.IsOpenEmbedding.toOpenPartialHomeomorph_left_inv.

                  @[deprecated Topology.IsOpenEmbedding.toOpenPartialHomeomorph_right_inv (since := "2025-08-29")]
                  theorem Topology.IsOpenEmbedding.toPartialHomeomorph_right_inv {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] (f : XY) (h : IsOpenEmbedding f) [Nonempty X] {x : Y} (hx : x Set.range f) :
                  f ((toOpenPartialHomeomorph f h).symm x) = x

                  Alias of Topology.IsOpenEmbedding.toOpenPartialHomeomorph_right_inv.

                  inclusion of an open set in a topological space

                  The inclusion of an open subset s of a space X into X is an open partial homeomorphism from the subtype s to X.

                  Equations
                  Instances For
                    @[deprecated TopologicalSpace.Opens.openPartialHomeomorphSubtypeCoe (since := "2025-08-29")]

                    Alias of TopologicalSpace.Opens.openPartialHomeomorphSubtypeCoe.


                    The inclusion of an open subset s of a space X into X is an open partial homeomorphism from the subtype s to X.

                    Equations
                    Instances For
                      @[deprecated TopologicalSpace.Opens.openPartialHomeomorphSubtypeCoe_coe (since := "2025-08-29")]

                      Alias of TopologicalSpace.Opens.openPartialHomeomorphSubtypeCoe_coe.

                      @[deprecated TopologicalSpace.Opens.openPartialHomeomorphSubtypeCoe_source (since := "2025-08-29")]

                      Alias of TopologicalSpace.Opens.openPartialHomeomorphSubtypeCoe_source.

                      @[deprecated TopologicalSpace.Opens.openPartialHomeomorphSubtypeCoe_target (since := "2025-08-29")]

                      Alias of TopologicalSpace.Opens.openPartialHomeomorphSubtypeCoe_target.