Documentation

Mathlib.Topology.PartialHomeomorph

Partial homeomorphisms #

This file defines homeomorphisms between open subsets of topological spaces. An element e of PartialHomeomorph X Y is an extension of PartialEquiv X Y, i.e., it is a pair of functions e.toFun and e.invFun, inverse of each other on the sets e.source and e.target. Additionally, we require that these sets are open, and that the functions are continuous on them. Equivalently, they are homeomorphisms there.

As in equivs, we register a coercion to functions, and we use e x and e.symm x throughout instead of e.toFun x and e.invFun x.

Main definitions #

Implementation notes #

Most statements are copied from their PartialEquiv versions, although some care is required especially when restricting to subsets, as these should be open subsets.

For design notes, see PartialEquiv.lean.

Local coding conventions #

If a lemma deals with the intersection of a set with either source or target of a PartialEquiv, then it should use e.source ∩ s or e.target ∩ t, not s ∩ e.source or t ∩ e.target.

structure PartialHomeomorph (X : Type u_7) (Y : Type u_8) [TopologicalSpace X] [TopologicalSpace Y] extends PartialEquiv :
Type (max u_7 u_8)

Partial homeomorphisms, defined on open subsets of the space

  • toFun : XY
  • invFun : YX
  • source : Set X
  • target : Set Y
  • map_source' : ∀ ⦃x : X⦄, x self.sourceself.toPartialEquiv x self.target
  • map_target' : ∀ ⦃x : Y⦄, x self.targetself.invFun x self.source
  • left_inv' : ∀ ⦃x : X⦄, x self.sourceself.invFun (self.toPartialEquiv x) = x
  • right_inv' : ∀ ⦃x : Y⦄, x self.targetself.toPartialEquiv (self.invFun x) = x
  • open_source : IsOpen self.source
  • open_target : IsOpen self.target
  • continuousOn_toFun : ContinuousOn (self.toPartialEquiv) self.source
  • continuousOn_invFun : ContinuousOn self.invFun self.target
Instances For

    Basic properties; inverse (symm instance)

    def PartialHomeomorph.toFun' {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] (e : PartialHomeomorph X Y) :
    XY

    Coercion of a partial homeomorphisms to a function. We don't use e.toFun because it is actually e.toPartialEquiv.toFun, so simp will apply lemmas about toPartialEquiv. While we may want to switch to this behavior later, doing it mid-port will break a lot of proofs.

    Equations
    • e = e.toPartialEquiv
    Instances For

      Coercion of a PartialHomeomorph to function. Note that a PartialHomeomorph is not DFunLike.

      Equations
      • PartialHomeomorph.instCoeFunPartialHomeomorphForAll = { coe := fun (e : PartialHomeomorph X Y) => e }

      The inverse of a partial homeomorphism

      Equations
      Instances For

        See Note [custom simps projection]. We need to specify this projection explicitly in this case, because it is a composition of multiple projections.

        Equations
        Instances For

          See Note [custom simps projection]

          Equations
          Instances For
            @[simp]
            theorem PartialHomeomorph.mk_coe {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] (e : PartialEquiv X Y) (a : IsOpen e.source) (b : IsOpen e.target) (c : ContinuousOn (e) e.source) (d : ContinuousOn e.invFun e.target) :
            { toPartialEquiv := e, open_source := a, open_target := b, continuousOn_toFun := c, continuousOn_invFun := d } = e
            @[simp]
            theorem PartialHomeomorph.mk_coe_symm {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] (e : PartialEquiv X Y) (a : IsOpen e.source) (b : IsOpen e.target) (c : ContinuousOn (e) e.source) (d : ContinuousOn e.invFun e.target) :
            (PartialHomeomorph.symm { toPartialEquiv := e, open_source := a, open_target := b, continuousOn_toFun := c, continuousOn_invFun := d }) = (PartialEquiv.symm e)
            theorem PartialHomeomorph.toPartialEquiv_injective {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] :
            Function.Injective PartialHomeomorph.toPartialEquiv
            @[simp]
            theorem PartialHomeomorph.toFun_eq_coe {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] (e : PartialHomeomorph X Y) :
            e.toPartialEquiv = e
            @[simp]
            @[simp]
            theorem PartialHomeomorph.coe_coe {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] (e : PartialHomeomorph X Y) :
            e.toPartialEquiv = e
            @[simp]
            @[simp]
            theorem PartialHomeomorph.map_source {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] (e : PartialHomeomorph X Y) {x : X} (h : x e.source) :
            e x e.target
            theorem PartialHomeomorph.map_source'' {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] (e : PartialHomeomorph X Y) :
            e '' e.source e.target

            Variant of map_source, stated for images of subsets of source.

            @[simp]
            theorem PartialHomeomorph.map_target {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] (e : PartialHomeomorph X Y) {x : Y} (h : x e.target) :
            (PartialHomeomorph.symm e) x e.source
            @[simp]
            theorem PartialHomeomorph.left_inv {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] (e : PartialHomeomorph X Y) {x : X} (h : x e.source) :
            (PartialHomeomorph.symm e) (e x) = x
            @[simp]
            theorem PartialHomeomorph.right_inv {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] (e : PartialHomeomorph X Y) {x : Y} (h : x e.target) :
            e ((PartialHomeomorph.symm e) x) = x
            theorem PartialHomeomorph.eq_symm_apply {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] (e : PartialHomeomorph X Y) {x : X} {y : Y} (hx : x e.source) (hy : y e.target) :
            x = (PartialHomeomorph.symm e) y e x = y
            theorem PartialHomeomorph.mapsTo {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] (e : PartialHomeomorph X Y) :
            Set.MapsTo (e) e.source e.target
            theorem PartialHomeomorph.invOn {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] (e : PartialHomeomorph X Y) :
            Set.InvOn ((PartialHomeomorph.symm e)) (e) e.source e.target
            theorem PartialHomeomorph.injOn {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] (e : PartialHomeomorph X Y) :
            Set.InjOn (e) e.source
            theorem PartialHomeomorph.bijOn {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] (e : PartialHomeomorph X Y) :
            Set.BijOn (e) e.source e.target
            theorem PartialHomeomorph.surjOn {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] (e : PartialHomeomorph X Y) :
            Set.SurjOn (e) e.source e.target
            @[simp]
            theorem Homeomorph.toPartialHomeomorphOfImageEq_apply {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] (e : X ≃ₜ Y) (s : Set X) (hs : IsOpen s) (t : Set Y) (h : e '' s = t) :
            @[simp]
            theorem Homeomorph.toPartialHomeomorphOfImageEq_symm_apply {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] (e : X ≃ₜ Y) (s : Set X) (hs : IsOpen s) (t : Set Y) (h : e '' s = t) :
            @[simp]
            theorem Homeomorph.toPartialHomeomorphOfImageEq_toPartialEquiv {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] (e : X ≃ₜ Y) (s : Set X) (hs : IsOpen s) (t : Set Y) (h : e '' s = t) :
            (Homeomorph.toPartialHomeomorphOfImageEq e s hs t h).toPartialEquiv = Equiv.toPartialEquivOfImageEq e.toEquiv s t h
            theorem Homeomorph.toPartialHomeomorphOfImageEq_source {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] (e : X ≃ₜ Y) (s : Set X) (hs : IsOpen s) (t : Set Y) (h : e '' s = t) :
            theorem Homeomorph.toPartialHomeomorphOfImageEq_target {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] (e : X ≃ₜ Y) (s : Set X) (hs : IsOpen s) (t : Set Y) (h : e '' s = t) :
            def Homeomorph.toPartialHomeomorphOfImageEq {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] (e : X ≃ₜ Y) (s : Set X) (hs : IsOpen s) (t : Set Y) (h : e '' s = t) :

            Interpret a Homeomorph as a PartialHomeomorph by restricting it to an open set s in the domain and to t in the codomain.

            Equations
            • One or more equations did not get rendered due to their size.
            Instances For
              @[simp]
              @[simp]

              A homeomorphism induces a partial homeomorphism on the whole space

              Equations
              Instances For
                def PartialHomeomorph.replaceEquiv {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] (e : PartialHomeomorph X Y) (e' : PartialEquiv X Y) (h : e.toPartialEquiv = e') :

                Replace toPartialEquiv field to provide better definitional equalities.

                Equations
                • PartialHomeomorph.replaceEquiv e e' h = { toPartialEquiv := e', open_source := , open_target := , continuousOn_toFun := , continuousOn_invFun := }
                Instances For
                  theorem PartialHomeomorph.replaceEquiv_eq_self {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] (e : PartialHomeomorph X Y) (e' : PartialEquiv X Y) (h : e.toPartialEquiv = e') :
                  theorem PartialHomeomorph.source_preimage_target {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] (e : PartialHomeomorph X Y) :
                  e.source e ⁻¹' e.target
                  @[deprecated PartialHomeomorph.toPartialEquiv_injective]
                  theorem PartialHomeomorph.eq_of_partialEquiv_eq {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] {e : PartialHomeomorph X Y} {e' : PartialHomeomorph X Y} (h : e.toPartialEquiv = e'.toPartialEquiv) :
                  e = e'
                  theorem PartialHomeomorph.eventually_left_inverse {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] (e : PartialHomeomorph X Y) {x : X} (hx : x e.source) :
                  ∀ᶠ (y : X) in nhds x, (PartialHomeomorph.symm e) (e y) = y
                  theorem PartialHomeomorph.eventually_left_inverse' {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] (e : PartialHomeomorph X Y) {x : Y} (hx : x e.target) :
                  ∀ᶠ (y : X) in nhds ((PartialHomeomorph.symm e) x), (PartialHomeomorph.symm e) (e y) = y
                  theorem PartialHomeomorph.eventually_right_inverse {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] (e : PartialHomeomorph X Y) {x : Y} (hx : x e.target) :
                  ∀ᶠ (y : Y) in nhds x, e ((PartialHomeomorph.symm e) y) = y
                  theorem PartialHomeomorph.eventually_right_inverse' {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] (e : PartialHomeomorph X Y) {x : X} (hx : x e.source) :
                  ∀ᶠ (y : Y) in nhds (e x), e ((PartialHomeomorph.symm e) y) = y
                  theorem PartialHomeomorph.eventually_ne_nhdsWithin {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] (e : PartialHomeomorph X Y) {x : X} (hx : x e.source) :
                  ∀ᶠ (x' : X) in nhdsWithin x {x}, e x' e x
                  theorem PartialHomeomorph.nhdsWithin_source_inter {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] (e : PartialHomeomorph X Y) {x : X} (hx : x e.source) (s : Set X) :
                  nhdsWithin x (e.source s) = nhdsWithin x s
                  theorem PartialHomeomorph.nhdsWithin_target_inter {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] (e : PartialHomeomorph X Y) {x : Y} (hx : x e.target) (s : Set Y) :
                  nhdsWithin x (e.target s) = nhdsWithin x s
                  theorem PartialHomeomorph.image_eq_target_inter_inv_preimage {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] (e : PartialHomeomorph X Y) {s : Set X} (h : s e.source) :
                  e '' s = e.target (PartialHomeomorph.symm e) ⁻¹' s
                  theorem PartialHomeomorph.image_source_inter_eq' {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] (e : PartialHomeomorph X Y) (s : Set X) :
                  e '' (e.source s) = e.target (PartialHomeomorph.symm e) ⁻¹' s
                  theorem PartialHomeomorph.image_source_inter_eq {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] (e : PartialHomeomorph X Y) (s : Set X) :
                  e '' (e.source s) = e.target (PartialHomeomorph.symm e) ⁻¹' (e.source s)
                  theorem PartialHomeomorph.symm_image_eq_source_inter_preimage {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] (e : PartialHomeomorph X Y) {s : Set Y} (h : s e.target) :
                  (PartialHomeomorph.symm e) '' s = e.source e ⁻¹' s
                  theorem PartialHomeomorph.symm_image_target_inter_eq {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] (e : PartialHomeomorph X Y) (s : Set Y) :
                  (PartialHomeomorph.symm e) '' (e.target s) = e.source e ⁻¹' (e.target s)
                  theorem PartialHomeomorph.source_inter_preimage_target_inter {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] (e : PartialHomeomorph X Y) (s : Set Y) :
                  e.source e ⁻¹' (e.target s) = e.source e ⁻¹' s
                  theorem PartialHomeomorph.image_source_eq_target {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] (e : PartialHomeomorph X Y) :
                  e '' e.source = e.target
                  theorem PartialHomeomorph.ext {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] (e : PartialHomeomorph X Y) (e' : PartialHomeomorph X Y) (h : ∀ (x : X), e x = e' x) (hinv : ∀ (x : Y), (PartialHomeomorph.symm e) x = (PartialHomeomorph.symm e') x) (hs : e.source = e'.source) :
                  e = e'

                  Two partial homeomorphisms are equal when they have equal toFun, invFun and source. It is not sufficient to have equal toFun and source, as this only determines invFun on the target. This would only be true for a weaker notion of equality, arguably the right one, called EqOnSource.

                  theorem PartialHomeomorph.ext_iff {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] {e : PartialHomeomorph X Y} {e' : PartialHomeomorph X Y} :
                  e = e' (∀ (x : X), e x = e' x) (∀ (x : Y), (PartialHomeomorph.symm e) x = (PartialHomeomorph.symm e') x) e.source = e'.source
                  @[simp]
                  theorem PartialHomeomorph.symm_toPartialEquiv {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] (e : PartialHomeomorph X Y) :
                  (PartialHomeomorph.symm e).toPartialEquiv = PartialEquiv.symm e.toPartialEquiv
                  theorem PartialHomeomorph.symm_bijective {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] :
                  Function.Bijective PartialHomeomorph.symm
                  theorem PartialHomeomorph.continuousAt {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] (e : PartialHomeomorph X Y) {x : X} (h : x e.source) :
                  ContinuousAt (e) x

                  A partial homeomorphism is continuous at any point of its source

                  theorem PartialHomeomorph.continuousAt_symm {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] (e : PartialHomeomorph X Y) {x : Y} (h : x e.target) :

                  A partial homeomorphism inverse is continuous at any point of its target

                  theorem PartialHomeomorph.tendsto_symm {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] (e : PartialHomeomorph X Y) {x : X} (hx : x e.source) :
                  theorem PartialHomeomorph.map_nhds_eq {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] (e : PartialHomeomorph X Y) {x : X} (hx : x e.source) :
                  Filter.map (e) (nhds x) = nhds (e x)
                  theorem PartialHomeomorph.symm_map_nhds_eq {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] (e : PartialHomeomorph X Y) {x : X} (hx : x e.source) :
                  theorem PartialHomeomorph.image_mem_nhds {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] (e : PartialHomeomorph X Y) {x : X} (hx : x e.source) {s : Set X} (hs : s nhds x) :
                  e '' s nhds (e x)
                  theorem PartialHomeomorph.map_nhdsWithin_eq {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] (e : PartialHomeomorph X Y) {x : X} (hx : x e.source) (s : Set X) :
                  Filter.map (e) (nhdsWithin x s) = nhdsWithin (e x) (e '' (e.source s))
                  theorem PartialHomeomorph.map_nhdsWithin_preimage_eq {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] (e : PartialHomeomorph X Y) {x : X} (hx : x e.source) (s : Set Y) :
                  Filter.map (e) (nhdsWithin x (e ⁻¹' s)) = nhdsWithin (e x) s
                  theorem PartialHomeomorph.eventually_nhds {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] (e : PartialHomeomorph X Y) {x : X} (p : YProp) (hx : x e.source) :
                  (∀ᶠ (y : Y) in nhds (e x), p y) ∀ᶠ (x : X) in nhds x, p (e x)
                  theorem PartialHomeomorph.eventually_nhds' {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] (e : PartialHomeomorph X Y) {x : X} (p : XProp) (hx : x e.source) :
                  (∀ᶠ (y : Y) in nhds (e x), p ((PartialHomeomorph.symm e) y)) ∀ᶠ (x : X) in nhds x, p x
                  theorem PartialHomeomorph.eventually_nhdsWithin {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] (e : PartialHomeomorph X Y) {x : X} (p : YProp) {s : Set X} (hx : x e.source) :
                  (∀ᶠ (y : Y) in nhdsWithin (e x) ((PartialHomeomorph.symm e) ⁻¹' s), p y) ∀ᶠ (x : X) in nhdsWithin x s, p (e x)
                  theorem PartialHomeomorph.eventually_nhdsWithin' {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] (e : PartialHomeomorph X Y) {x : X} (p : XProp) {s : Set X} (hx : x e.source) :
                  (∀ᶠ (y : Y) in nhdsWithin (e x) ((PartialHomeomorph.symm e) ⁻¹' s), p ((PartialHomeomorph.symm e) y)) ∀ᶠ (x : X) in nhdsWithin x s, p x
                  theorem PartialHomeomorph.preimage_eventuallyEq_target_inter_preimage_inter {X : Type u_1} {Y : Type u_3} {Z : Type u_5} [TopologicalSpace X] [TopologicalSpace Y] [TopologicalSpace Z] {e : PartialHomeomorph X Y} {s : Set X} {t : Set Z} {x : X} {f : XZ} (hf : ContinuousWithinAt f s x) (hxe : x e.source) (ht : t nhds (f x)) :

                  This lemma is useful in the manifold library in the case that e is a chart. It states that locally around e x the set e.symm ⁻¹' s is the same as the set intersected with the target of e and some other neighborhood of f x (which will be the source of a chart on Z).

                  theorem PartialHomeomorph.isOpen_inter_preimage {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] (e : PartialHomeomorph X Y) {s : Set Y} (hs : IsOpen s) :
                  IsOpen (e.source e ⁻¹' s)
                  theorem PartialHomeomorph.isOpen_image_of_subset_source {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] (e : PartialHomeomorph X Y) {s : Set X} (hs : IsOpen s) (hse : s e.source) :
                  IsOpen (e '' s)

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

                  theorem PartialHomeomorph.isOpen_image_source_inter {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] (e : PartialHomeomorph X Y) {s : Set X} (hs : IsOpen s) :
                  IsOpen (e '' (e.source s))

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

                  theorem PartialHomeomorph.isOpen_image_symm_of_subset_target {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] (e : PartialHomeomorph X Y) {t : Set Y} (ht : IsOpen t) (hte : t e.target) :

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

                  theorem PartialHomeomorph.isOpen_image_iff_of_subset_source {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] (e : PartialHomeomorph X Y) {s : Set X} (hs : s e.source) :
                  IsOpen (e '' s) IsOpen s

                  PartialHomeomorph.IsImage relation #

                  We say that t : Set Y is an image of s : Set X under a partial homeomorphism e if any of the following equivalent conditions hold:

                  This definition is a restatement of PartialEquiv.IsImage for partial homeomorphisms. In this section we transfer API about PartialEquiv.IsImage to partial homeomorphisms and add a few PartialHomeomorph-specific lemmas like PartialHomeomorph.IsImage.closure.

                  def PartialHomeomorph.IsImage {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] (e : PartialHomeomorph X Y) (s : Set X) (t : Set Y) :

                  We say that t : Set Y is an image of s : Set X under a partial homeomorphism e if any of the following equivalent conditions hold:

                  • e '' (e.source ∩ s) = e.target ∩ t;
                  • e.source ∩ e ⁻¹ t = e.source ∩ s;
                  • ∀ x ∈ e.source, e x ∈ t ↔ x ∈ s (this one is used in the definition).
                  Equations
                  Instances For
                    theorem PartialHomeomorph.IsImage.toPartialEquiv {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] {e : PartialHomeomorph X Y} {s : Set X} {t : Set Y} (h : PartialHomeomorph.IsImage e s t) :
                    PartialEquiv.IsImage e.toPartialEquiv s t
                    theorem PartialHomeomorph.IsImage.apply_mem_iff {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] {e : PartialHomeomorph X Y} {s : Set X} {t : Set Y} {x : X} (h : PartialHomeomorph.IsImage e s t) (hx : x e.source) :
                    e x t x s
                    theorem PartialHomeomorph.IsImage.symm_apply_mem_iff {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] {e : PartialHomeomorph X Y} {s : Set X} {t : Set Y} {y : Y} (h : PartialHomeomorph.IsImage e s t) (hy : y e.target) :
                    theorem PartialHomeomorph.IsImage.mapsTo {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] {e : PartialHomeomorph X Y} {s : Set X} {t : Set Y} (h : PartialHomeomorph.IsImage e s t) :
                    Set.MapsTo (e) (e.source s) (e.target t)
                    theorem PartialHomeomorph.IsImage.symm_mapsTo {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] {e : PartialHomeomorph X Y} {s : Set X} {t : Set Y} (h : PartialHomeomorph.IsImage e s t) :
                    Set.MapsTo ((PartialHomeomorph.symm e)) (e.target t) (e.source s)
                    theorem PartialHomeomorph.IsImage.image_eq {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] {e : PartialHomeomorph X Y} {s : Set X} {t : Set Y} (h : PartialHomeomorph.IsImage e s t) :
                    e '' (e.source s) = e.target t
                    theorem PartialHomeomorph.IsImage.symm_image_eq {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] {e : PartialHomeomorph X Y} {s : Set X} {t : Set Y} (h : PartialHomeomorph.IsImage e s t) :
                    (PartialHomeomorph.symm e) '' (e.target t) = e.source s
                    theorem PartialHomeomorph.IsImage.iff_preimage_eq {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] {e : PartialHomeomorph X Y} {s : Set X} {t : Set Y} :
                    PartialHomeomorph.IsImage e s t e.source e ⁻¹' t = e.source s
                    theorem PartialHomeomorph.IsImage.of_preimage_eq {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] {e : PartialHomeomorph X Y} {s : Set X} {t : Set Y} :
                    e.source e ⁻¹' t = e.source sPartialHomeomorph.IsImage e s t

                    Alias of the reverse direction of PartialHomeomorph.IsImage.iff_preimage_eq.

                    theorem PartialHomeomorph.IsImage.preimage_eq {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] {e : PartialHomeomorph X Y} {s : Set X} {t : Set Y} :
                    PartialHomeomorph.IsImage e s te.source e ⁻¹' t = e.source s

                    Alias of the forward direction of PartialHomeomorph.IsImage.iff_preimage_eq.

                    theorem PartialHomeomorph.IsImage.of_symm_preimage_eq {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] {e : PartialHomeomorph X Y} {s : Set X} {t : Set Y} :
                    e.target (PartialHomeomorph.symm e) ⁻¹' s = e.target tPartialHomeomorph.IsImage e s t

                    Alias of the reverse direction of PartialHomeomorph.IsImage.iff_symm_preimage_eq.

                    theorem PartialHomeomorph.IsImage.symm_preimage_eq {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] {e : PartialHomeomorph X Y} {s : Set X} {t : Set Y} :
                    PartialHomeomorph.IsImage e s te.target (PartialHomeomorph.symm e) ⁻¹' s = e.target t

                    Alias of the forward direction of PartialHomeomorph.IsImage.iff_symm_preimage_eq.

                    theorem PartialHomeomorph.IsImage.iff_symm_preimage_eq' {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] {e : PartialHomeomorph X Y} {s : Set X} {t : Set Y} :
                    PartialHomeomorph.IsImage e s t e.target (PartialHomeomorph.symm e) ⁻¹' (e.source s) = e.target t
                    theorem PartialHomeomorph.IsImage.of_symm_preimage_eq' {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] {e : PartialHomeomorph X Y} {s : Set X} {t : Set Y} :
                    e.target (PartialHomeomorph.symm e) ⁻¹' (e.source s) = e.target tPartialHomeomorph.IsImage e s t

                    Alias of the reverse direction of PartialHomeomorph.IsImage.iff_symm_preimage_eq'.

                    theorem PartialHomeomorph.IsImage.symm_preimage_eq' {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] {e : PartialHomeomorph X Y} {s : Set X} {t : Set Y} :
                    PartialHomeomorph.IsImage e s te.target (PartialHomeomorph.symm e) ⁻¹' (e.source s) = e.target t

                    Alias of the forward direction of PartialHomeomorph.IsImage.iff_symm_preimage_eq'.

                    theorem PartialHomeomorph.IsImage.iff_preimage_eq' {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] {e : PartialHomeomorph X Y} {s : Set X} {t : Set Y} :
                    PartialHomeomorph.IsImage e s t e.source e ⁻¹' (e.target t) = e.source s
                    theorem PartialHomeomorph.IsImage.preimage_eq' {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] {e : PartialHomeomorph X Y} {s : Set X} {t : Set Y} :
                    PartialHomeomorph.IsImage e s te.source e ⁻¹' (e.target t) = e.source s

                    Alias of the forward direction of PartialHomeomorph.IsImage.iff_preimage_eq'.

                    theorem PartialHomeomorph.IsImage.of_preimage_eq' {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] {e : PartialHomeomorph X Y} {s : Set X} {t : Set Y} :
                    e.source e ⁻¹' (e.target t) = e.source sPartialHomeomorph.IsImage e s t

                    Alias of the reverse direction of PartialHomeomorph.IsImage.iff_preimage_eq'.

                    theorem PartialHomeomorph.IsImage.of_image_eq {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] {e : PartialHomeomorph X Y} {s : Set X} {t : Set Y} (h : e '' (e.source s) = e.target t) :
                    theorem PartialHomeomorph.IsImage.of_symm_image_eq {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] {e : PartialHomeomorph X Y} {s : Set X} {t : Set Y} (h : (PartialHomeomorph.symm e) '' (e.target t) = e.source s) :
                    theorem PartialHomeomorph.IsImage.inter {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] {e : PartialHomeomorph X Y} {s : Set X} {t : Set Y} {s' : Set X} {t' : Set Y} (h : PartialHomeomorph.IsImage e s t) (h' : PartialHomeomorph.IsImage e s' t') :
                    theorem PartialHomeomorph.IsImage.union {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] {e : PartialHomeomorph X Y} {s : Set X} {t : Set Y} {s' : Set X} {t' : Set Y} (h : PartialHomeomorph.IsImage e s t) (h' : PartialHomeomorph.IsImage e s' t') :
                    theorem PartialHomeomorph.IsImage.diff {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] {e : PartialHomeomorph X Y} {s : Set X} {t : Set Y} {s' : Set X} {t' : Set Y} (h : PartialHomeomorph.IsImage e s t) (h' : PartialHomeomorph.IsImage e s' t') :
                    PartialHomeomorph.IsImage e (s \ s') (t \ t')
                    theorem PartialHomeomorph.IsImage.leftInvOn_piecewise {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] {e : PartialHomeomorph X Y} {s : Set X} {t : Set Y} {e' : PartialHomeomorph X Y} [(i : X) → Decidable (i s)] [(i : Y) → Decidable (i t)] (h : PartialHomeomorph.IsImage e s t) (h' : PartialHomeomorph.IsImage e' s t) :
                    Set.LeftInvOn (Set.piecewise t (PartialHomeomorph.symm e) (PartialHomeomorph.symm e')) (Set.piecewise s e e') (Set.ite s e.source e'.source)
                    theorem PartialHomeomorph.IsImage.inter_eq_of_inter_eq_of_eqOn {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] {e : PartialHomeomorph X Y} {s : Set X} {t : Set Y} {e' : PartialHomeomorph X Y} (h : PartialHomeomorph.IsImage e s t) (h' : PartialHomeomorph.IsImage e' s t) (hs : e.source s = e'.source s) (Heq : Set.EqOn (e) (e') (e.source s)) :
                    e.target t = e'.target t
                    theorem PartialHomeomorph.IsImage.symm_eqOn_of_inter_eq_of_eqOn {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] {e : PartialHomeomorph X Y} {s : Set X} {t : Set Y} {e' : PartialHomeomorph X Y} (h : PartialHomeomorph.IsImage e s t) (hs : e.source s = e'.source s) (Heq : Set.EqOn (e) (e') (e.source s)) :
                    theorem PartialHomeomorph.IsImage.map_nhdsWithin_eq {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] {e : PartialHomeomorph X Y} {s : Set X} {t : Set Y} {x : X} (h : PartialHomeomorph.IsImage e s t) (hx : x e.source) :
                    Filter.map (e) (nhdsWithin x s) = nhdsWithin (e x) t
                    theorem PartialHomeomorph.IsImage.isOpen_iff {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] {e : PartialHomeomorph X Y} {s : Set X} {t : Set Y} (h : PartialHomeomorph.IsImage e s t) :
                    IsOpen (e.source s) IsOpen (e.target t)
                    @[simp]
                    theorem PartialHomeomorph.IsImage.restr_toPartialEquiv {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] {e : PartialHomeomorph X Y} {s : Set X} {t : Set Y} (h : PartialHomeomorph.IsImage e s t) (hs : IsOpen (e.source s)) :
                    def PartialHomeomorph.IsImage.restr {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] {e : PartialHomeomorph X Y} {s : Set X} {t : Set Y} (h : PartialHomeomorph.IsImage e s t) (hs : IsOpen (e.source s)) :

                    Restrict a PartialHomeomorph to a pair of corresponding open sets.

                    Equations
                    Instances For
                      theorem PartialHomeomorph.isImage_source_target_of_disjoint {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] (e : PartialHomeomorph X Y) (e' : PartialHomeomorph X Y) (hs : Disjoint e.source e'.source) (ht : Disjoint e.target e'.target) :
                      PartialHomeomorph.IsImage e e'.source e'.target
                      theorem PartialHomeomorph.preimage_interior {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] (e : PartialHomeomorph X Y) (s : Set Y) :
                      e.source e ⁻¹' interior s = e.source interior (e ⁻¹' s)

                      Preimage of interior or interior of preimage coincide for partial homeomorphisms, when restricted to the source.

                      theorem PartialHomeomorph.preimage_closure {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] (e : PartialHomeomorph X Y) (s : Set Y) :
                      e.source e ⁻¹' closure s = e.source closure (e ⁻¹' s)
                      theorem PartialHomeomorph.preimage_frontier {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] (e : PartialHomeomorph X Y) (s : Set Y) :
                      e.source e ⁻¹' frontier s = e.source frontier (e ⁻¹' s)
                      def PartialHomeomorph.ofContinuousOpenRestrict {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] (e : PartialEquiv X Y) (hc : ContinuousOn (e) e.source) (ho : IsOpenMap (Set.restrict e.source e)) (hs : IsOpen e.source) :

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

                      Equations
                      Instances For
                        def PartialHomeomorph.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) :

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

                        Equations
                        Instances For

                          Restricting a partial homeomorphism e to e.source ∩ s when s is open. This is sometimes hard to use because of the openness assumption, but it has the advantage that when it can be used then its PartialEquiv is defeq to PartialEquiv.restr.

                          Equations
                          Instances For
                            @[simp]
                            theorem PartialHomeomorph.restrOpen_toPartialEquiv {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] (e : PartialHomeomorph X Y) (s : Set X) (hs : IsOpen s) :
                            (PartialHomeomorph.restrOpen e s hs).toPartialEquiv = PartialEquiv.restr e.toPartialEquiv s
                            theorem PartialHomeomorph.restrOpen_source {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] (e : PartialHomeomorph X Y) (s : Set X) (hs : IsOpen s) :
                            (PartialHomeomorph.restrOpen e s hs).source = e.source s
                            @[simp]
                            theorem PartialHomeomorph.restr_apply {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] (e : PartialHomeomorph X Y) (s : Set X) :
                            theorem PartialHomeomorph.restr_source {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] (e : PartialHomeomorph X Y) (s : Set X) :
                            (PartialHomeomorph.restr e s).source = e.source interior s

                            Restricting a partial homeomorphism e to e.source ∩ interior s. We use the interior to make sure that the restriction is well defined whatever the set s, since partial homeomorphisms are by definition defined on open sets. In applications where s is open, this coincides with the restriction of partial equivalences

                            Equations
                            Instances For
                              @[simp]
                              theorem PartialHomeomorph.restr_toPartialEquiv {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] (e : PartialHomeomorph X Y) (s : Set X) :
                              (PartialHomeomorph.restr e s).toPartialEquiv = PartialEquiv.restr e.toPartialEquiv (interior s)
                              theorem PartialHomeomorph.restr_source' {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] (e : PartialHomeomorph X Y) (s : Set X) (hs : IsOpen s) :
                              (PartialHomeomorph.restr e s).source = e.source s
                              theorem PartialHomeomorph.restr_toPartialEquiv' {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] (e : PartialHomeomorph X Y) (s : Set X) (hs : IsOpen s) :
                              (PartialHomeomorph.restr e s).toPartialEquiv = PartialEquiv.restr e.toPartialEquiv s

                              The identity on the whole space as a partial homeomorphism.

                              Equations
                              Instances For

                                ofSet: the identity on a set s

                                @[simp]
                                theorem PartialHomeomorph.ofSet_apply {X : Type u_1} [TopologicalSpace X] (s : Set X) (hs : IsOpen s) :
                                theorem PartialHomeomorph.ofSet_target {X : Type u_1} [TopologicalSpace X] (s : Set X) (hs : IsOpen s) :
                                (PartialHomeomorph.ofSet s hs).target = s
                                theorem PartialHomeomorph.ofSet_source {X : Type u_1} [TopologicalSpace X] (s : Set X) (hs : IsOpen s) :
                                (PartialHomeomorph.ofSet s hs).source = s

                                The identity partial equivalence on a set s

                                Equations
                                Instances For
                                  @[simp]

                                  trans: composition of two partial homeomorphisms

                                  @[simp]
                                  theorem PartialHomeomorph.trans'_symm_apply {X : Type u_1} {Y : Type u_3} {Z : Type u_5} [TopologicalSpace X] [TopologicalSpace Y] [TopologicalSpace Z] (e : PartialHomeomorph X Y) (e' : PartialHomeomorph Y Z) (h : e.target = e'.source) :
                                  @[simp]
                                  theorem PartialHomeomorph.trans'_toPartialEquiv {X : Type u_1} {Y : Type u_3} {Z : Type u_5} [TopologicalSpace X] [TopologicalSpace Y] [TopologicalSpace Z] (e : PartialHomeomorph X Y) (e' : PartialHomeomorph Y Z) (h : e.target = e'.source) :
                                  (PartialHomeomorph.trans' e e' h).toPartialEquiv = PartialEquiv.trans' e.toPartialEquiv e'.toPartialEquiv h
                                  @[simp]
                                  theorem PartialHomeomorph.trans'_apply {X : Type u_1} {Y : Type u_3} {Z : Type u_5} [TopologicalSpace X] [TopologicalSpace Y] [TopologicalSpace Z] (e : PartialHomeomorph X Y) (e' : PartialHomeomorph Y Z) (h : e.target = e'.source) :
                                  ∀ (a : X), (PartialHomeomorph.trans' e e' h) a = e' (e a)
                                  theorem PartialHomeomorph.trans'_target {X : Type u_1} {Y : Type u_3} {Z : Type u_5} [TopologicalSpace X] [TopologicalSpace Y] [TopologicalSpace Z] (e : PartialHomeomorph X Y) (e' : PartialHomeomorph Y Z) (h : e.target = e'.source) :
                                  (PartialHomeomorph.trans' e e' h).target = e'.target
                                  theorem PartialHomeomorph.trans'_source {X : Type u_1} {Y : Type u_3} {Z : Type u_5} [TopologicalSpace X] [TopologicalSpace Y] [TopologicalSpace Z] (e : PartialHomeomorph X Y) (e' : PartialHomeomorph Y Z) (h : e.target = e'.source) :
                                  (PartialHomeomorph.trans' e e' h).source = e.source
                                  def PartialHomeomorph.trans' {X : Type u_1} {Y : Type u_3} {Z : Type u_5} [TopologicalSpace X] [TopologicalSpace Y] [TopologicalSpace Z] (e : PartialHomeomorph X Y) (e' : PartialHomeomorph Y Z) (h : e.target = e'.source) :

                                  Composition of two partial homeomorphisms when the target of the first and the source of the second coincide.

                                  Equations
                                  • One or more equations did not get rendered due to their size.
                                  Instances For

                                    Composing two partial homeomorphisms, by restricting to the maximal domain where their composition is well defined.

                                    Equations
                                    • One or more equations did not get rendered due to their size.
                                    Instances For
                                      @[simp]
                                      theorem PartialHomeomorph.trans_toPartialEquiv {X : Type u_1} {Y : Type u_3} {Z : Type u_5} [TopologicalSpace X] [TopologicalSpace Y] [TopologicalSpace Z] (e : PartialHomeomorph X Y) (e' : PartialHomeomorph Y Z) :
                                      (PartialHomeomorph.trans e e').toPartialEquiv = PartialEquiv.trans e.toPartialEquiv e'.toPartialEquiv
                                      @[simp]
                                      theorem PartialHomeomorph.coe_trans {X : Type u_1} {Y : Type u_3} {Z : Type u_5} [TopologicalSpace X] [TopologicalSpace Y] [TopologicalSpace Z] (e : PartialHomeomorph X Y) (e' : PartialHomeomorph Y Z) :
                                      (PartialHomeomorph.trans e e') = e' e
                                      theorem PartialHomeomorph.trans_apply {X : Type u_1} {Y : Type u_3} {Z : Type u_5} [TopologicalSpace X] [TopologicalSpace Y] [TopologicalSpace Z] (e : PartialHomeomorph X Y) (e' : PartialHomeomorph Y Z) {x : X} :
                                      (PartialHomeomorph.trans e e') x = e' (e x)
                                      theorem PartialHomeomorph.trans_source {X : Type u_1} {Y : Type u_3} {Z : Type u_5} [TopologicalSpace X] [TopologicalSpace Y] [TopologicalSpace Z] (e : PartialHomeomorph X Y) (e' : PartialHomeomorph Y Z) :
                                      (PartialHomeomorph.trans e e').source = e.source e ⁻¹' e'.source
                                      theorem PartialHomeomorph.trans_source' {X : Type u_1} {Y : Type u_3} {Z : Type u_5} [TopologicalSpace X] [TopologicalSpace Y] [TopologicalSpace Z] (e : PartialHomeomorph X Y) (e' : PartialHomeomorph Y Z) :
                                      (PartialHomeomorph.trans e e').source = e.source e ⁻¹' (e.target e'.source)
                                      theorem PartialHomeomorph.trans_source'' {X : Type u_1} {Y : Type u_3} {Z : Type u_5} [TopologicalSpace X] [TopologicalSpace Y] [TopologicalSpace Z] (e : PartialHomeomorph X Y) (e' : PartialHomeomorph Y Z) :
                                      (PartialHomeomorph.trans e e').source = (PartialHomeomorph.symm e) '' (e.target e'.source)
                                      theorem PartialHomeomorph.image_trans_source {X : Type u_1} {Y : Type u_3} {Z : Type u_5} [TopologicalSpace X] [TopologicalSpace Y] [TopologicalSpace Z] (e : PartialHomeomorph X Y) (e' : PartialHomeomorph Y Z) :
                                      e '' (PartialHomeomorph.trans e e').source = e.target e'.source
                                      theorem PartialHomeomorph.trans_target {X : Type u_1} {Y : Type u_3} {Z : Type u_5} [TopologicalSpace X] [TopologicalSpace Y] [TopologicalSpace Z] (e : PartialHomeomorph X Y) (e' : PartialHomeomorph Y Z) :
                                      (PartialHomeomorph.trans e e').target = e'.target (PartialHomeomorph.symm e') ⁻¹' e.target
                                      theorem PartialHomeomorph.trans_target' {X : Type u_1} {Y : Type u_3} {Z : Type u_5} [TopologicalSpace X] [TopologicalSpace Y] [TopologicalSpace Z] (e : PartialHomeomorph X Y) (e' : PartialHomeomorph Y Z) :
                                      (PartialHomeomorph.trans e e').target = e'.target (PartialHomeomorph.symm e') ⁻¹' (e'.source e.target)
                                      theorem PartialHomeomorph.trans_target'' {X : Type u_1} {Y : Type u_3} {Z : Type u_5} [TopologicalSpace X] [TopologicalSpace Y] [TopologicalSpace Z] (e : PartialHomeomorph X Y) (e' : PartialHomeomorph Y Z) :
                                      (PartialHomeomorph.trans e e').target = e' '' (e'.source e.target)

                                      EqOnSource: equivalence on their source

                                      EqOnSource e e' means that e and e' have the same source, and coincide there. They should really be considered the same partial equivalence.

                                      Equations
                                      Instances For

                                        EqOnSource is an equivalence relation.

                                        Equations
                                        • PartialHomeomorph.eqOnSourceSetoid = let __src := Setoid.comap PartialHomeomorph.toPartialEquiv PartialEquiv.eqOnSourceSetoid; { r := PartialHomeomorph.EqOnSource, iseqv := }

                                        If two partial homeomorphisms are equivalent, so are their inverses.

                                        theorem PartialHomeomorph.EqOnSource.source_eq {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] {e : PartialHomeomorph X Y} {e' : PartialHomeomorph X Y} (h : e e') :
                                        e.source = e'.source

                                        Two equivalent partial homeomorphisms have the same source.

                                        theorem PartialHomeomorph.EqOnSource.target_eq {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] {e : PartialHomeomorph X Y} {e' : PartialHomeomorph X Y} (h : e e') :
                                        e.target = e'.target

                                        Two equivalent partial homeomorphisms have the same target.

                                        theorem PartialHomeomorph.EqOnSource.eqOn {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] {e : PartialHomeomorph X Y} {e' : PartialHomeomorph X Y} (h : e e') :
                                        Set.EqOn (e) (e') e.source

                                        Two equivalent partial homeomorphisms have coinciding toFun on the source

                                        Two equivalent partial homeomorphisms have coinciding invFun on the target

                                        Composition of partial homeomorphisms respects equivalence.

                                        Restriction of partial homeomorphisms respects equivalence

                                        theorem PartialHomeomorph.Set.EqOn.restr_eqOn_source {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] {e : PartialHomeomorph X Y} {e' : PartialHomeomorph X Y} (h : Set.EqOn (e) (e') (e.source e'.source)) :

                                        Two equivalent partial homeomorphisms are equal when the source and target are univ.

                                        Composition of a partial homeomorphism and its inverse is equivalent to the restriction of the identity to the source

                                        theorem PartialHomeomorph.eq_of_eqOnSource_univ {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] {e : PartialHomeomorph X Y} {e' : PartialHomeomorph X Y} (h : e e') (s : e.source = Set.univ) (t : e.target = Set.univ) :
                                        e = e'

                                        product of two partial homeomorphisms

                                        @[simp]
                                        theorem PartialHomeomorph.prod_toPartialEquiv {X : Type u_1} {X' : Type u_2} {Y : Type u_3} {Y' : Type u_4} [TopologicalSpace X] [TopologicalSpace X'] [TopologicalSpace Y] [TopologicalSpace Y'] (eX : PartialHomeomorph X X') (eY : PartialHomeomorph Y Y') :
                                        (PartialHomeomorph.prod eX eY).toPartialEquiv = PartialEquiv.prod eX.toPartialEquiv eY.toPartialEquiv
                                        @[simp]
                                        theorem PartialHomeomorph.prod_apply {X : Type u_1} {X' : Type u_2} {Y : Type u_3} {Y' : Type u_4} [TopologicalSpace X] [TopologicalSpace X'] [TopologicalSpace Y] [TopologicalSpace Y'] (eX : PartialHomeomorph X X') (eY : PartialHomeomorph Y Y') :
                                        (PartialHomeomorph.prod eX eY) = fun (p : X × Y) => (eX p.1, eY p.2)
                                        theorem PartialHomeomorph.prod_symm_apply {X : Type u_1} {X' : Type u_2} {Y : Type u_3} {Y' : Type u_4} [TopologicalSpace X] [TopologicalSpace X'] [TopologicalSpace Y] [TopologicalSpace Y'] (eX : PartialHomeomorph X X') (eY : PartialHomeomorph Y Y') (p : X' × Y') :
                                        theorem PartialHomeomorph.prod_target {X : Type u_1} {X' : Type u_2} {Y : Type u_3} {Y' : Type u_4} [TopologicalSpace X] [TopologicalSpace X'] [TopologicalSpace Y] [TopologicalSpace Y'] (eX : PartialHomeomorph X X') (eY : PartialHomeomorph Y Y') :
                                        (PartialHomeomorph.prod eX eY).target = eX.target ×ˢ eY.target
                                        theorem PartialHomeomorph.prod_source {X : Type u_1} {X' : Type u_2} {Y : Type u_3} {Y' : Type u_4} [TopologicalSpace X] [TopologicalSpace X'] [TopologicalSpace Y] [TopologicalSpace Y'] (eX : PartialHomeomorph X X') (eY : PartialHomeomorph Y Y') :
                                        (PartialHomeomorph.prod eX eY).source = eX.source ×ˢ eY.source
                                        def PartialHomeomorph.prod {X : Type u_1} {X' : Type u_2} {Y : Type u_3} {Y' : Type u_4} [TopologicalSpace X] [TopologicalSpace X'] [TopologicalSpace Y] [TopologicalSpace Y'] (eX : PartialHomeomorph X X') (eY : PartialHomeomorph Y Y') :
                                        PartialHomeomorph (X × Y) (X' × Y')

                                        The product of two partial homeomorphisms, as a partial homeomorphism on the product space.

                                        Equations
                                        • PartialHomeomorph.prod eX eY = { toPartialEquiv := PartialEquiv.prod eX.toPartialEquiv eY.toPartialEquiv, open_source := , open_target := , continuousOn_toFun := , continuousOn_invFun := }
                                        Instances For
                                          theorem PartialHomeomorph.prod_eq_prod_of_nonempty {X : Type u_1} {X' : Type u_2} {Y : Type u_3} {Y' : Type u_4} [TopologicalSpace X] [TopologicalSpace X'] [TopologicalSpace Y] [TopologicalSpace Y'] {eX : PartialHomeomorph X X'} {eX' : PartialHomeomorph X X'} {eY : PartialHomeomorph Y Y'} {eY' : PartialHomeomorph Y Y'} (h : Set.Nonempty (PartialHomeomorph.prod eX eY).source) :
                                          theorem PartialHomeomorph.prod_eq_prod_of_nonempty' {X : Type u_1} {X' : Type u_2} {Y : Type u_3} {Y' : Type u_4} [TopologicalSpace X] [TopologicalSpace X'] [TopologicalSpace Y] [TopologicalSpace Y'] {eX : PartialHomeomorph X X'} {eX' : PartialHomeomorph X X'} {eY : PartialHomeomorph Y Y'} {eY' : PartialHomeomorph Y Y'} (h : Set.Nonempty (PartialHomeomorph.prod eX' eY').source) :

                                          finite product of partial homeomorphisms

                                          @[simp]
                                          theorem PartialHomeomorph.pi_toPartialEquiv {ι : Type u_7} [Finite ι] {X : ιType u_8} {Y : ιType u_9} [(i : ι) → TopologicalSpace (X i)] [(i : ι) → TopologicalSpace (Y i)] (ei : (i : ι) → PartialHomeomorph (X i) (Y i)) :
                                          (PartialHomeomorph.pi ei).toPartialEquiv = PartialEquiv.pi fun (i : ι) => (ei i).toPartialEquiv
                                          def PartialHomeomorph.pi {ι : Type u_7} [Finite ι] {X : ιType u_8} {Y : ιType u_9} [(i : ι) → TopologicalSpace (X i)] [(i : ι) → TopologicalSpace (Y i)] (ei : (i : ι) → PartialHomeomorph (X i) (Y i)) :
                                          PartialHomeomorph ((i : ι) → X i) ((i : ι) → Y i)

                                          The product of a finite family of PartialHomeomorphs.

                                          Equations
                                          • PartialHomeomorph.pi ei = { toPartialEquiv := PartialEquiv.pi fun (i : ι) => (ei i).toPartialEquiv, open_source := , open_target := , continuousOn_toFun := , continuousOn_invFun := }
                                          Instances For

                                            combining two partial homeomorphisms using Set.piecewise

                                            @[simp]
                                            theorem PartialHomeomorph.piecewise_apply {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] (e : PartialHomeomorph X Y) (e' : PartialHomeomorph X Y) (s : Set X) (t : Set Y) [(x : X) → Decidable (x s)] [(y : Y) → Decidable (y t)] (H : PartialHomeomorph.IsImage e s t) (H' : PartialHomeomorph.IsImage e' s t) (Hs : e.source frontier s = e'.source frontier s) (Heq : Set.EqOn (e) (e') (e.source frontier s)) :
                                            (PartialHomeomorph.piecewise e e' s t H H' Hs Heq) = Set.piecewise s e e'
                                            @[simp]
                                            theorem PartialHomeomorph.piecewise_toPartialEquiv {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] (e : PartialHomeomorph X Y) (e' : PartialHomeomorph X Y) (s : Set X) (t : Set Y) [(x : X) → Decidable (x s)] [(y : Y) → Decidable (y t)] (H : PartialHomeomorph.IsImage e s t) (H' : PartialHomeomorph.IsImage e' s t) (Hs : e.source frontier s = e'.source frontier s) (Heq : Set.EqOn (e) (e') (e.source frontier s)) :
                                            (PartialHomeomorph.piecewise e e' s t H H' Hs Heq).toPartialEquiv = PartialEquiv.piecewise e.toPartialEquiv e'.toPartialEquiv s t H H'
                                            def PartialHomeomorph.piecewise {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] (e : PartialHomeomorph X Y) (e' : PartialHomeomorph X Y) (s : Set X) (t : Set Y) [(x : X) → Decidable (x s)] [(y : Y) → Decidable (y t)] (H : PartialHomeomorph.IsImage e s t) (H' : PartialHomeomorph.IsImage e' s t) (Hs : e.source frontier s = e'.source frontier s) (Heq : Set.EqOn (e) (e') (e.source frontier s)) :

                                            Combine two PartialHomeomorphs using Set.piecewise. The source of the new PartialHomeomorph is s.ite e.source e'.source = e.source ∩ s ∪ e'.source \ s, and similarly for target. The function sends e.source ∩ s to e.target ∩ t using e and e'.source \ s to e'.target \ t using e', and similarly for the inverse function. To ensure the maps toFun and invFun are inverse of each other on the new source and target, the definition assumes that the sets s and t are related both by e.is_image and e'.is_image. To ensure that the new maps are continuous on source/target, it also assumes that e.source and e'.source meet frontier s on the same set and e x = e' x on this intersection.

                                            Equations
                                            • One or more equations did not get rendered due to their size.
                                            Instances For
                                              @[simp]
                                              theorem PartialHomeomorph.symm_piecewise {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] (e : PartialHomeomorph X Y) (e' : PartialHomeomorph X Y) {s : Set X} {t : Set Y} [(x : X) → Decidable (x s)] [(y : Y) → Decidable (y t)] (H : PartialHomeomorph.IsImage e s t) (H' : PartialHomeomorph.IsImage e' s t) (Hs : e.source frontier s = e'.source frontier s) (Heq : Set.EqOn (e) (e') (e.source frontier s)) :
                                              def PartialHomeomorph.disjointUnion {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] (e : PartialHomeomorph X Y) (e' : PartialHomeomorph X Y) [(x : X) → Decidable (x e.source)] [(y : Y) → Decidable (y e.target)] (Hs : Disjoint e.source e'.source) (Ht : Disjoint e.target e'.target) :

                                              Combine two PartialHomeomorphs with disjoint sources and disjoint targets. We reuse PartialHomeomorph.piecewise then override toPartialEquiv to PartialEquiv.disjointUnion. This way we have better definitional equalities for source and target.

                                              Equations
                                              • One or more equations did not get rendered due to their size.
                                              Instances For
                                                theorem PartialHomeomorph.continuousWithinAt_iff_continuousWithinAt_comp_right {X : Type u_1} {Y : Type u_3} {Z : Type u_5} [TopologicalSpace X] [TopologicalSpace Y] [TopologicalSpace Z] (e : PartialHomeomorph X Y) {f : YZ} {s : Set Y} {x : Y} (h : x e.target) :

                                                Continuity within a set at a point can be read under right composition with a local homeomorphism, if the point is in its target

                                                theorem PartialHomeomorph.continuousAt_iff_continuousAt_comp_right {X : Type u_1} {Y : Type u_3} {Z : Type u_5} [TopologicalSpace X] [TopologicalSpace Y] [TopologicalSpace Z] (e : PartialHomeomorph X Y) {f : YZ} {x : Y} (h : x e.target) :

                                                Continuity at a point can be read under right composition with a partial homeomorphism, if the point is in its target

                                                theorem PartialHomeomorph.continuousOn_iff_continuousOn_comp_right {X : Type u_1} {Y : Type u_3} {Z : Type u_5} [TopologicalSpace X] [TopologicalSpace Y] [TopologicalSpace Z] (e : PartialHomeomorph X Y) {f : YZ} {s : Set Y} (h : s e.target) :
                                                ContinuousOn f s ContinuousOn (f e) (e.source e ⁻¹' s)

                                                A function is continuous on a set if and only if its composition with a partial homeomorphism on the right is continuous on the corresponding set.

                                                theorem PartialHomeomorph.continuousWithinAt_iff_continuousWithinAt_comp_left {X : Type u_1} {Y : Type u_3} {Z : Type u_5} [TopologicalSpace X] [TopologicalSpace Y] [TopologicalSpace Z] (e : PartialHomeomorph X Y) {f : ZX} {s : Set Z} {x : Z} (hx : f x e.source) (h : f ⁻¹' e.source nhdsWithin x s) :

                                                Continuity within a set at a point can be read under left composition with a local homeomorphism if a neighborhood of the initial point is sent to the source of the local homeomorphism

                                                theorem PartialHomeomorph.continuousAt_iff_continuousAt_comp_left {X : Type u_1} {Y : Type u_3} {Z : Type u_5} [TopologicalSpace X] [TopologicalSpace Y] [TopologicalSpace Z] (e : PartialHomeomorph X Y) {f : ZX} {x : Z} (h : f ⁻¹' e.source nhds x) :

                                                Continuity at a point can be read under left composition with a partial homeomorphism if a neighborhood of the initial point is sent to the source of the partial homeomorphism

                                                theorem PartialHomeomorph.continuousOn_iff_continuousOn_comp_left {X : Type u_1} {Y : Type u_3} {Z : Type u_5} [TopologicalSpace X] [TopologicalSpace Y] [TopologicalSpace Z] (e : PartialHomeomorph X Y) {f : ZX} {s : Set Z} (h : s f ⁻¹' e.source) :

                                                A function is continuous on a set if and only if its composition with a partial homeomorphism on the left is continuous on the corresponding set.

                                                theorem PartialHomeomorph.continuous_iff_continuous_comp_left {X : Type u_1} {Y : Type u_3} {Z : Type u_5} [TopologicalSpace X] [TopologicalSpace Y] [TopologicalSpace Z] (e : PartialHomeomorph X Y) {f : ZX} (h : f ⁻¹' e.source = Set.univ) :

                                                A function is continuous if and only if its composition with a partial homeomorphism on the left is continuous and its image is contained in the source.

                                                @[simp]
                                                theorem PartialHomeomorph.homeomorphOfImageSubsetSource_apply {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] (e : PartialHomeomorph X Y) {s : Set X} {t : Set Y} (hs : s e.source) (ht : e '' s = t) :
                                                ∀ (a : s), (PartialHomeomorph.homeomorphOfImageSubsetSource e hs ht) a = Set.MapsTo.restrict (e) s t a
                                                @[simp]
                                                theorem PartialHomeomorph.homeomorphOfImageSubsetSource_symm_apply {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] (e : PartialHomeomorph X Y) {s : Set X} {t : Set Y} (hs : s e.source) (ht : e '' s = t) :
                                                def PartialHomeomorph.homeomorphOfImageSubsetSource {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] (e : PartialHomeomorph X Y) {s : Set X} {t : Set Y} (hs : s e.source) (ht : e '' s = t) :
                                                s ≃ₜ t

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

                                                Equations
                                                • One or more equations did not get rendered due to their size.
                                                Instances For
                                                  @[simp]
                                                  def PartialHomeomorph.toHomeomorphSourceTarget {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] (e : PartialHomeomorph X Y) :
                                                  e.source ≃ₜ e.target

                                                  A partial homeomorphism defines a homeomorphism between its source and target.

                                                  Equations
                                                  Instances For
                                                    theorem PartialHomeomorph.nhds_eq_comap_inf_principal {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] (e : PartialHomeomorph X Y) {x : X} (hx : x e.source) :
                                                    nhds x = Filter.comap (e) (nhds (e x)) Filter.principal e.source
                                                    @[simp]
                                                    theorem PartialHomeomorph.toHomeomorphOfSourceEqUnivTargetEqUniv_apply {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] (e : PartialHomeomorph X Y) (h : e.source = Set.univ) (h' : e.target = Set.univ) :
                                                    def PartialHomeomorph.toHomeomorphOfSourceEqUnivTargetEqUniv {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] (e : PartialHomeomorph X Y) (h : e.source = Set.univ) (h' : e.target = Set.univ) :
                                                    X ≃ₜ Y

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

                                                    Equations
                                                    • One or more equations did not get rendered due to their size.
                                                    Instances For
                                                      theorem PartialHomeomorph.to_openEmbedding {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] (e : PartialHomeomorph X Y) (h : e.source = Set.univ) :

                                                      A partial homeomorphism whose source is all of X defines an open embedding of X into Y. The converse is also true; see OpenEmbedding.toPartialHomeomorph.

                                                      @[simp]
                                                      @[simp]
                                                      theorem Homeomorph.transPartialHomeomorph_target {X : Type u_1} {Y : Type u_3} {Z : Type u_5} [TopologicalSpace X] [TopologicalSpace Y] [TopologicalSpace Z] (e : X ≃ₜ Y) (f' : PartialHomeomorph Y Z) :
                                                      (Homeomorph.transPartialHomeomorph e f').target = f'.target
                                                      @[simp]
                                                      theorem Homeomorph.transPartialHomeomorph_source {X : Type u_1} {Y : Type u_3} {Z : Type u_5} [TopologicalSpace X] [TopologicalSpace Y] [TopologicalSpace Z] (e : X ≃ₜ Y) (f' : PartialHomeomorph Y Z) :
                                                      (Homeomorph.transPartialHomeomorph e f').source = e ⁻¹' f'.source

                                                      Precompose a partial homeomorphism with a homeomorphism. We modify the source and target to have better definitional behavior.

                                                      Equations
                                                      • One or more equations did not get rendered due to their size.
                                                      Instances For
                                                        @[simp]
                                                        theorem OpenEmbedding.toPartialHomeomorph_source {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] (f : XY) (h : OpenEmbedding f) [Nonempty X] :
                                                        (OpenEmbedding.toPartialHomeomorph f h).source = Set.univ
                                                        noncomputable def OpenEmbedding.toPartialHomeomorph {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] (f : XY) (h : OpenEmbedding f) [Nonempty X] :

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

                                                        Equations
                                                        Instances For

                                                          inclusion of an open set in a topological space

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

                                                          Equations
                                                          Instances For
                                                            @[simp]
                                                            @[simp]
                                                            theorem PartialHomeomorph.transHomeomorph_source {X : Type u_1} {Y : Type u_3} {Z : Type u_5} [TopologicalSpace X] [TopologicalSpace Y] [TopologicalSpace Z] (e : PartialHomeomorph X Y) (f' : Y ≃ₜ Z) :
                                                            (PartialHomeomorph.transHomeomorph e f').source = e.source
                                                            @[simp]

                                                            Postcompose a partial homeomorphism with a homeomorphism. We modify the source and target to have better definitional behavior.

                                                            Equations
                                                            • One or more equations did not get rendered due to their size.
                                                            Instances For

                                                              subtypeRestr: restriction to a subtype

                                                              noncomputable def PartialHomeomorph.subtypeRestr {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] (e : PartialHomeomorph X Y) {s : TopologicalSpace.Opens X} (hs : Nonempty s) :

                                                              The restriction of a partial homeomorphism e to an open subset s of the domain type produces a partial homeomorphism whose domain is the subtype s.

                                                              Equations
                                                              Instances For
                                                                @[simp]
                                                                theorem PartialHomeomorph.subtypeRestr_source {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] (e : PartialHomeomorph X Y) {s : TopologicalSpace.Opens X} (hs : Nonempty s) :
                                                                (PartialHomeomorph.subtypeRestr e hs).source = Subtype.val ⁻¹' e.source
                                                                theorem PartialHomeomorph.map_subtype_source {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] (e : PartialHomeomorph X Y) {s : TopologicalSpace.Opens X} (hs : Nonempty s) {x : s} (hxe : x e.source) :
                                                                e x (PartialHomeomorph.subtypeRestr e hs).target

                                                                This lemma characterizes the transition functions of an open subset in terms of the transition functions of the original space.