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 X Y :
Type (max u_7 u_8)

Partial homeomorphisms, defined on open subsets of the space

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
    Instances For

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

      Equations
      Instances For

        The inverse of a partial homeomorphism

        Equations
        • Eq e.symm { toPartialEquiv := e.symm, open_source := , open_target := , continuousOn_toFun := , continuousOn_invFun := }
        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
              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.toFun e.source) (d : ContinuousOn e.invFun e.target) :
              Eq { toPartialEquiv := e, open_source := a, open_target := b, continuousOn_toFun := c, continuousOn_invFun := d }.toFun' e.toFun
              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.toFun e.source) (d : ContinuousOn e.invFun e.target) :
              Eq { toPartialEquiv := e, open_source := a, open_target := b, continuousOn_toFun := c, continuousOn_invFun := d }.symm.toFun' e.symm.toFun

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

              theorem PartialHomeomorph.left_inv {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] (e : PartialHomeomorph X Y) {x : X} (h : Membership.mem e.source x) :
              Eq (e.symm.toFun' (e.toFun' x)) x
              theorem PartialHomeomorph.right_inv {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] (e : PartialHomeomorph X Y) {x : Y} (h : Membership.mem e.target x) :
              Eq (e.toFun' (e.symm.toFun' 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 : Membership.mem e.source x) (hy : Membership.mem e.target y) :
              Iff (Eq x (e.symm.toFun' y)) (Eq (e.toFun' x) y)
              def Homeomorph.toPartialHomeomorphOfImageEq {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] (e : Homeomorph X Y) (s : Set X) (hs : IsOpen s) (t : Set Y) (h : Eq (Set.image (DFunLike.coe 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
              Instances For
                theorem Homeomorph.toPartialHomeomorphOfImageEq_source {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] (e : Homeomorph X Y) (s : Set X) (hs : IsOpen s) (t : Set Y) (h : Eq (Set.image (DFunLike.coe e) s) t) :
                theorem Homeomorph.toPartialHomeomorphOfImageEq_target {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] (e : Homeomorph X Y) (s : Set X) (hs : IsOpen s) (t : Set Y) (h : Eq (Set.image (DFunLike.coe e) s) t) :

                A homeomorphism induces a partial homeomorphism on the whole space

                Equations
                Instances For

                  Replace toPartialEquiv field to provide better definitional equalities.

                  Equations
                  • Eq (e.replaceEquiv e' h) { toPartialEquiv := e', open_source := , open_target := , continuousOn_toFun := , continuousOn_invFun := }
                  Instances For
                    theorem PartialHomeomorph.eventually_left_inverse {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] (e : PartialHomeomorph X Y) {x : X} (hx : Membership.mem e.source x) :
                    Filter.Eventually (fun (y : X) => Eq (e.symm.toFun' (e.toFun' y)) y) (nhds x)
                    theorem PartialHomeomorph.eventually_left_inverse' {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] (e : PartialHomeomorph X Y) {x : Y} (hx : Membership.mem e.target x) :
                    Filter.Eventually (fun (y : X) => Eq (e.symm.toFun' (e.toFun' y)) y) (nhds (e.symm.toFun' x))
                    theorem PartialHomeomorph.eventually_right_inverse {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] (e : PartialHomeomorph X Y) {x : Y} (hx : Membership.mem e.target x) :
                    Filter.Eventually (fun (y : Y) => Eq (e.toFun' (e.symm.toFun' y)) y) (nhds x)
                    theorem PartialHomeomorph.eventually_right_inverse' {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] (e : PartialHomeomorph X Y) {x : X} (hx : Membership.mem e.source x) :
                    Filter.Eventually (fun (y : Y) => Eq (e.toFun' (e.symm.toFun' y)) y) (nhds (e.toFun' x))
                    theorem PartialHomeomorph.ext {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] (e e' : PartialHomeomorph X Y) (h : ∀ (x : X), Eq (e.toFun' x) (e'.toFun' x)) (hinv : ∀ (x : Y), Eq (e.symm.toFun' x) (e'.symm.toFun' x)) (hs : Eq e.source e'.source) :
                    Eq 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 e' : PartialHomeomorph X Y} :
                    Iff (Eq e e') (And (∀ (x : X), Eq (e.toFun' x) (e'.toFun' x)) (And (∀ (x : Y), Eq (e.symm.toFun' x) (e'.symm.toFun' x)) (Eq e.source e'.source)))

                    A partial homeomorphism is continuous at any point of its source

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

                    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 : Membership.mem e.source x) :
                    Iff (Filter.Eventually (fun (y : Y) => p y) (nhds (e.toFun' x))) (Filter.Eventually (fun (x : X) => p (e.toFun' x)) (nhds 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 : Membership.mem e.source x) :
                    Iff (Filter.Eventually (fun (y : Y) => p (e.symm.toFun' y)) (nhds (e.toFun' x))) (Filter.Eventually (fun (x : X) => p x) (nhds 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 : Membership.mem e.source x) :
                    Iff (Filter.Eventually (fun (y : Y) => p y) (nhdsWithin (e.toFun' x) (Set.preimage e.symm.toFun' s))) (Filter.Eventually (fun (x : X) => p (e.toFun' x)) (nhdsWithin x s))
                    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 : Membership.mem e.source x) :
                    Iff (Filter.Eventually (fun (y : Y) => p (e.symm.toFun' y)) (nhdsWithin (e.toFun' x) (Set.preimage e.symm.toFun' s))) (Filter.Eventually (fun (x : X) => p x) (nhdsWithin x s))

                    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).

                    A 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 a partial homeomorphism e is an open map on e.target.

                    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 : e.IsImage s t) :
                      e.IsImage 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 : e.IsImage s t) (hx : Membership.mem e.source x) :
                      theorem PartialHomeomorph.IsImage.symm {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] {e : PartialHomeomorph X Y} {s : Set X} {t : Set Y} (h : e.IsImage s t) :
                      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 : e.IsImage s t) (hy : Membership.mem e.target y) :
                      theorem PartialHomeomorph.IsImage.symm_iff {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] {e : PartialHomeomorph X Y} {s : Set X} {t : Set Y} :
                      Iff (e.symm.IsImage t s) (e.IsImage s t)

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

                      theorem PartialHomeomorph.IsImage.compl {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] {e : PartialHomeomorph X Y} {s : Set X} {t : Set Y} (h : e.IsImage s t) :
                      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 : e.IsImage s t) (h' : e.IsImage 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 : e.IsImage s t) (h' : e.IsImage 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 : e.IsImage s t) (h' : e.IsImage s' 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 (Membership.mem s i)] [(i : Y) → Decidable (Membership.mem t i)] (h : e.IsImage s t) (h' : e'.IsImage s t) :
                      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 : e.IsImage s t) (h' : e'.IsImage s t) (hs : Eq (Inter.inter e.source s) (Inter.inter e'.source s)) (Heq : Set.EqOn e.toFun' e'.toFun' (Inter.inter 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 : e.IsImage s t) (hx : Membership.mem e.source x) :
                      theorem PartialHomeomorph.IsImage.closure {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] {e : PartialHomeomorph X Y} {s : Set X} {t : Set Y} (h : e.IsImage s t) :
                      theorem PartialHomeomorph.IsImage.interior {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] {e : PartialHomeomorph X Y} {s : Set X} {t : Set Y} (h : e.IsImage s t) :
                      theorem PartialHomeomorph.IsImage.frontier {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] {e : PartialHomeomorph X Y} {s : Set X} {t : Set Y} (h : e.IsImage s t) :
                      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 : e.IsImage s t) (hs : IsOpen (Inter.inter e.source s)) :

                      Restrict a PartialHomeomorph to a pair of corresponding open sets.

                      Equations
                      • Eq (h.restr hs) { toPartialEquiv := .restr, open_source := hs, open_target := , continuousOn_toFun := , continuousOn_invFun := }
                      Instances For
                        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 : e.IsImage s t) (hs : IsOpen (Inter.inter e.source s)) :

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

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

                        Equations
                        Instances For

                          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

                              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

                                The identity on the whole space as a partial homeomorphism.

                                Equations
                                Instances For

                                  const: PartialEquiv.const as a partial homeomorphism

                                  This is PartialEquiv.single as a partial homeomorphism: a constant map, whose source and target are necessarily singleton sets.

                                  Equations
                                  Instances For
                                    theorem PartialHomeomorph.const_apply {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] {a : X} {b : Y} (ha : IsOpen (Singleton.singleton a)) (hb : IsOpen (Singleton.singleton b)) (x : X) :
                                    Eq ((const ha hb).toFun' x) b

                                    ofSet: the identity on a set s

                                    The identity partial equivalence on a set s

                                    Equations
                                    Instances For
                                      theorem PartialHomeomorph.ofSet_apply {X : Type u_1} [TopologicalSpace X] (s : Set X) (hs : IsOpen s) :
                                      Eq (ofSet s hs).toFun' id
                                      theorem PartialHomeomorph.ofSet_source {X : Type u_1} [TopologicalSpace X] (s : Set X) (hs : IsOpen s) :
                                      Eq (ofSet s hs).source s
                                      theorem PartialHomeomorph.ofSet_target {X : Type u_1} [TopologicalSpace X] (s : Set X) (hs : IsOpen s) :
                                      Eq (ofSet s hs).target s
                                      theorem PartialHomeomorph.ofSet_symm {X : Type u_1} [TopologicalSpace X] {s : Set X} (hs : IsOpen s) :
                                      Eq (ofSet s hs).symm (ofSet s hs)

                                      trans: composition of two partial homeomorphisms

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

                                      Equations
                                      • Eq (e.trans' e' h) { toPartialEquiv := e.trans' e'.toPartialEquiv h, open_source := , open_target := , continuousOn_toFun := , continuousOn_invFun := }
                                      Instances For
                                        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 : Eq e.target e'.source) (a✝ : X) :
                                        Eq ((e.trans' e' h).toFun' a✝) (e'.toFun' (e.toFun' a✝))
                                        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 : Eq e.target e'.source) (a✝ : Z) :
                                        Eq ((e.trans' e' h).symm.toFun' a✝) (e.symm.toFun' (e'.symm.toFun' 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 : Eq e.target e'.source) :
                                        Eq (e.trans' e' h).target e'.target

                                        Composing two partial homeomorphisms, by restricting to the maximal domain where their composition is well defined. Within the Manifold namespace, there is the notation e ≫ₕ f for this.

                                        Equations
                                        Instances For
                                          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} :
                                          Eq ((e.trans e').toFun' x) (e'.toFun' (e.toFun' x))
                                          theorem PartialHomeomorph.trans_assoc {X : Type u_1} {Y : Type u_3} {Z : Type u_5} {Z' : Type u_6} [TopologicalSpace X] [TopologicalSpace Y] [TopologicalSpace Z] [TopologicalSpace Z'] (e : PartialHomeomorph X Y) (e' : PartialHomeomorph Y Z) (e'' : PartialHomeomorph Z Z') :
                                          Eq ((e.trans e').trans e'') (e.trans (e'.trans e''))
                                          theorem PartialHomeomorph.trans_ofSet {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] (e : PartialHomeomorph X Y) {s : Set Y} (hs : IsOpen s) :
                                          Eq (e.trans (ofSet s hs)) (e.restr (Set.preimage e.toFun' s))
                                          theorem PartialHomeomorph.trans_of_set' {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] (e : PartialHomeomorph X Y) {s : Set Y} (hs : IsOpen s) :
                                          theorem PartialHomeomorph.ofSet_trans {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] (e : PartialHomeomorph X Y) {s : Set X} (hs : IsOpen s) :
                                          Eq ((ofSet s hs).trans e) (e.restr s)
                                          theorem PartialHomeomorph.ofSet_trans' {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] (e : PartialHomeomorph X Y) {s : Set X} (hs : IsOpen s) :
                                          Eq ((ofSet s hs).trans e) (e.restr (Inter.inter e.source s))
                                          theorem PartialHomeomorph.ofSet_trans_ofSet {X : Type u_1} [TopologicalSpace X] {s : Set X} (hs : IsOpen s) {s' : Set X} (hs' : IsOpen s') :
                                          Eq ((ofSet s hs).trans (ofSet s' hs')) (ofSet (Inter.inter s s') )
                                          theorem PartialHomeomorph.restr_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) (s : Set X) :
                                          Eq ((e.restr s).trans e') ((e.trans e').restr s)

                                          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
                                            Instances For

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

                                              Two equivalent partial homeomorphisms have the same source.

                                              Two equivalent partial homeomorphisms have the same target.

                                              Two equivalent partial homeomorphisms have coinciding toFun on the source

                                              Two equivalent partial homeomorphisms have coinciding invFun on the target

                                              theorem PartialHomeomorph.EqOnSource.trans' {X : Type u_1} {Y : Type u_3} {Z : Type u_5} [TopologicalSpace X] [TopologicalSpace Y] [TopologicalSpace Z] {e e' : PartialHomeomorph X Y} {f f' : PartialHomeomorph Y Z} (he : HasEquiv.Equiv e e') (hf : HasEquiv.Equiv f f') :
                                              HasEquiv.Equiv (e.trans f) (e'.trans f')

                                              Composition of partial homeomorphisms respects equivalence.

                                              theorem PartialHomeomorph.EqOnSource.restr {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] {e e' : PartialHomeomorph X Y} (he : HasEquiv.Equiv e e') (s : Set X) :

                                              Restriction of partial homeomorphisms respects equivalence

                                              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

                                              product of two partial homeomorphisms

                                              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') :

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

                                              Equations
                                              • Eq (eX.prod eY) { toPartialEquiv := eX.prod eY.toPartialEquiv, open_source := , open_target := , continuousOn_toFun := , continuousOn_invFun := }
                                              Instances For
                                                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') :
                                                Eq (eX.prod eY).toFun' fun (p : Prod X Y) => { fst := eX.toFun' p.fst, snd := eY.toFun' p.snd }
                                                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 : Prod X' Y') :
                                                Eq ((eX.prod eY).symm.toFun' p) { fst := eX.symm.toFun' p.fst, snd := eY.symm.toFun' p.snd }
                                                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') :
                                                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') :
                                                theorem PartialHomeomorph.prod_symm {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') :
                                                Eq (eX.prod eY).symm (eX.symm.prod eY.symm)
                                                theorem PartialHomeomorph.prod_trans {X : Type u_1} {X' : Type u_2} {Y : Type u_3} {Y' : Type u_4} {Z : Type u_5} {Z' : Type u_6} [TopologicalSpace X] [TopologicalSpace X'] [TopologicalSpace Y] [TopologicalSpace Y'] [TopologicalSpace Z] [TopologicalSpace Z'] (e : PartialHomeomorph X Y) (f : PartialHomeomorph Y Z) (e' : PartialHomeomorph X' Y') (f' : PartialHomeomorph Y' Z') :
                                                Eq ((e.prod e').trans (f.prod f')) ((e.trans f).prod (e'.trans f'))
                                                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 eX' : PartialHomeomorph X X'} {eY eY' : PartialHomeomorph Y Y'} (h : (eX.prod eY).source.Nonempty) :
                                                Iff (Eq (eX.prod eY) (eX'.prod eY')) (And (Eq eX eX') (Eq eY eY'))
                                                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 eX' : PartialHomeomorph X X'} {eY eY' : PartialHomeomorph Y Y'} (h : (eX'.prod eY').source.Nonempty) :
                                                Iff (Eq (eX.prod eY) (eX'.prod eY')) (And (Eq eX eX') (Eq eY eY'))

                                                finite product of partial homeomorphisms

                                                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
                                                Instances For
                                                  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)) :
                                                  Eq (pi ei).toPartialEquiv (PartialEquiv.pi fun (i : ι) => (ei i).toPartialEquiv)
                                                  theorem PartialHomeomorph.pi_source {ι : 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)) :
                                                  Eq (pi ei).source (Set.univ.pi fun (i : ι) => (ei i).source)
                                                  theorem PartialHomeomorph.pi_symm_apply {ι : 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)) (a✝ : (i : ι) → Y i) (i : ι) :
                                                  Eq ((pi ei).symm.toFun' a✝ i) ((ei i).symm.toFun' (a✝ i))
                                                  theorem PartialHomeomorph.pi_apply {ι : 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)) (a✝ : (i : ι) → X i) (i : ι) :
                                                  Eq ((pi ei).toFun' a✝ i) ((ei i).toFun' (a✝ i))
                                                  theorem PartialHomeomorph.pi_target {ι : 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)) :
                                                  Eq (pi ei).target (Set.univ.pi fun (i : ι) => (ei i).target)

                                                  combining two partial homeomorphisms using Set.piecewise

                                                  def PartialHomeomorph.piecewise {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] (e e' : PartialHomeomorph X Y) (s : Set X) (t : Set Y) [(x : X) → Decidable (Membership.mem s x)] [(y : Y) → Decidable (Membership.mem t y)] (H : e.IsImage s t) (H' : e'.IsImage s t) (Hs : Eq (Inter.inter e.source (frontier s)) (Inter.inter e'.source (frontier s))) (Heq : Set.EqOn e.toFun' e'.toFun' (Inter.inter 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
                                                  • Eq (e.piecewise e' s t H H' Hs Heq) { toPartialEquiv := e.piecewise e'.toPartialEquiv s t H H', open_source := , open_target := , continuousOn_toFun := , continuousOn_invFun := }
                                                  Instances For
                                                    theorem PartialHomeomorph.piecewise_toPartialEquiv {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] (e e' : PartialHomeomorph X Y) (s : Set X) (t : Set Y) [(x : X) → Decidable (Membership.mem s x)] [(y : Y) → Decidable (Membership.mem t y)] (H : e.IsImage s t) (H' : e'.IsImage s t) (Hs : Eq (Inter.inter e.source (frontier s)) (Inter.inter e'.source (frontier s))) (Heq : Set.EqOn e.toFun' e'.toFun' (Inter.inter e.source (frontier s))) :
                                                    Eq (e.piecewise e' s t H H' Hs Heq).toPartialEquiv (e.piecewise e'.toPartialEquiv s t H H')
                                                    theorem PartialHomeomorph.piecewise_apply {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] (e e' : PartialHomeomorph X Y) (s : Set X) (t : Set Y) [(x : X) → Decidable (Membership.mem s x)] [(y : Y) → Decidable (Membership.mem t y)] (H : e.IsImage s t) (H' : e'.IsImage s t) (Hs : Eq (Inter.inter e.source (frontier s)) (Inter.inter e'.source (frontier s))) (Heq : Set.EqOn e.toFun' e'.toFun' (Inter.inter e.source (frontier s))) :
                                                    Eq (e.piecewise e' s t H H' Hs Heq).toFun' (s.piecewise e.toFun' e'.toFun')
                                                    theorem PartialHomeomorph.symm_piecewise {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] (e e' : PartialHomeomorph X Y) {s : Set X} {t : Set Y} [(x : X) → Decidable (Membership.mem s x)] [(y : Y) → Decidable (Membership.mem t y)] (H : e.IsImage s t) (H' : e'.IsImage s t) (Hs : Eq (Inter.inter e.source (frontier s)) (Inter.inter e'.source (frontier s))) (Heq : Set.EqOn e.toFun' e'.toFun' (Inter.inter e.source (frontier s))) :
                                                    Eq (e.piecewise e' s t H H' Hs Heq).symm (e.symm.piecewise e'.symm t s )

                                                    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
                                                    Instances For

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

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

                                                      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.

                                                      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

                                                      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

                                                      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.

                                                      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.

                                                      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

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

                                                        Equations
                                                        Instances For

                                                          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
                                                          Instances For
                                                            @[deprecated PartialHomeomorph.isOpenEmbedding_restrict (since := "2024-10-18")]

                                                            Alias of PartialHomeomorph.isOpenEmbedding_restrict.

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

                                                            @[deprecated PartialHomeomorph.to_isOpenEmbedding (since := "2024-10-18")]

                                                            Alias of PartialHomeomorph.to_isOpenEmbedding.


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

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

                                                            Equations
                                                            Instances For
                                                              noncomputable def Topology.IsOpenEmbedding.toPartialHomeomorph {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] (f : XY) (h : IsOpenEmbedding 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_isOpenEmbedding.

                                                              Equations
                                                              Instances For

                                                                inclusion of an open set in a topological space

                                                                noncomputable def TopologicalSpace.Opens.partialHomeomorphSubtypeCoe {X : Type u_1} [TopologicalSpace X] (s : Opens X) (hs : Nonempty (Subtype fun (x : X) => Membership.mem s x)) :

                                                                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

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

                                                                  Equations
                                                                  Instances For
                                                                    theorem PartialHomeomorph.trans_transHomeomorph {X : Type u_1} {Y : Type u_3} {Z : Type u_5} {Z' : Type u_6} [TopologicalSpace X] [TopologicalSpace Y] [TopologicalSpace Z] [TopologicalSpace Z'] (e : PartialHomeomorph X Y) (e' : PartialHomeomorph Y Z) (f'' : Homeomorph Z Z') :
                                                                    Eq ((e.trans e').transHomeomorph f'') (e.trans (e'.transHomeomorph f''))

                                                                    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 (Subtype fun (x : X) => Membership.mem s x)) :

                                                                    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

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

                                                                      noncomputable def PartialHomeomorph.lift_openEmbedding {X : Type u_7} {X' : Type u_8} {Z : Type u_9} [TopologicalSpace X] [TopologicalSpace X'] [TopologicalSpace Z] [Nonempty Z] {f : XX'} (e : PartialHomeomorph X Z) (hf : Topology.IsOpenEmbedding f) :

                                                                      Extend a partial homeomorphism e : X → Z to X' → Z, using an open embedding ι : X → X'. On ι(X), the extension is specified by e; its value elsewhere is arbitrary (and uninteresting).

                                                                      Equations
                                                                      • One or more equations did not get rendered due to their size.
                                                                      Instances For
                                                                        theorem PartialHomeomorph.lift_openEmbedding_apply {X : Type u_7} {X' : Type u_8} {Z : Type u_9} [TopologicalSpace X] [TopologicalSpace X'] [TopologicalSpace Z] [Nonempty Z] {f : XX'} (e : PartialHomeomorph X Z) (hf : Topology.IsOpenEmbedding f) {x : X} :
                                                                        Eq ((e.lift_openEmbedding hf).toFun' (f x)) (e.toFun' x)