Documentation

Mathlib.Topology.OpenPartialHomeomorph.Composition

Partial homeomorphisms: composition #

Main definitions #

Composition #

trans: composition of two open partial homeomorphisms

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

Equations
  • e.trans' e' h = { toPartialEquiv := e.trans' e'.toPartialEquiv h, open_source := , open_target := , continuousOn_toFun := , continuousOn_invFun := }
Instances For
    @[simp]
    theorem OpenPartialHomeomorph.trans'_symm_apply {X : Type u_1} {Y : Type u_3} {Z : Type u_5} [TopologicalSpace X] [TopologicalSpace Y] [TopologicalSpace Z] (e : OpenPartialHomeomorph X Y) (e' : OpenPartialHomeomorph Y Z) (h : e.target = e'.source) (a✝ : Z) :
    (e.trans' e' h).symm a✝ = e.symm (e'.symm a✝)
    @[simp]
    theorem OpenPartialHomeomorph.trans'_apply {X : Type u_1} {Y : Type u_3} {Z : Type u_5} [TopologicalSpace X] [TopologicalSpace Y] [TopologicalSpace Z] (e : OpenPartialHomeomorph X Y) (e' : OpenPartialHomeomorph Y Z) (h : e.target = e'.source) (a✝ : X) :
    (e.trans' e' h) a✝ = e' (e a✝)

    Composing two open 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
      @[simp]
      theorem OpenPartialHomeomorph.coe_trans {X : Type u_1} {Y : Type u_3} {Z : Type u_5} [TopologicalSpace X] [TopologicalSpace Y] [TopologicalSpace Z] (e : OpenPartialHomeomorph X Y) (e' : OpenPartialHomeomorph Y Z) :
      (e.trans e') = e' e
      @[simp]
      theorem OpenPartialHomeomorph.coe_trans_symm {X : Type u_1} {Y : Type u_3} {Z : Type u_5} [TopologicalSpace X] [TopologicalSpace Y] [TopologicalSpace Z] (e : OpenPartialHomeomorph X Y) (e' : OpenPartialHomeomorph Y Z) :
      (e.trans e').symm = e.symm e'.symm
      theorem OpenPartialHomeomorph.trans_apply {X : Type u_1} {Y : Type u_3} {Z : Type u_5} [TopologicalSpace X] [TopologicalSpace Y] [TopologicalSpace Z] (e : OpenPartialHomeomorph X Y) (e' : OpenPartialHomeomorph Y Z) {x : X} :
      (e.trans e') x = e' (e x)
      theorem OpenPartialHomeomorph.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 : OpenPartialHomeomorph X Y) (e' : OpenPartialHomeomorph Y Z) (e'' : OpenPartialHomeomorph Z Z') :
      (e.trans e').trans e'' = e.trans (e'.trans e'')
      theorem OpenPartialHomeomorph.trans_ofSet {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] (e : OpenPartialHomeomorph X Y) {s : Set Y} (hs : IsOpen s) :
      e.trans (ofSet s hs) = e.restr (e ⁻¹' s)
      theorem OpenPartialHomeomorph.trans_of_set' {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] (e : OpenPartialHomeomorph X Y) {s : Set Y} (hs : IsOpen s) :
      e.trans (ofSet s hs) = e.restr (e.source e ⁻¹' s)
      theorem OpenPartialHomeomorph.ofSet_trans {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] (e : OpenPartialHomeomorph X Y) {s : Set X} (hs : IsOpen s) :
      (ofSet s hs).trans e = e.restr s
      theorem OpenPartialHomeomorph.ofSet_trans' {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] (e : OpenPartialHomeomorph X Y) {s : Set X} (hs : IsOpen s) :
      (ofSet s hs).trans e = e.restr (e.source s)
      @[simp]
      theorem OpenPartialHomeomorph.ofSet_trans_ofSet {X : Type u_1} [TopologicalSpace X] {s : Set X} (hs : IsOpen s) {s' : Set X} (hs' : IsOpen s') :
      (ofSet s hs).trans (ofSet s' hs') = ofSet (s s')
      theorem OpenPartialHomeomorph.restr_trans {X : Type u_1} {Y : Type u_3} {Z : Type u_5} [TopologicalSpace X] [TopologicalSpace Y] [TopologicalSpace Z] (e : OpenPartialHomeomorph X Y) (e' : OpenPartialHomeomorph Y Z) (s : Set X) :
      (e.restr s).trans e' = (e.trans e').restr s
      theorem OpenPartialHomeomorph.EqOnSource.trans' {X : Type u_1} {Y : Type u_3} {Z : Type u_5} [TopologicalSpace X] [TopologicalSpace Y] [TopologicalSpace Z] {e e' : OpenPartialHomeomorph X Y} {f f' : OpenPartialHomeomorph Y Z} (he : e e') (hf : f f') :
      e.trans f e'.trans f'

      Composition of open partial homeomorphisms respects equivalence.

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

      theorem OpenPartialHomeomorph.restr_symm_trans {X : Type u_1} {Y : Type u_3} [TopologicalSpace X] [TopologicalSpace Y] (e : OpenPartialHomeomorph X Y) {s : Set X} {e' : OpenPartialHomeomorph X Y} (hs : IsOpen s) (hs' : IsOpen (e '' s)) (hs'' : s e.source) :
      (e.restr s).symm.trans e' (e.symm.trans e').restr (e '' s)
      @[deprecated Homeomorph.trans_toOpenPartialHomeomorph (since := "2025-08-29")]

      Alias of Homeomorph.trans_toOpenPartialHomeomorph.

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

      Equations
      Instances For
        @[simp]
        @[deprecated Homeomorph.transOpenPartialHomeomorph (since := "2025-08-29")]

        Alias of Homeomorph.transOpenPartialHomeomorph.


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

        Equations
        Instances For
          @[deprecated Homeomorph.transOpenPartialHomeomorph_eq_trans (since := "2025-08-29")]

          Alias of Homeomorph.transOpenPartialHomeomorph_eq_trans.

          @[deprecated Homeomorph.transOpenPartialHomeomorph_trans (since := "2025-08-29")]

          Alias of Homeomorph.transOpenPartialHomeomorph_trans.

          @[deprecated Homeomorph.trans_transOpenPartialHomeomorph (since := "2025-08-29")]

          Alias of Homeomorph.trans_transOpenPartialHomeomorph.