Documentation

Mathlib.Geometry.Manifold.IsManifold

C^n manifolds (possibly with boundary or corners) #

A C^n manifold is a manifold modelled on a normed vector space, or a subset like a half-space (to get manifolds with boundaries) for which the changes of coordinates are C^n maps. We define a model with corners as a map I : H β†’ E embedding nicely the topological space H in the vector space E (or more precisely as a structure containing all the relevant properties). Given such a model with corners I on (E, H), we define the groupoid of local homeomorphisms of H which are C^n when read in E (for any regularity n : WithTop β„•βˆž). With this groupoid at hand and the general machinery of charted spaces, we thus get the notion of C^n manifold with respect to any model with corners I on (E, H).

Some texts assume manifolds to be Hausdorff and second countable. We (in mathlib) assume neither, but add these assumptions later as needed. (Quite a few results still do not require them.)

Main definitions #

As specific examples of models with corners, we define (in Geometry.Manifold.Instances.Real)

With these definitions at hand, to invoke an n-dimensional C^∞ real manifold without boundary, one could use

variable {n : β„•} {M : Type*} [TopologicalSpace M] [ChartedSpace (EuclideanSpace ℝ (Fin n)) M] [IsManifold (𝓑 n) ∞ M].

However, this is not the recommended way: a theorem proved using this assumption would not apply for instance to the tangent space of such a manifold, which is modelled on (EuclideanSpace ℝ (Fin n)) Γ— (EuclideanSpace ℝ (Fin n)) and not on EuclideanSpace ℝ (Fin (2 * n))! In the same way, it would not apply to product manifolds, modelled on (EuclideanSpace ℝ (Fin n)) Γ— (EuclideanSpace ℝ (Fin m)). The right invocation does not focus on one specific construction, but on all constructions sharing the right properties, like

variable {E : Type*} [NormedAddCommGroup E] [NormedSpace ℝ E] [FiniteDimensional ℝ E] {I : ModelWithCorners ℝ E E} [I.Boundaryless] {M : Type*} [TopologicalSpace M] [ChartedSpace E M] [IsManifold I ∞ M]

Here, I.Boundaryless is a typeclass property ensuring that there is no boundary (this is for instance the case for modelWithCornersSelf, or products of these). Note that one could consider as a natural assumption to only use the trivial model with corners modelWithCornersSelf ℝ E, but again in product manifolds the natural model with corners will not be this one but the product one (and they are not defeq as (fun p : E Γ— F ↦ (p.1, p.2)) is not defeq to the identity). So, it is important to use the above incantation to maximize the applicability of theorems.

Even better, if the result should apply in a parallel way to smooth manifolds and to analytic manifolds, the last typeclass should be replaced with [IsManifold I n M] for n : WithTop β„•βˆž.

We also define TangentSpace I (x : M) as a type synonym of E, and TangentBundle I M as a type synonym for Ξ  (x : M), TangentSpace I x (in the form of an abbrev of Bundle.TotalSpace E (TangentSpace I : M β†’ Type _)). Apart from basic typeclasses on TangentSpace I x, nothing is proved about them in this file, but it is useful to have them available as definitions early on to get a clean import structure below. The smooth bundle structure is defined in VectorBundle.Tangent, while the definition is used to talk about manifold derivatives in MFDeriv.Basic, and neither file needs import the other.

Implementation notes #

We want to talk about manifolds modelled on a vector space, but also on manifolds with boundary, modelled on a half space (or even manifolds with corners). For the latter examples, we still want to define smooth functions, tangent bundles, and so on. As smooth functions are well defined on vector spaces or subsets of these, one could take for model space a subtype of a vector space. With the drawback that the whole vector space itself (which is the most basic example) is not directly a subtype of itself: the inclusion of univ : Set E in Set E would show up in the definition, instead of id.

A good abstraction covering both cases it to have a vector space E (with basic example the Euclidean space), a model space H (with basic example the upper half space), and an embedding of H into E (which can be the identity for H = E, or Subtype.val for manifolds with corners). We say that the pair (E, H) with their embedding is a model with corners, and we encompass all the relevant properties (in particular the fact that the image of H in E should have unique differentials) in the definition of ModelWithCorners.

I have considered using the model with corners I as a typeclass argument, possibly outParam, to get lighter notations later on, but it did not turn out right, as on E Γ— F there are two natural model with corners, the trivial (identity) one, and the product one, and they are not defeq and one needs to indicate to Lean which one we want to use. This means that when talking on objects on manifolds one will most often need to specify the model with corners one is using. For instance, the tangent bundle will be TangentBundle I M and the derivative will be mfderiv I I' f, instead of the more natural notations TangentBundle π•œ M and mfderiv π•œ f (the field has to be explicit anyway, as some manifolds could be considered both as real and complex manifolds).

Models with corners. #

structure ModelWithCorners (π•œ : Type u_1) [NontriviallyNormedField π•œ] (E : Type u_2) [NormedAddCommGroup E] [NormedSpace π•œ E] (H : Type u_3) [TopologicalSpace H] extends PartialEquiv H E :
Type (max u_2 u_3)

A structure containing information on the way a space H embeds in a model vector space E over the field π•œ. This is all what is needed to define a C^n manifold with model space H, and model vector space E.

We require two conditions uniqueDiffOn' and target_subset_closure_interior, which are satisfied in the relevant cases (where range I = univ or a half space or a quadrant) and useful for technical reasons. The former makes sure that manifold derivatives are uniquely defined, the latter ensures that for C^2 maps the second derivatives are symmetric even for points on the boundary, as these are limit points of interior points where symmetry holds. If further conditions turn out to be useful, they can be added here.

Instances For
    theorem ModelWithCorners.ext {π•œ : Type u_1} {inst✝ : NontriviallyNormedField π•œ} {E : Type u_2} {inst✝¹ : NormedAddCommGroup E} {inst✝² : NormedSpace π•œ E} {H : Type u_3} {inst✝³ : TopologicalSpace H} {x y : ModelWithCorners π•œ E H} (toFun : ↑x.toPartialEquiv = ↑y.toPartialEquiv) (invFun : x.invFun = y.invFun) (source : x.source = y.source) (target : x.target = y.target) :
    x = y
    def modelWithCornersSelf (π•œ : Type u_1) [NontriviallyNormedField π•œ] (E : Type u_2) [NormedAddCommGroup E] [NormedSpace π•œ E] :
    ModelWithCorners π•œ E E

    A vector space is a model with corners.

    Equations
    • modelWithCornersSelf π•œ E = { toPartialEquiv := PartialEquiv.refl E, source_eq := β‹―, uniqueDiffOn' := β‹―, target_subset_closure_interior := β‹―, continuous_toFun := β‹―, continuous_invFun := β‹― }
    Instances For

      A vector space is a model with corners.

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

        A normed field is a model with corners.

        Equations
        • One or more equations did not get rendered due to their size.
        Instances For
          def ModelWithCorners.toFun' {π•œ : Type u_1} [NontriviallyNormedField π•œ] {E : Type u_2} [NormedAddCommGroup E] [NormedSpace π•œ E] {H : Type u_3} [TopologicalSpace H] (e : ModelWithCorners π•œ E H) :
          H β†’ E

          Coercion of a model with corners 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
            instance ModelWithCorners.instCoeFunForall {π•œ : Type u_1} [NontriviallyNormedField π•œ] {E : Type u_2} [NormedAddCommGroup E] [NormedSpace π•œ E] {H : Type u_3} [TopologicalSpace H] :
            CoeFun (ModelWithCorners π•œ E H) fun (x : ModelWithCorners π•œ E H) => H β†’ E
            Equations
            def ModelWithCorners.symm {π•œ : Type u_1} [NontriviallyNormedField π•œ] {E : Type u_2} [NormedAddCommGroup E] [NormedSpace π•œ E] {H : Type u_3} [TopologicalSpace H] (I : ModelWithCorners π•œ E H) :

            The inverse to a model with corners, only registered as a PartialEquiv.

            Equations
            • I.symm = I.symm
            Instances For
              def ModelWithCorners.Simps.apply (π•œ : Type u_4) [NontriviallyNormedField π•œ] (E : Type u_5) [NormedAddCommGroup E] [NormedSpace π•œ E] (H : Type u_6) [TopologicalSpace H] (I : ModelWithCorners π•œ E H) :
              H β†’ E

              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
                def ModelWithCorners.Simps.symm_apply (π•œ : Type u_4) [NontriviallyNormedField π•œ] (E : Type u_5) [NormedAddCommGroup E] [NormedSpace π•œ E] (H : Type u_6) [TopologicalSpace H] (I : ModelWithCorners π•œ E H) :
                E β†’ H

                See Note [custom simps projection]

                Equations
                Instances For
                  @[simp]
                  theorem ModelWithCorners.toPartialEquiv_coe {π•œ : Type u_1} [NontriviallyNormedField π•œ] {E : Type u_2} [NormedAddCommGroup E] [NormedSpace π•œ E] {H : Type u_3} [TopologicalSpace H] (I : ModelWithCorners π•œ E H) :
                  ↑I.toPartialEquiv = ↑I
                  @[simp]
                  theorem ModelWithCorners.mk_coe {π•œ : Type u_1} [NontriviallyNormedField π•œ] {E : Type u_2} [NormedAddCommGroup E] [NormedSpace π•œ E] {H : Type u_3} [TopologicalSpace H] (e : PartialEquiv H E) (a : e.source = Set.univ) (b : UniqueDiffOn π•œ e.target) (c : e.target βŠ† closure (interior e.target)) (d : Continuous ↑e) (d' : Continuous e.invFun) :
                  ↑{ toPartialEquiv := e, source_eq := a, uniqueDiffOn' := b, target_subset_closure_interior := c, continuous_toFun := d, continuous_invFun := d' } = ↑e
                  @[simp]
                  theorem ModelWithCorners.toPartialEquiv_coe_symm {π•œ : Type u_1} [NontriviallyNormedField π•œ] {E : Type u_2} [NormedAddCommGroup E] [NormedSpace π•œ E] {H : Type u_3} [TopologicalSpace H] (I : ModelWithCorners π•œ E H) :
                  ↑I.symm = ↑I.symm
                  @[simp]
                  theorem ModelWithCorners.mk_symm {π•œ : Type u_1} [NontriviallyNormedField π•œ] {E : Type u_2} [NormedAddCommGroup E] [NormedSpace π•œ E] {H : Type u_3} [TopologicalSpace H] (e : PartialEquiv H E) (a : e.source = Set.univ) (b : UniqueDiffOn π•œ e.target) (c : e.target βŠ† closure (interior e.target)) (d : Continuous ↑e) (d' : Continuous e.invFun) :
                  { toPartialEquiv := e, source_eq := a, uniqueDiffOn' := b, target_subset_closure_interior := c, continuous_toFun := d, continuous_invFun := d' }.symm = e.symm
                  theorem ModelWithCorners.continuous {π•œ : Type u_1} [NontriviallyNormedField π•œ] {E : Type u_2} [NormedAddCommGroup E] [NormedSpace π•œ E] {H : Type u_3} [TopologicalSpace H] (I : ModelWithCorners π•œ E H) :
                  Continuous ↑I
                  theorem ModelWithCorners.continuousAt {π•œ : Type u_1} [NontriviallyNormedField π•œ] {E : Type u_2} [NormedAddCommGroup E] [NormedSpace π•œ E] {H : Type u_3} [TopologicalSpace H] (I : ModelWithCorners π•œ E H) {x : H} :
                  ContinuousAt (↑I) x
                  theorem ModelWithCorners.continuousWithinAt {π•œ : Type u_1} [NontriviallyNormedField π•œ] {E : Type u_2} [NormedAddCommGroup E] [NormedSpace π•œ E] {H : Type u_3} [TopologicalSpace H] (I : ModelWithCorners π•œ E H) {s : Set H} {x : H} :
                  ContinuousWithinAt (↑I) s x
                  theorem ModelWithCorners.continuous_symm {π•œ : Type u_1} [NontriviallyNormedField π•œ] {E : Type u_2} [NormedAddCommGroup E] [NormedSpace π•œ E] {H : Type u_3} [TopologicalSpace H] (I : ModelWithCorners π•œ E H) :
                  Continuous ↑I.symm
                  theorem ModelWithCorners.continuousAt_symm {π•œ : Type u_1} [NontriviallyNormedField π•œ] {E : Type u_2} [NormedAddCommGroup E] [NormedSpace π•œ E] {H : Type u_3} [TopologicalSpace H] (I : ModelWithCorners π•œ E H) {x : E} :
                  ContinuousAt (↑I.symm) x
                  theorem ModelWithCorners.continuousWithinAt_symm {π•œ : Type u_1} [NontriviallyNormedField π•œ] {E : Type u_2} [NormedAddCommGroup E] [NormedSpace π•œ E] {H : Type u_3} [TopologicalSpace H] (I : ModelWithCorners π•œ E H) {s : Set E} {x : E} :
                  ContinuousWithinAt (↑I.symm) s x
                  theorem ModelWithCorners.continuousOn_symm {π•œ : Type u_1} [NontriviallyNormedField π•œ] {E : Type u_2} [NormedAddCommGroup E] [NormedSpace π•œ E] {H : Type u_3} [TopologicalSpace H] (I : ModelWithCorners π•œ E H) {s : Set E} :
                  ContinuousOn (↑I.symm) s
                  @[simp]
                  theorem ModelWithCorners.target_eq {π•œ : Type u_1} [NontriviallyNormedField π•œ] {E : Type u_2} [NormedAddCommGroup E] [NormedSpace π•œ E] {H : Type u_3} [TopologicalSpace H] (I : ModelWithCorners π•œ E H) :
                  I.target = Set.range ↑I
                  theorem ModelWithCorners.uniqueDiffOn {π•œ : Type u_1} [NontriviallyNormedField π•œ] {E : Type u_2} [NormedAddCommGroup E] [NormedSpace π•œ E] {H : Type u_3} [TopologicalSpace H] (I : ModelWithCorners π•œ E H) :
                  UniqueDiffOn π•œ (Set.range ↑I)
                  @[deprecated ModelWithCorners.uniqueDiffOn (since := "2024-09-30")]
                  theorem ModelWithCorners.unique_diff {π•œ : Type u_1} [NontriviallyNormedField π•œ] {E : Type u_2} [NormedAddCommGroup E] [NormedSpace π•œ E] {H : Type u_3} [TopologicalSpace H] (I : ModelWithCorners π•œ E H) :
                  UniqueDiffOn π•œ (Set.range ↑I)

                  Alias of ModelWithCorners.uniqueDiffOn.

                  theorem ModelWithCorners.range_subset_closure_interior {π•œ : Type u_1} [NontriviallyNormedField π•œ] {E : Type u_2} [NormedAddCommGroup E] [NormedSpace π•œ E] {H : Type u_3} [TopologicalSpace H] (I : ModelWithCorners π•œ E H) :
                  @[simp]
                  theorem ModelWithCorners.left_inv {π•œ : Type u_1} [NontriviallyNormedField π•œ] {E : Type u_2} [NormedAddCommGroup E] [NormedSpace π•œ E] {H : Type u_3} [TopologicalSpace H] (I : ModelWithCorners π•œ E H) (x : H) :
                  ↑I.symm (↑I x) = x
                  theorem ModelWithCorners.leftInverse {π•œ : Type u_1} [NontriviallyNormedField π•œ] {E : Type u_2} [NormedAddCommGroup E] [NormedSpace π•œ E] {H : Type u_3} [TopologicalSpace H] (I : ModelWithCorners π•œ E H) :
                  Function.LeftInverse ↑I.symm ↑I
                  theorem ModelWithCorners.injective {π•œ : Type u_1} [NontriviallyNormedField π•œ] {E : Type u_2} [NormedAddCommGroup E] [NormedSpace π•œ E] {H : Type u_3} [TopologicalSpace H] (I : ModelWithCorners π•œ E H) :
                  @[simp]
                  theorem ModelWithCorners.symm_comp_self {π•œ : Type u_1} [NontriviallyNormedField π•œ] {E : Type u_2} [NormedAddCommGroup E] [NormedSpace π•œ E] {H : Type u_3} [TopologicalSpace H] (I : ModelWithCorners π•œ E H) :
                  ↑I.symm ∘ ↑I = id
                  theorem ModelWithCorners.rightInvOn {π•œ : Type u_1} [NontriviallyNormedField π•œ] {E : Type u_2} [NormedAddCommGroup E] [NormedSpace π•œ E] {H : Type u_3} [TopologicalSpace H] (I : ModelWithCorners π•œ E H) :
                  Set.RightInvOn (↑I.symm) (↑I) (Set.range ↑I)
                  @[simp]
                  theorem ModelWithCorners.right_inv {π•œ : Type u_1} [NontriviallyNormedField π•œ] {E : Type u_2} [NormedAddCommGroup E] [NormedSpace π•œ E] {H : Type u_3} [TopologicalSpace H] (I : ModelWithCorners π•œ E H) {x : E} (hx : x ∈ Set.range ↑I) :
                  ↑I (↑I.symm x) = x
                  theorem ModelWithCorners.preimage_image {π•œ : Type u_1} [NontriviallyNormedField π•œ] {E : Type u_2} [NormedAddCommGroup E] [NormedSpace π•œ E] {H : Type u_3} [TopologicalSpace H] (I : ModelWithCorners π•œ E H) (s : Set H) :
                  ↑I ⁻¹' (↑I '' s) = s
                  theorem ModelWithCorners.image_eq {π•œ : Type u_1} [NontriviallyNormedField π•œ] {E : Type u_2} [NormedAddCommGroup E] [NormedSpace π•œ E] {H : Type u_3} [TopologicalSpace H] (I : ModelWithCorners π•œ E H) (s : Set H) :
                  ↑I '' s = ↑I.symm ⁻¹' s ∩ Set.range ↑I
                  theorem ModelWithCorners.isClosedEmbedding {π•œ : Type u_1} [NontriviallyNormedField π•œ] {E : Type u_2} [NormedAddCommGroup E] [NormedSpace π•œ E] {H : Type u_3} [TopologicalSpace H] (I : ModelWithCorners π•œ E H) :
                  @[deprecated ModelWithCorners.isClosedEmbedding (since := "2024-10-20")]
                  theorem ModelWithCorners.closedEmbedding {π•œ : Type u_1} [NontriviallyNormedField π•œ] {E : Type u_2} [NormedAddCommGroup E] [NormedSpace π•œ E] {H : Type u_3} [TopologicalSpace H] (I : ModelWithCorners π•œ E H) :

                  Alias of ModelWithCorners.isClosedEmbedding.

                  theorem ModelWithCorners.isClosed_range {π•œ : Type u_1} [NontriviallyNormedField π•œ] {E : Type u_2} [NormedAddCommGroup E] [NormedSpace π•œ E] {H : Type u_3} [TopologicalSpace H] (I : ModelWithCorners π•œ E H) :
                  theorem ModelWithCorners.range_eq_closure_interior {π•œ : Type u_1} [NontriviallyNormedField π•œ] {E : Type u_2} [NormedAddCommGroup E] [NormedSpace π•œ E] {H : Type u_3} [TopologicalSpace H] (I : ModelWithCorners π•œ E H) :
                  Set.range ↑I = closure (interior (Set.range ↑I))
                  theorem ModelWithCorners.map_nhds_eq {π•œ : Type u_1} [NontriviallyNormedField π•œ] {E : Type u_2} [NormedAddCommGroup E] [NormedSpace π•œ E] {H : Type u_3} [TopologicalSpace H] (I : ModelWithCorners π•œ E H) (x : H) :
                  Filter.map (↑I) (nhds x) = nhdsWithin (↑I x) (Set.range ↑I)
                  theorem ModelWithCorners.map_nhdsWithin_eq {π•œ : Type u_1} [NontriviallyNormedField π•œ] {E : Type u_2} [NormedAddCommGroup E] [NormedSpace π•œ E] {H : Type u_3} [TopologicalSpace H] (I : ModelWithCorners π•œ E H) (s : Set H) (x : H) :
                  Filter.map (↑I) (nhdsWithin x s) = nhdsWithin (↑I x) (↑I '' s)
                  theorem ModelWithCorners.image_mem_nhdsWithin {π•œ : Type u_1} [NontriviallyNormedField π•œ] {E : Type u_2} [NormedAddCommGroup E] [NormedSpace π•œ E] {H : Type u_3} [TopologicalSpace H] (I : ModelWithCorners π•œ E H) {x : H} {s : Set H} (hs : s ∈ nhds x) :
                  ↑I '' s ∈ nhdsWithin (↑I x) (Set.range ↑I)
                  theorem ModelWithCorners.symm_map_nhdsWithin_image {π•œ : Type u_1} [NontriviallyNormedField π•œ] {E : Type u_2} [NormedAddCommGroup E] [NormedSpace π•œ E] {H : Type u_3} [TopologicalSpace H] (I : ModelWithCorners π•œ E H) {x : H} {s : Set H} :
                  Filter.map (↑I.symm) (nhdsWithin (↑I x) (↑I '' s)) = nhdsWithin x s
                  theorem ModelWithCorners.symm_map_nhdsWithin_range {π•œ : Type u_1} [NontriviallyNormedField π•œ] {E : Type u_2} [NormedAddCommGroup E] [NormedSpace π•œ E] {H : Type u_3} [TopologicalSpace H] (I : ModelWithCorners π•œ E H) (x : H) :
                  Filter.map (↑I.symm) (nhdsWithin (↑I x) (Set.range ↑I)) = nhds x
                  theorem ModelWithCorners.uniqueDiffOn_preimage {π•œ : Type u_1} [NontriviallyNormedField π•œ] {E : Type u_2} [NormedAddCommGroup E] [NormedSpace π•œ E] {H : Type u_3} [TopologicalSpace H] (I : ModelWithCorners π•œ E H) {s : Set H} (hs : IsOpen s) :
                  UniqueDiffOn π•œ (↑I.symm ⁻¹' s ∩ Set.range ↑I)
                  @[deprecated ModelWithCorners.uniqueDiffOn_preimage (since := "2024-09-30")]
                  theorem ModelWithCorners.unique_diff_preimage {π•œ : Type u_1} [NontriviallyNormedField π•œ] {E : Type u_2} [NormedAddCommGroup E] [NormedSpace π•œ E] {H : Type u_3} [TopologicalSpace H] (I : ModelWithCorners π•œ E H) {s : Set H} (hs : IsOpen s) :
                  UniqueDiffOn π•œ (↑I.symm ⁻¹' s ∩ Set.range ↑I)

                  Alias of ModelWithCorners.uniqueDiffOn_preimage.

                  theorem ModelWithCorners.uniqueDiffOn_preimage_source {π•œ : Type u_1} [NontriviallyNormedField π•œ] {E : Type u_2} [NormedAddCommGroup E] [NormedSpace π•œ E] {H : Type u_3} [TopologicalSpace H] (I : ModelWithCorners π•œ E H) {Ξ² : Type u_4} [TopologicalSpace Ξ²] {e : PartialHomeomorph H Ξ²} :
                  UniqueDiffOn π•œ (↑I.symm ⁻¹' e.source ∩ Set.range ↑I)
                  @[deprecated ModelWithCorners.uniqueDiffOn_preimage_source (since := "2024-09-30")]
                  theorem ModelWithCorners.unique_diff_preimage_source {π•œ : Type u_1} [NontriviallyNormedField π•œ] {E : Type u_2} [NormedAddCommGroup E] [NormedSpace π•œ E] {H : Type u_3} [TopologicalSpace H] (I : ModelWithCorners π•œ E H) {Ξ² : Type u_4} [TopologicalSpace Ξ²] {e : PartialHomeomorph H Ξ²} :
                  UniqueDiffOn π•œ (↑I.symm ⁻¹' e.source ∩ Set.range ↑I)

                  Alias of ModelWithCorners.uniqueDiffOn_preimage_source.

                  theorem ModelWithCorners.uniqueDiffWithinAt_image {π•œ : Type u_1} [NontriviallyNormedField π•œ] {E : Type u_2} [NormedAddCommGroup E] [NormedSpace π•œ E] {H : Type u_3} [TopologicalSpace H] (I : ModelWithCorners π•œ E H) {x : H} :
                  UniqueDiffWithinAt π•œ (Set.range ↑I) (↑I x)
                  @[deprecated ModelWithCorners.uniqueDiffWithinAt_image (since := "2024-09-30")]
                  theorem ModelWithCorners.unique_diff_at_image {π•œ : Type u_1} [NontriviallyNormedField π•œ] {E : Type u_2} [NormedAddCommGroup E] [NormedSpace π•œ E] {H : Type u_3} [TopologicalSpace H] (I : ModelWithCorners π•œ E H) {x : H} :
                  UniqueDiffWithinAt π•œ (Set.range ↑I) (↑I x)

                  Alias of ModelWithCorners.uniqueDiffWithinAt_image.

                  theorem ModelWithCorners.symm_continuousWithinAt_comp_right_iff {π•œ : Type u_1} [NontriviallyNormedField π•œ] {E : Type u_2} [NormedAddCommGroup E] [NormedSpace π•œ E] {H : Type u_3} [TopologicalSpace H] (I : ModelWithCorners π•œ E H) {X : Type u_4} [TopologicalSpace X] {f : H β†’ X} {s : Set H} {x : H} :
                  ContinuousWithinAt (f ∘ ↑I.symm) (↑I.symm ⁻¹' s ∩ Set.range ↑I) (↑I x) ↔ ContinuousWithinAt f s x
                  theorem ModelWithCorners.t1Space {π•œ : Type u_1} [NontriviallyNormedField π•œ] {E : Type u_2} [NormedAddCommGroup E] [NormedSpace π•œ E] {H : Type u_3} [TopologicalSpace H] (I : ModelWithCorners π•œ E H) (M : Type u_4) [TopologicalSpace M] [ChartedSpace H M] :

                  Every manifold is a FrΓ©chet space (T1 space) -- regardless of whether it is Hausdorff.

                  @[simp]
                  theorem modelWithCornersSelf_partialEquiv (π•œ : Type u_1) [NontriviallyNormedField π•œ] (E : Type u_2) [NormedAddCommGroup E] [NormedSpace π•œ E] :
                  (modelWithCornersSelf π•œ E).toPartialEquiv = PartialEquiv.refl E

                  In the trivial model with corners, the associated PartialEquiv is the identity.

                  @[simp]
                  theorem modelWithCornersSelf_coe (π•œ : Type u_1) [NontriviallyNormedField π•œ] (E : Type u_2) [NormedAddCommGroup E] [NormedSpace π•œ E] :
                  ↑(modelWithCornersSelf π•œ E) = id
                  @[simp]
                  theorem modelWithCornersSelf_coe_symm (π•œ : Type u_1) [NontriviallyNormedField π•œ] (E : Type u_2) [NormedAddCommGroup E] [NormedSpace π•œ E] :
                  ↑(modelWithCornersSelf π•œ E).symm = id
                  def ModelWithCorners.prod {π•œ : Type u} [NontriviallyNormedField π•œ] {E : Type v} [NormedAddCommGroup E] [NormedSpace π•œ E] {H : Type w} [TopologicalSpace H] (I : ModelWithCorners π•œ E H) {E' : Type v'} [NormedAddCommGroup E'] [NormedSpace π•œ E'] {H' : Type w'} [TopologicalSpace H'] (I' : ModelWithCorners π•œ E' H') :
                  ModelWithCorners π•œ (E Γ— E') (ModelProd H H')

                  Given two model_with_corners I on (E, H) and I' on (E', H'), we define the model with corners I.prod I' on (E Γ— E', ModelProd H H'). This appears in particular for the manifold structure on the tangent bundle to a manifold modelled on (E, H): it will be modelled on (E Γ— E, H Γ— E). See note [Manifold type tags] for explanation about ModelProd H H' vs H Γ— H'.

                  Equations
                  • One or more equations did not get rendered due to their size.
                  Instances For
                    theorem ModelWithCorners.prod_source {π•œ : Type u} [NontriviallyNormedField π•œ] {E : Type v} [NormedAddCommGroup E] [NormedSpace π•œ E] {H : Type w} [TopologicalSpace H] (I : ModelWithCorners π•œ E H) {E' : Type v'} [NormedAddCommGroup E'] [NormedSpace π•œ E'] {H' : Type w'} [TopologicalSpace H'] (I' : ModelWithCorners π•œ E' H') :
                    (I.prod I').source = {x : ModelProd H H' | x.1 ∈ I.source ∧ x.2 ∈ I'.source}
                    theorem ModelWithCorners.prod_target {π•œ : Type u} [NontriviallyNormedField π•œ] {E : Type v} [NormedAddCommGroup E] [NormedSpace π•œ E] {H : Type w} [TopologicalSpace H] (I : ModelWithCorners π•œ E H) {E' : Type v'} [NormedAddCommGroup E'] [NormedSpace π•œ E'] {H' : Type w'} [TopologicalSpace H'] (I' : ModelWithCorners π•œ E' H') :
                    (I.prod I').target = (I.prod I'.toPartialEquiv).target
                    theorem ModelWithCorners.prod_symm_apply {π•œ : Type u} [NontriviallyNormedField π•œ] {E : Type v} [NormedAddCommGroup E] [NormedSpace π•œ E] {H : Type w} [TopologicalSpace H] (I : ModelWithCorners π•œ E H) {E' : Type v'} [NormedAddCommGroup E'] [NormedSpace π•œ E'] {H' : Type w'} [TopologicalSpace H'] (I' : ModelWithCorners π•œ E' H') (x : E Γ— E') :
                    ↑(I.prod I').symm x = (↑I.symm x.1, ↑I'.symm x.2)
                    theorem ModelWithCorners.prod_apply {π•œ : Type u} [NontriviallyNormedField π•œ] {E : Type v} [NormedAddCommGroup E] [NormedSpace π•œ E] {H : Type w} [TopologicalSpace H] (I : ModelWithCorners π•œ E H) {E' : Type v'} [NormedAddCommGroup E'] [NormedSpace π•œ E'] {H' : Type w'} [TopologicalSpace H'] (I' : ModelWithCorners π•œ E' H') (x : ModelProd H H') :
                    ↑(I.prod I') x = (↑I x.1, ↑I' x.2)
                    def ModelWithCorners.pi {π•œ : Type u} [NontriviallyNormedField π•œ] {ΞΉ : Type v} [Fintype ΞΉ] {E : ΞΉ β†’ Type w} [(i : ΞΉ) β†’ NormedAddCommGroup (E i)] [(i : ΞΉ) β†’ NormedSpace π•œ (E i)] {H : ΞΉ β†’ Type u'} [(i : ΞΉ) β†’ TopologicalSpace (H i)] (I : (i : ΞΉ) β†’ ModelWithCorners π•œ (E i) (H i)) :
                    ModelWithCorners π•œ ((i : ΞΉ) β†’ E i) (ModelPi H)

                    Given a finite family of ModelWithCorners I i on (E i, H i), we define the model with corners pi I on (Ξ  i, E i, ModelPi H). See note [Manifold type tags] for explanation about ModelPi H.

                    Equations
                    • One or more equations did not get rendered due to their size.
                    Instances For
                      @[reducible, inline]
                      abbrev ModelWithCorners.tangent {π•œ : Type u} [NontriviallyNormedField π•œ] {E : Type v} [NormedAddCommGroup E] [NormedSpace π•œ E] {H : Type w} [TopologicalSpace H] (I : ModelWithCorners π•œ E H) :
                      ModelWithCorners π•œ (E Γ— E) (ModelProd H E)

                      Special case of product model with corners, which is trivial on the second factor. This shows up as the model to tangent bundles.

                      Equations
                      Instances For
                        @[simp]
                        theorem modelWithCorners_prod_toPartialEquiv {π•œ : Type u_1} [NontriviallyNormedField π•œ] {E : Type u_2} [NormedAddCommGroup E] [NormedSpace π•œ E] {F : Type u_4} [NormedAddCommGroup F] [NormedSpace π•œ F] {H : Type u_5} [TopologicalSpace H] {G : Type u_7} [TopologicalSpace G] {I : ModelWithCorners π•œ E H} {J : ModelWithCorners π•œ F G} :
                        (I.prod J).toPartialEquiv = I.prod J.toPartialEquiv
                        @[simp]
                        theorem modelWithCorners_prod_coe {π•œ : Type u_1} [NontriviallyNormedField π•œ] {E : Type u_2} [NormedAddCommGroup E] [NormedSpace π•œ E] {E' : Type u_3} [NormedAddCommGroup E'] [NormedSpace π•œ E'] {H : Type u_5} [TopologicalSpace H] {H' : Type u_6} [TopologicalSpace H'] (I : ModelWithCorners π•œ E H) (I' : ModelWithCorners π•œ E' H') :
                        ↑(I.prod I') = Prod.map ↑I ↑I'
                        @[simp]
                        theorem modelWithCorners_prod_coe_symm {π•œ : Type u_1} [NontriviallyNormedField π•œ] {E : Type u_2} [NormedAddCommGroup E] [NormedSpace π•œ E] {E' : Type u_3} [NormedAddCommGroup E'] [NormedSpace π•œ E'] {H : Type u_5} [TopologicalSpace H] {H' : Type u_6} [TopologicalSpace H'] (I : ModelWithCorners π•œ E H) (I' : ModelWithCorners π•œ E' H') :
                        ↑(I.prod I').symm = Prod.map ↑I.symm ↑I'.symm
                        theorem modelWithCornersSelf_prod {π•œ : Type u_1} [NontriviallyNormedField π•œ] {E : Type u_2} [NormedAddCommGroup E] [NormedSpace π•œ E] {F : Type u_4} [NormedAddCommGroup F] [NormedSpace π•œ F] :
                        modelWithCornersSelf π•œ (E Γ— F) = (modelWithCornersSelf π•œ E).prod (modelWithCornersSelf π•œ F)

                        This lemma should be erased, or at least burn in hell, as it uses bad defeq: the left model with corners is for E times F, the right one for ModelProd E F, and there's a good reason we are distinguishing them.

                        theorem ModelWithCorners.range_prod {π•œ : Type u_1} [NontriviallyNormedField π•œ] {E : Type u_2} [NormedAddCommGroup E] [NormedSpace π•œ E] {F : Type u_4} [NormedAddCommGroup F] [NormedSpace π•œ F] {H : Type u_5} [TopologicalSpace H] {G : Type u_7} [TopologicalSpace G] {I : ModelWithCorners π•œ E H} {J : ModelWithCorners π•œ F G} :
                        Set.range ↑(I.prod J) = Set.range ↑I Γ—Λ’ Set.range ↑J
                        class ModelWithCorners.Boundaryless {π•œ : Type u_1} [NontriviallyNormedField π•œ] {E : Type u_2} [NormedAddCommGroup E] [NormedSpace π•œ E] {H : Type u_3} [TopologicalSpace H] (I : ModelWithCorners π•œ E H) :

                        Property ensuring that the model with corners I defines manifolds without boundary. This differs from the more general BoundarylessManifold, which requires every point on the manifold to be an interior point.

                        Instances
                          theorem ModelWithCorners.range_eq_univ {π•œ : Type u_1} [NontriviallyNormedField π•œ] {E : Type u_2} [NormedAddCommGroup E] [NormedSpace π•œ E] {H : Type u_3} [TopologicalSpace H] (I : ModelWithCorners π•œ E H) [I.Boundaryless] :
                          def ModelWithCorners.toHomeomorph {π•œ : Type u_1} [NontriviallyNormedField π•œ] {E : Type u_2} [NormedAddCommGroup E] [NormedSpace π•œ E] {H : Type u_3} [TopologicalSpace H] (I : ModelWithCorners π•œ E H) [I.Boundaryless] :

                          If I is a ModelWithCorners.Boundaryless model, then it is a homeomorphism.

                          Equations
                          • I.toHomeomorph = { toFun := ↑I.toPartialEquiv, invFun := I.invFun, left_inv := β‹―, right_inv := β‹―, continuous_toFun := β‹―, continuous_invFun := β‹― }
                          Instances For
                            @[simp]
                            theorem ModelWithCorners.toHomeomorph_symm_apply {π•œ : Type u_1} [NontriviallyNormedField π•œ] {E : Type u_2} [NormedAddCommGroup E] [NormedSpace π•œ E] {H : Type u_3} [TopologicalSpace H] (I : ModelWithCorners π•œ E H) [I.Boundaryless] (a✝ : E) :
                            I.toHomeomorph.symm a✝ = ↑I.symm a✝
                            @[simp]
                            theorem ModelWithCorners.toHomeomorph_apply {π•œ : Type u_1} [NontriviallyNormedField π•œ] {E : Type u_2} [NormedAddCommGroup E] [NormedSpace π•œ E] {H : Type u_3} [TopologicalSpace H] (I : ModelWithCorners π•œ E H) [I.Boundaryless] (a✝ : H) :
                            I.toHomeomorph a✝ = ↑I a✝
                            instance modelWithCornersSelf_boundaryless (π•œ : Type u_1) [NontriviallyNormedField π•œ] (E : Type u_2) [NormedAddCommGroup E] [NormedSpace π•œ E] :
                            (modelWithCornersSelf π•œ E).Boundaryless

                            The trivial model with corners has no boundary

                            instance ModelWithCorners.range_eq_univ_prod {π•œ : Type u} [NontriviallyNormedField π•œ] {E : Type v} [NormedAddCommGroup E] [NormedSpace π•œ E] {H : Type w} [TopologicalSpace H] (I : ModelWithCorners π•œ E H) [I.Boundaryless] {E' : Type v'} [NormedAddCommGroup E'] [NormedSpace π•œ E'] {H' : Type w'} [TopologicalSpace H'] (I' : ModelWithCorners π•œ E' H') [I'.Boundaryless] :
                            (I.prod I').Boundaryless

                            If two model with corners are boundaryless, their product also is

                            C^n functions on models with corners #

                            def contDiffPregroupoid (n : WithTop β„•βˆž) {π•œ : Type u_1} [NontriviallyNormedField π•œ] {E : Type u_2} [NormedAddCommGroup E] [NormedSpace π•œ E] {H : Type u_3} [TopologicalSpace H] (I : ModelWithCorners π•œ E H) :

                            Given a model with corners (E, H), we define the pregroupoid of C^n transformations of H as the maps that are C^n when read in E through I.

                            Equations
                            Instances For
                              def contDiffGroupoid (n : WithTop β„•βˆž) {π•œ : Type u_1} [NontriviallyNormedField π•œ] {E : Type u_2} [NormedAddCommGroup E] [NormedSpace π•œ E] {H : Type u_3} [TopologicalSpace H] (I : ModelWithCorners π•œ E H) :

                              Given a model with corners (E, H), we define the groupoid of invertible C^n transformations of H as the invertible maps that are C^n when read in E through I.

                              Equations
                              Instances For
                                theorem contDiffGroupoid_le {m n : WithTop β„•βˆž} {π•œ : Type u_1} [NontriviallyNormedField π•œ] {E : Type u_2} [NormedAddCommGroup E] [NormedSpace π•œ E] {H : Type u_3} [TopologicalSpace H] {I : ModelWithCorners π•œ E H} (h : m ≀ n) :

                                Inclusion of the groupoid of C^n local diffeos in the groupoid of C^m local diffeos when m ≀ n

                                theorem contDiffGroupoid_zero_eq {π•œ : Type u_1} [NontriviallyNormedField π•œ] {E : Type u_2} [NormedAddCommGroup E] [NormedSpace π•œ E] {H : Type u_3} [TopologicalSpace H] {I : ModelWithCorners π•œ E H} :

                                The groupoid of 0-times continuously differentiable maps is just the groupoid of all partial homeomorphisms

                                theorem ofSet_mem_contDiffGroupoid {n : WithTop β„•βˆž} {π•œ : Type u_1} [NontriviallyNormedField π•œ] {E : Type u_2} [NormedAddCommGroup E] [NormedSpace π•œ E] {H : Type u_3} [TopologicalSpace H] {I : ModelWithCorners π•œ E H} {s : Set H} (hs : IsOpen s) :

                                An identity partial homeomorphism belongs to the C^n groupoid.

                                theorem symm_trans_mem_contDiffGroupoid {n : WithTop β„•βˆž} {π•œ : Type u_1} [NontriviallyNormedField π•œ] {E : Type u_2} [NormedAddCommGroup E] [NormedSpace π•œ E] {H : Type u_3} [TopologicalSpace H] {I : ModelWithCorners π•œ E H} {M : Type u_4} [TopologicalSpace M] (e : PartialHomeomorph M H) :
                                e.symm.trans e ∈ contDiffGroupoid n I

                                The composition of a partial homeomorphism from H to M and its inverse belongs to the C^n groupoid.

                                theorem contDiffGroupoid_prod {n : WithTop β„•βˆž} {π•œ : Type u_1} [NontriviallyNormedField π•œ] {E : Type u_2} [NormedAddCommGroup E] [NormedSpace π•œ E] {H : Type u_3} [TopologicalSpace H] {E' : Type u_5} {H' : Type u_6} [NormedAddCommGroup E'] [NormedSpace π•œ E'] [TopologicalSpace H'] {I : ModelWithCorners π•œ E H} {I' : ModelWithCorners π•œ E' H'} {e : PartialHomeomorph H H} {e' : PartialHomeomorph H' H'} (he : e ∈ contDiffGroupoid n I) (he' : e' ∈ contDiffGroupoid n I') :
                                e.prod e' ∈ contDiffGroupoid n (I.prod I')

                                The product of two C^n partial homeomorphisms is C^n.

                                The C^n groupoid is closed under restriction.

                                C^n manifolds (possibly with boundary or corners) #

                                class IsManifold {π•œ : Type u_1} [NontriviallyNormedField π•œ] {E : Type u_2} [NormedAddCommGroup E] [NormedSpace π•œ E] {H : Type u_3} [TopologicalSpace H] (I : ModelWithCorners π•œ E H) (n : WithTop β„•βˆž) (M : Type u_4) [TopologicalSpace M] [ChartedSpace H M] extends HasGroupoid M (contDiffGroupoid n I) :

                                Typeclass defining manifolds with respect to a model with corners, over a field π•œ. This definition includes the model with corners I (which might allow boundary, corners, or not, so this class covers both manifolds with boundary and manifolds without boundary), and a smoothness parameter n : WithTop β„•βˆž (where n = 0 means topological manifold, n = ∞ means smooth manifold and n = Ο‰ means analytic manifold).

                                Instances
                                  @[deprecated IsManifold (since := "2025-01-09")]
                                  def SmoothManifoldWithCorners {π•œ : Type u_1} [NontriviallyNormedField π•œ] {E : Type u_2} [NormedAddCommGroup E] [NormedSpace π•œ E] {H : Type u_3} [TopologicalSpace H] (I : ModelWithCorners π•œ E H) (n : WithTop β„•βˆž) (M : Type u_4) [TopologicalSpace M] [ChartedSpace H M] :

                                  Alias of IsManifold.


                                  Typeclass defining manifolds with respect to a model with corners, over a field π•œ. This definition includes the model with corners I (which might allow boundary, corners, or not, so this class covers both manifolds with boundary and manifolds without boundary), and a smoothness parameter n : WithTop β„•βˆž (where n = 0 means topological manifold, n = ∞ means smooth manifold and n = Ο‰ means analytic manifold).

                                  Equations
                                  Instances For
                                    theorem IsManifold.mk' {π•œ : Type u_1} [NontriviallyNormedField π•œ] {E : Type u_2} [NormedAddCommGroup E] [NormedSpace π•œ E] {H : Type u_3} [TopologicalSpace H] (I : ModelWithCorners π•œ E H) (n : WithTop β„•βˆž) (M : Type u_4) [TopologicalSpace M] [ChartedSpace H M] [gr : HasGroupoid M (contDiffGroupoid n I)] :

                                    Building a C^n manifold from a HasGroupoid assumption.

                                    @[deprecated IsManifold.mk' (since := "2025-01-09")]
                                    theorem SmoothManifoldWithCorners.mk' {π•œ : Type u_1} [NontriviallyNormedField π•œ] {E : Type u_2} [NormedAddCommGroup E] [NormedSpace π•œ E] {H : Type u_3} [TopologicalSpace H] (I : ModelWithCorners π•œ E H) (n : WithTop β„•βˆž) (M : Type u_4) [TopologicalSpace M] [ChartedSpace H M] [gr : HasGroupoid M (contDiffGroupoid n I)] :

                                    Alias of IsManifold.mk'.


                                    Building a C^n manifold from a HasGroupoid assumption.

                                    theorem isManifold_of_contDiffOn {π•œ : Type u_1} [NontriviallyNormedField π•œ] {E : Type u_2} [NormedAddCommGroup E] [NormedSpace π•œ E] {H : Type u_3} [TopologicalSpace H] (I : ModelWithCorners π•œ E H) (n : WithTop β„•βˆž) (M : Type u_4) [TopologicalSpace M] [ChartedSpace H M] (h : βˆ€ (e e' : PartialHomeomorph M H), e ∈ atlas H M β†’ e' ∈ atlas H M β†’ ContDiffOn π•œ n (↑I ∘ ↑(e.symm.trans e') ∘ ↑I.symm) (↑I.symm ⁻¹' (e.symm.trans e').source ∩ Set.range ↑I)) :
                                    @[deprecated isManifold_of_contDiffOn (since := "2025-01-09")]
                                    theorem smoothManifoldWithCorners_of_contDiffOn {π•œ : Type u_1} [NontriviallyNormedField π•œ] {E : Type u_2} [NormedAddCommGroup E] [NormedSpace π•œ E] {H : Type u_3} [TopologicalSpace H] (I : ModelWithCorners π•œ E H) (n : WithTop β„•βˆž) (M : Type u_4) [TopologicalSpace M] [ChartedSpace H M] (h : βˆ€ (e e' : PartialHomeomorph M H), e ∈ atlas H M β†’ e' ∈ atlas H M β†’ ContDiffOn π•œ n (↑I ∘ ↑(e.symm.trans e') ∘ ↑I.symm) (↑I.symm ⁻¹' (e.symm.trans e').source ∩ Set.range ↑I)) :

                                    Alias of isManifold_of_contDiffOn.

                                    instance intIsManifoldModelSpace {π•œ : Type u_1} [NontriviallyNormedField π•œ] {E : Type u_2} [NormedAddCommGroup E] [NormedSpace π•œ E] {H : Type u_3} [TopologicalSpace H] {I : ModelWithCorners π•œ E H} {n : WithTop β„•βˆž} :

                                    For any model with corners, the model space is a C^n manifold

                                    theorem IsManifold.of_le {π•œ : Type u_1} [NontriviallyNormedField π•œ] {E : Type u_2} [NormedAddCommGroup E] [NormedSpace π•œ E] {H : Type u_3} [TopologicalSpace H] {I : ModelWithCorners π•œ E H} {M : Type u_4} [TopologicalSpace M] [ChartedSpace H M] {m n : WithTop β„•βˆž} (hmn : m ≀ n) [IsManifold I n M] :

                                    A typeclass registering that a smoothness exponent is smaller than ∞. Used to deduce that some manifolds are C^n when they are C^∞.

                                    Instances
                                      instance IsManifold.instOfSomeENatTopOfLEInfty {π•œ : Type u_1} [NontriviallyNormedField π•œ] {E : Type u_2} [NormedAddCommGroup E] [NormedSpace π•œ E] {H : Type u_3} [TopologicalSpace H] {I : ModelWithCorners π•œ E H} {M : Type u_4} [TopologicalSpace M] [ChartedSpace H M] {a : WithTop β„•βˆž} [IsManifold I (β†‘βŠ€) M] [h : ENat.LEInfty a] :
                                      instance IsManifold.instOfTopWithTopENat {π•œ : Type u_1} [NontriviallyNormedField π•œ] {E : Type u_2} [NormedAddCommGroup E] [NormedSpace π•œ E] {H : Type u_3} [TopologicalSpace H] {I : ModelWithCorners π•œ E H} {M : Type u_4} [TopologicalSpace M] [ChartedSpace H M] {a : WithTop β„•βˆž} [IsManifold I ⊀ M] :
                                      instance IsManifold.instOfNatWithTopENat {π•œ : Type u_1} [NontriviallyNormedField π•œ] {E : Type u_2} [NormedAddCommGroup E] [NormedSpace π•œ E] {H : Type u_3} [TopologicalSpace H] {I : ModelWithCorners π•œ E H} {M : Type u_4} [TopologicalSpace M] [ChartedSpace H M] :
                                      instance IsManifold.instOfNatWithTopENat_1 {π•œ : Type u_1} [NontriviallyNormedField π•œ] {E : Type u_2} [NormedAddCommGroup E] [NormedSpace π•œ E] {H : Type u_3} [TopologicalSpace H] {I : ModelWithCorners π•œ E H} {M : Type u_4} [TopologicalSpace M] [ChartedSpace H M] [IsManifold I 2 M] :
                                      def IsManifold.maximalAtlas {π•œ : Type u_1} [NontriviallyNormedField π•œ] {E : Type u_2} [NormedAddCommGroup E] [NormedSpace π•œ E] {H : Type u_3} [TopologicalSpace H] (I : ModelWithCorners π•œ E H) (n : WithTop β„•βˆž) (M : Type u_4) [TopologicalSpace M] [ChartedSpace H M] :

                                      The maximal atlas of M for the C^n manifold with corners structure corresponding to the model with corners I.

                                      Equations
                                      Instances For
                                        theorem IsManifold.subset_maximalAtlas {π•œ : Type u_1} [NontriviallyNormedField π•œ] {E : Type u_2} [NormedAddCommGroup E] [NormedSpace π•œ E] {H : Type u_3} [TopologicalSpace H] {I : ModelWithCorners π•œ E H} {n : WithTop β„•βˆž} {M : Type u_4} [TopologicalSpace M] [ChartedSpace H M] [IsManifold I n M] :
                                        theorem IsManifold.chart_mem_maximalAtlas {π•œ : Type u_1} [NontriviallyNormedField π•œ] {E : Type u_2} [NormedAddCommGroup E] [NormedSpace π•œ E] {H : Type u_3} [TopologicalSpace H] {I : ModelWithCorners π•œ E H} {n : WithTop β„•βˆž} {M : Type u_4} [TopologicalSpace M] [ChartedSpace H M] [IsManifold I n M] (x : M) :
                                        theorem IsManifold.compatible_of_mem_maximalAtlas {π•œ : Type u_1} [NontriviallyNormedField π•œ] {E : Type u_2} [NormedAddCommGroup E] [NormedSpace π•œ E] {H : Type u_3} [TopologicalSpace H] {I : ModelWithCorners π•œ E H} {n : WithTop β„•βˆž} {M : Type u_4} [TopologicalSpace M] [ChartedSpace H M] {e e' : PartialHomeomorph M H} (he : e ∈ maximalAtlas I n M) (he' : e' ∈ maximalAtlas I n M) :
                                        e.symm.trans e' ∈ contDiffGroupoid n I
                                        theorem IsManifold.maximalAtlas_subset_of_le {π•œ : Type u_1} [NontriviallyNormedField π•œ] {E : Type u_2} [NormedAddCommGroup E] [NormedSpace π•œ E] {H : Type u_3} [TopologicalSpace H] {I : ModelWithCorners π•œ E H} {M : Type u_4} [TopologicalSpace M] [ChartedSpace H M] {m n : WithTop β„•βˆž} (h : m ≀ n) :
                                        instance IsManifold.empty {π•œ : Type u_1} [NontriviallyNormedField π•œ] {E : Type u_2} [NormedAddCommGroup E] [NormedSpace π•œ E] {H : Type u_3} [TopologicalSpace H] {I : ModelWithCorners π•œ E H} {n : WithTop β„•βˆž} {M : Type u_4} [TopologicalSpace M] [ChartedSpace H M] [IsEmpty M] :

                                        The empty set is a C^n manifold w.r.t. any charted space and model.

                                        instance IsManifold.prod {n : WithTop β„•βˆž} {π•œ : Type u_5} [NontriviallyNormedField π•œ] {E : Type u_6} [NormedAddCommGroup E] [NormedSpace π•œ E] {E' : Type u_7} [NormedAddCommGroup E'] [NormedSpace π•œ E'] {H : Type u_8} [TopologicalSpace H] {I : ModelWithCorners π•œ E H} {H' : Type u_9} [TopologicalSpace H'] {I' : ModelWithCorners π•œ E' H'} (M : Type u_10) [TopologicalSpace M] [ChartedSpace H M] [IsManifold I n M] (M' : Type u_11) [TopologicalSpace M'] [ChartedSpace H' M'] [IsManifold I' n M'] :
                                        IsManifold (I.prod I') n (M Γ— M')

                                        The product of two C^n manifolds is naturally a C^n manifold.

                                        theorem PartialHomeomorph.isManifold_singleton {π•œ : Type u_1} [NontriviallyNormedField π•œ] {E : Type u_2} [NormedAddCommGroup E] [NormedSpace π•œ E] {H : Type u_3} [TopologicalSpace H] {I : ModelWithCorners π•œ E H} {n : WithTop β„•βˆž} {M : Type u_4} [TopologicalSpace M] (e : PartialHomeomorph M H) (h : e.source = Set.univ) :
                                        @[deprecated PartialHomeomorph.isManifold_singleton (since := "2025-01-09")]
                                        theorem PartialHomeomorph.singleton_smoothManifoldWithCorners {π•œ : Type u_1} [NontriviallyNormedField π•œ] {E : Type u_2} [NormedAddCommGroup E] [NormedSpace π•œ E] {H : Type u_3} [TopologicalSpace H] {I : ModelWithCorners π•œ E H} {n : WithTop β„•βˆž} {M : Type u_4} [TopologicalSpace M] (e : PartialHomeomorph M H) (h : e.source = Set.univ) :

                                        Alias of PartialHomeomorph.isManifold_singleton.

                                        theorem Topology.IsOpenEmbedding.isManifold_singleton {π•œ : Type u_1} {E : Type u_2} {H : Type u_3} [NontriviallyNormedField π•œ] [NormedAddCommGroup E] [NormedSpace π•œ E] [TopologicalSpace H] {I : ModelWithCorners π•œ E H} {n : WithTop β„•βˆž} {M : Type u_4} [TopologicalSpace M] [Nonempty M] {f : M β†’ H} (h : IsOpenEmbedding f) :
                                        @[deprecated Topology.IsOpenEmbedding.isManifold_singleton (since := "2025-01-09")]
                                        theorem Topology.IsOpenEmbedding.singleton_smoothManifoldWithCorners {π•œ : Type u_1} {E : Type u_2} {H : Type u_3} [NontriviallyNormedField π•œ] [NormedAddCommGroup E] [NormedSpace π•œ E] [TopologicalSpace H] {I : ModelWithCorners π•œ E H} {n : WithTop β„•βˆž} {M : Type u_4} [TopologicalSpace M] [Nonempty M] {f : M β†’ H} (h : IsOpenEmbedding f) :

                                        Alias of Topology.IsOpenEmbedding.isManifold_singleton.

                                        @[deprecated Topology.IsOpenEmbedding.isManifold_singleton (since := "2024-10-18")]
                                        theorem OpenEmbedding.singleton_smoothManifoldWithCorners {π•œ : Type u_1} {E : Type u_2} {H : Type u_3} [NontriviallyNormedField π•œ] [NormedAddCommGroup E] [NormedSpace π•œ E] [TopologicalSpace H] {I : ModelWithCorners π•œ E H} {n : WithTop β„•βˆž} {M : Type u_4} [TopologicalSpace M] [Nonempty M] {f : M β†’ H} (h : Topology.IsOpenEmbedding f) :

                                        Alias of Topology.IsOpenEmbedding.isManifold_singleton.

                                        instance TopologicalSpace.Opens.instIsManifoldSubtypeMem {π•œ : Type u_1} [NontriviallyNormedField π•œ] {E : Type u_2} [NormedAddCommGroup E] [NormedSpace π•œ E] {H : Type u_3} [TopologicalSpace H] {I : ModelWithCorners π•œ E H} {n : WithTop β„•βˆž} {M : Type u_4} [TopologicalSpace M] [ChartedSpace H M] [IsManifold I n M] (s : Opens M) :
                                        IsManifold I n β†₯s

                                        Extended charts #

                                        In a C^n manifold with corners, the model space is the space H. However, we will also need to use extended charts taking values in the model vector space E. These extended charts are not PartialHomeomorph as the target is not open in E in general, but we can still register them as PartialEquiv.

                                        def PartialHomeomorph.extend {π•œ : Type u_1} {E : Type u_2} {M : Type u_3} {H : Type u_4} [NontriviallyNormedField π•œ] [NormedAddCommGroup E] [NormedSpace π•œ E] [TopologicalSpace H] [TopologicalSpace M] (f : PartialHomeomorph M H) (I : ModelWithCorners π•œ E H) :

                                        Given a chart f on a manifold with corners, f.extend I is the extended chart to the model vector space.

                                        Equations
                                        • f.extend I = f.trans I.toPartialEquiv
                                        Instances For
                                          theorem PartialHomeomorph.extend_coe {π•œ : Type u_1} {E : Type u_2} {M : Type u_3} {H : Type u_4} [NontriviallyNormedField π•œ] [NormedAddCommGroup E] [NormedSpace π•œ E] [TopologicalSpace H] [TopologicalSpace M] (f : PartialHomeomorph M H) {I : ModelWithCorners π•œ E H} :
                                          ↑(f.extend I) = ↑I ∘ ↑f
                                          theorem PartialHomeomorph.extend_coe_symm {π•œ : Type u_1} {E : Type u_2} {M : Type u_3} {H : Type u_4} [NontriviallyNormedField π•œ] [NormedAddCommGroup E] [NormedSpace π•œ E] [TopologicalSpace H] [TopologicalSpace M] (f : PartialHomeomorph M H) {I : ModelWithCorners π•œ E H} :
                                          ↑(f.extend I).symm = ↑f.symm ∘ ↑I.symm
                                          theorem PartialHomeomorph.extend_source {π•œ : Type u_1} {E : Type u_2} {M : Type u_3} {H : Type u_4} [NontriviallyNormedField π•œ] [NormedAddCommGroup E] [NormedSpace π•œ E] [TopologicalSpace H] [TopologicalSpace M] (f : PartialHomeomorph M H) {I : ModelWithCorners π•œ E H} :
                                          (f.extend I).source = f.source
                                          theorem PartialHomeomorph.isOpen_extend_source {π•œ : Type u_1} {E : Type u_2} {M : Type u_3} {H : Type u_4} [NontriviallyNormedField π•œ] [NormedAddCommGroup E] [NormedSpace π•œ E] [TopologicalSpace H] [TopologicalSpace M] (f : PartialHomeomorph M H) {I : ModelWithCorners π•œ E H} :
                                          IsOpen (f.extend I).source
                                          theorem PartialHomeomorph.extend_target {π•œ : Type u_1} {E : Type u_2} {M : Type u_3} {H : Type u_4} [NontriviallyNormedField π•œ] [NormedAddCommGroup E] [NormedSpace π•œ E] [TopologicalSpace H] [TopologicalSpace M] (f : PartialHomeomorph M H) {I : ModelWithCorners π•œ E H} :
                                          (f.extend I).target = ↑I.symm ⁻¹' f.target ∩ Set.range ↑I
                                          theorem PartialHomeomorph.extend_target' {π•œ : Type u_1} {E : Type u_2} {M : Type u_3} {H : Type u_4} [NontriviallyNormedField π•œ] [NormedAddCommGroup E] [NormedSpace π•œ E] [TopologicalSpace H] [TopologicalSpace M] (f : PartialHomeomorph M H) {I : ModelWithCorners π•œ E H} :
                                          (f.extend I).target = ↑I '' f.target
                                          theorem PartialHomeomorph.isOpen_extend_target {π•œ : Type u_1} {E : Type u_2} {M : Type u_3} {H : Type u_4} [NontriviallyNormedField π•œ] [NormedAddCommGroup E] [NormedSpace π•œ E] [TopologicalSpace H] [TopologicalSpace M] (f : PartialHomeomorph M H) {I : ModelWithCorners π•œ E H} [I.Boundaryless] :
                                          IsOpen (f.extend I).target
                                          theorem PartialHomeomorph.mapsTo_extend {π•œ : Type u_1} {E : Type u_2} {M : Type u_3} {H : Type u_4} [NontriviallyNormedField π•œ] [NormedAddCommGroup E] [NormedSpace π•œ E] [TopologicalSpace H] [TopologicalSpace M] (f : PartialHomeomorph M H) {I : ModelWithCorners π•œ E H} {s : Set M} (hs : s βŠ† f.source) :
                                          Set.MapsTo (↑(f.extend I)) s (↑(f.extend I).symm ⁻¹' s ∩ Set.range ↑I)
                                          theorem PartialHomeomorph.extend_left_inv {π•œ : Type u_1} {E : Type u_2} {M : Type u_3} {H : Type u_4} [NontriviallyNormedField π•œ] [NormedAddCommGroup E] [NormedSpace π•œ E] [TopologicalSpace H] [TopologicalSpace M] (f : PartialHomeomorph M H) {I : ModelWithCorners π•œ E H} {x : M} (hxf : x ∈ f.source) :
                                          ↑(f.extend I).symm (↑(f.extend I) x) = x
                                          theorem PartialHomeomorph.extend_left_inv' {π•œ : Type u_1} {E : Type u_2} {M : Type u_3} {H : Type u_4} [NontriviallyNormedField π•œ] [NormedAddCommGroup E] [NormedSpace π•œ E] [TopologicalSpace H] [TopologicalSpace M] (f : PartialHomeomorph M H) {I : ModelWithCorners π•œ E H} {t : Set M} (ht : t βŠ† f.source) :
                                          ↑(f.extend I).symm ∘ ↑(f.extend I) '' t = t

                                          Variant of f.extend_left_inv I, stated in terms of images.

                                          theorem PartialHomeomorph.extend_source_mem_nhds {π•œ : Type u_1} {E : Type u_2} {M : Type u_3} {H : Type u_4} [NontriviallyNormedField π•œ] [NormedAddCommGroup E] [NormedSpace π•œ E] [TopologicalSpace H] [TopologicalSpace M] (f : PartialHomeomorph M H) {I : ModelWithCorners π•œ E H} {x : M} (h : x ∈ f.source) :
                                          (f.extend I).source ∈ nhds x
                                          theorem PartialHomeomorph.extend_source_mem_nhdsWithin {π•œ : Type u_1} {E : Type u_2} {M : Type u_3} {H : Type u_4} [NontriviallyNormedField π•œ] [NormedAddCommGroup E] [NormedSpace π•œ E] [TopologicalSpace H] [TopologicalSpace M] (f : PartialHomeomorph M H) {I : ModelWithCorners π•œ E H} {s : Set M} {x : M} (h : x ∈ f.source) :
                                          (f.extend I).source ∈ nhdsWithin x s
                                          theorem PartialHomeomorph.continuousOn_extend {π•œ : Type u_1} {E : Type u_2} {M : Type u_3} {H : Type u_4} [NontriviallyNormedField π•œ] [NormedAddCommGroup E] [NormedSpace π•œ E] [TopologicalSpace H] [TopologicalSpace M] (f : PartialHomeomorph M H) {I : ModelWithCorners π•œ E H} :
                                          ContinuousOn (↑(f.extend I)) (f.extend I).source
                                          theorem PartialHomeomorph.continuousAt_extend {π•œ : Type u_1} {E : Type u_2} {M : Type u_3} {H : Type u_4} [NontriviallyNormedField π•œ] [NormedAddCommGroup E] [NormedSpace π•œ E] [TopologicalSpace H] [TopologicalSpace M] (f : PartialHomeomorph M H) {I : ModelWithCorners π•œ E H} {x : M} (h : x ∈ f.source) :
                                          ContinuousAt (↑(f.extend I)) x
                                          theorem PartialHomeomorph.map_extend_nhds {π•œ : Type u_1} {E : Type u_2} {M : Type u_3} {H : Type u_4} [NontriviallyNormedField π•œ] [NormedAddCommGroup E] [NormedSpace π•œ E] [TopologicalSpace H] [TopologicalSpace M] (f : PartialHomeomorph M H) {I : ModelWithCorners π•œ E H} {x : M} (hy : x ∈ f.source) :
                                          Filter.map (↑(f.extend I)) (nhds x) = nhdsWithin (↑(f.extend I) x) (Set.range ↑I)
                                          theorem PartialHomeomorph.map_extend_nhds_of_mem_interior_range {π•œ : Type u_1} {E : Type u_2} {M : Type u_3} {H : Type u_4} [NontriviallyNormedField π•œ] [NormedAddCommGroup E] [NormedSpace π•œ E] [TopologicalSpace H] [TopologicalSpace M] (f : PartialHomeomorph M H) {I : ModelWithCorners π•œ E H} {x : M} (hx : x ∈ f.source) (h'x : ↑(f.extend I) x ∈ interior (Set.range ↑I)) :
                                          Filter.map (↑(f.extend I)) (nhds x) = nhds (↑(f.extend I) x)
                                          theorem PartialHomeomorph.map_extend_nhds_of_boundaryless {π•œ : Type u_1} {E : Type u_2} {M : Type u_3} {H : Type u_4} [NontriviallyNormedField π•œ] [NormedAddCommGroup E] [NormedSpace π•œ E] [TopologicalSpace H] [TopologicalSpace M] (f : PartialHomeomorph M H) {I : ModelWithCorners π•œ E H} [I.Boundaryless] {x : M} (hx : x ∈ f.source) :
                                          Filter.map (↑(f.extend I)) (nhds x) = nhds (↑(f.extend I) x)
                                          theorem PartialHomeomorph.extend_target_mem_nhdsWithin {π•œ : Type u_1} {E : Type u_2} {M : Type u_3} {H : Type u_4} [NontriviallyNormedField π•œ] [NormedAddCommGroup E] [NormedSpace π•œ E] [TopologicalSpace H] [TopologicalSpace M] (f : PartialHomeomorph M H) {I : ModelWithCorners π•œ E H} {y : M} (hy : y ∈ f.source) :
                                          (f.extend I).target ∈ nhdsWithin (↑(f.extend I) y) (Set.range ↑I)
                                          theorem PartialHomeomorph.extend_image_nhd_mem_nhds_of_boundaryless {π•œ : Type u_1} {E : Type u_2} {M : Type u_3} {H : Type u_4} [NontriviallyNormedField π•œ] [NormedAddCommGroup E] [NormedSpace π•œ E] [TopologicalSpace H] [TopologicalSpace M] (f : PartialHomeomorph M H) {I : ModelWithCorners π•œ E H} [I.Boundaryless] {x : M} (hx : x ∈ f.source) {s : Set M} (h : s ∈ nhds x) :
                                          ↑(f.extend I) '' s ∈ nhds (↑(f.extend I) x)
                                          theorem PartialHomeomorph.extend_image_nhd_mem_nhds_of_mem_interior_range {π•œ : Type u_1} {E : Type u_2} {M : Type u_3} {H : Type u_4} [NontriviallyNormedField π•œ] [NormedAddCommGroup E] [NormedSpace π•œ E] [TopologicalSpace H] [TopologicalSpace M] (f : PartialHomeomorph M H) {I : ModelWithCorners π•œ E H} {x : M} (hx : x ∈ f.source) (h'x : ↑(f.extend I) x ∈ interior (Set.range ↑I)) {s : Set M} (h : s ∈ nhds x) :
                                          ↑(f.extend I) '' s ∈ nhds (↑(f.extend I) x)
                                          theorem PartialHomeomorph.extend_target_subset_range {π•œ : Type u_1} {E : Type u_2} {M : Type u_3} {H : Type u_4} [NontriviallyNormedField π•œ] [NormedAddCommGroup E] [NormedSpace π•œ E] [TopologicalSpace H] [TopologicalSpace M] (f : PartialHomeomorph M H) {I : ModelWithCorners π•œ E H} :
                                          (f.extend I).target βŠ† Set.range ↑I
                                          theorem PartialHomeomorph.interior_extend_target_subset_interior_range {π•œ : Type u_1} {E : Type u_2} {M : Type u_3} {H : Type u_4} [NontriviallyNormedField π•œ] [NormedAddCommGroup E] [NormedSpace π•œ E] [TopologicalSpace H] [TopologicalSpace M] (f : PartialHomeomorph M H) {I : ModelWithCorners π•œ E H} :
                                          interior (f.extend I).target βŠ† interior (Set.range ↑I)
                                          theorem PartialHomeomorph.mem_interior_extend_target {π•œ : Type u_1} {E : Type u_2} {M : Type u_3} {H : Type u_4} [NontriviallyNormedField π•œ] [NormedAddCommGroup E] [NormedSpace π•œ E] [TopologicalSpace H] [TopologicalSpace M] (f : PartialHomeomorph M H) {I : ModelWithCorners π•œ E H} {y : H} (hy : y ∈ f.target) (hy' : ↑I y ∈ interior (Set.range ↑I)) :
                                          ↑I y ∈ interior (f.extend I).target

                                          If y ∈ f.target and I y ∈ interior (range I), then I y is an interior point of (I ∘ f).target.

                                          theorem PartialHomeomorph.nhdsWithin_extend_target_eq {π•œ : Type u_1} {E : Type u_2} {M : Type u_3} {H : Type u_4} [NontriviallyNormedField π•œ] [NormedAddCommGroup E] [NormedSpace π•œ E] [TopologicalSpace H] [TopologicalSpace M] (f : PartialHomeomorph M H) {I : ModelWithCorners π•œ E H} {y : M} (hy : y ∈ f.source) :
                                          nhdsWithin (↑(f.extend I) y) (f.extend I).target = nhdsWithin (↑(f.extend I) y) (Set.range ↑I)
                                          theorem PartialHomeomorph.extend_target_eventuallyEq {π•œ : Type u_1} {E : Type u_2} {M : Type u_3} {H : Type u_4} [NontriviallyNormedField π•œ] [NormedAddCommGroup E] [NormedSpace π•œ E] [TopologicalSpace H] [TopologicalSpace M] (f : PartialHomeomorph M H) {I : ModelWithCorners π•œ E H} {y : M} (hy : y ∈ f.source) :
                                          (f.extend I).target =αΆ [nhds (↑(f.extend I) y)] Set.range ↑I
                                          theorem PartialHomeomorph.continuousAt_extend_symm' {π•œ : Type u_1} {E : Type u_2} {M : Type u_3} {H : Type u_4} [NontriviallyNormedField π•œ] [NormedAddCommGroup E] [NormedSpace π•œ E] [TopologicalSpace H] [TopologicalSpace M] (f : PartialHomeomorph M H) {I : ModelWithCorners π•œ E H} {x : E} (h : x ∈ (f.extend I).target) :
                                          ContinuousAt (↑(f.extend I).symm) x
                                          theorem PartialHomeomorph.continuousAt_extend_symm {π•œ : Type u_1} {E : Type u_2} {M : Type u_3} {H : Type u_4} [NontriviallyNormedField π•œ] [NormedAddCommGroup E] [NormedSpace π•œ E] [TopologicalSpace H] [TopologicalSpace M] (f : PartialHomeomorph M H) {I : ModelWithCorners π•œ E H} {x : M} (h : x ∈ f.source) :
                                          ContinuousAt (↑(f.extend I).symm) (↑(f.extend I) x)
                                          theorem PartialHomeomorph.continuousOn_extend_symm {π•œ : Type u_1} {E : Type u_2} {M : Type u_3} {H : Type u_4} [NontriviallyNormedField π•œ] [NormedAddCommGroup E] [NormedSpace π•œ E] [TopologicalSpace H] [TopologicalSpace M] (f : PartialHomeomorph M H) {I : ModelWithCorners π•œ E H} :
                                          ContinuousOn (↑(f.extend I).symm) (f.extend I).target
                                          theorem PartialHomeomorph.extend_symm_continuousWithinAt_comp_right_iff {π•œ : Type u_1} {E : Type u_2} {M : Type u_3} {H : Type u_4} [NontriviallyNormedField π•œ] [NormedAddCommGroup E] [NormedSpace π•œ E] [TopologicalSpace H] [TopologicalSpace M] (f : PartialHomeomorph M H) {I : ModelWithCorners π•œ E H} {X : Type u_8} [TopologicalSpace X] {g : M β†’ X} {s : Set M} {x : M} :
                                          ContinuousWithinAt (g ∘ ↑(f.extend I).symm) (↑(f.extend I).symm ⁻¹' s ∩ Set.range ↑I) (↑(f.extend I) x) ↔ ContinuousWithinAt (g ∘ ↑f.symm) (↑f.symm ⁻¹' s) (↑f x)
                                          theorem PartialHomeomorph.isOpen_extend_preimage' {π•œ : Type u_1} {E : Type u_2} {M : Type u_3} {H : Type u_4} [NontriviallyNormedField π•œ] [NormedAddCommGroup E] [NormedSpace π•œ E] [TopologicalSpace H] [TopologicalSpace M] (f : PartialHomeomorph M H) {I : ModelWithCorners π•œ E H} {s : Set E} (hs : IsOpen s) :
                                          IsOpen ((f.extend I).source ∩ ↑(f.extend I) ⁻¹' s)
                                          theorem PartialHomeomorph.isOpen_extend_preimage {π•œ : Type u_1} {E : Type u_2} {M : Type u_3} {H : Type u_4} [NontriviallyNormedField π•œ] [NormedAddCommGroup E] [NormedSpace π•œ E] [TopologicalSpace H] [TopologicalSpace M] (f : PartialHomeomorph M H) {I : ModelWithCorners π•œ E H} {s : Set E} (hs : IsOpen s) :
                                          IsOpen (f.source ∩ ↑(f.extend I) ⁻¹' s)
                                          theorem PartialHomeomorph.map_extend_nhdsWithin_eq_image {π•œ : Type u_1} {E : Type u_2} {M : Type u_3} {H : Type u_4} [NontriviallyNormedField π•œ] [NormedAddCommGroup E] [NormedSpace π•œ E] [TopologicalSpace H] [TopologicalSpace M] (f : PartialHomeomorph M H) {I : ModelWithCorners π•œ E H} {s : Set M} {y : M} (hy : y ∈ f.source) :
                                          Filter.map (↑(f.extend I)) (nhdsWithin y s) = nhdsWithin (↑(f.extend I) y) (↑(f.extend I) '' ((f.extend I).source ∩ s))
                                          theorem PartialHomeomorph.map_extend_nhdsWithin_eq_image_of_subset {π•œ : Type u_1} {E : Type u_2} {M : Type u_3} {H : Type u_4} [NontriviallyNormedField π•œ] [NormedAddCommGroup E] [NormedSpace π•œ E] [TopologicalSpace H] [TopologicalSpace M] (f : PartialHomeomorph M H) {I : ModelWithCorners π•œ E H} {s : Set M} {y : M} (hy : y ∈ f.source) (hs : s βŠ† f.source) :
                                          Filter.map (↑(f.extend I)) (nhdsWithin y s) = nhdsWithin (↑(f.extend I) y) (↑(f.extend I) '' s)
                                          theorem PartialHomeomorph.map_extend_nhdsWithin {π•œ : Type u_1} {E : Type u_2} {M : Type u_3} {H : Type u_4} [NontriviallyNormedField π•œ] [NormedAddCommGroup E] [NormedSpace π•œ E] [TopologicalSpace H] [TopologicalSpace M] (f : PartialHomeomorph M H) {I : ModelWithCorners π•œ E H} {s : Set M} {y : M} (hy : y ∈ f.source) :
                                          Filter.map (↑(f.extend I)) (nhdsWithin y s) = nhdsWithin (↑(f.extend I) y) (↑(f.extend I).symm ⁻¹' s ∩ Set.range ↑I)
                                          theorem PartialHomeomorph.map_extend_symm_nhdsWithin {π•œ : Type u_1} {E : Type u_2} {M : Type u_3} {H : Type u_4} [NontriviallyNormedField π•œ] [NormedAddCommGroup E] [NormedSpace π•œ E] [TopologicalSpace H] [TopologicalSpace M] (f : PartialHomeomorph M H) {I : ModelWithCorners π•œ E H} {s : Set M} {y : M} (hy : y ∈ f.source) :
                                          Filter.map (↑(f.extend I).symm) (nhdsWithin (↑(f.extend I) y) (↑(f.extend I).symm ⁻¹' s ∩ Set.range ↑I)) = nhdsWithin y s
                                          theorem PartialHomeomorph.map_extend_symm_nhdsWithin_range {π•œ : Type u_1} {E : Type u_2} {M : Type u_3} {H : Type u_4} [NontriviallyNormedField π•œ] [NormedAddCommGroup E] [NormedSpace π•œ E] [TopologicalSpace H] [TopologicalSpace M] (f : PartialHomeomorph M H) {I : ModelWithCorners π•œ E H} {y : M} (hy : y ∈ f.source) :
                                          Filter.map (↑(f.extend I).symm) (nhdsWithin (↑(f.extend I) y) (Set.range ↑I)) = nhds y
                                          theorem PartialHomeomorph.tendsto_extend_comp_iff {π•œ : Type u_1} {E : Type u_2} {M : Type u_3} {H : Type u_4} [NontriviallyNormedField π•œ] [NormedAddCommGroup E] [NormedSpace π•œ E] [TopologicalSpace H] [TopologicalSpace M] (f : PartialHomeomorph M H) {I : ModelWithCorners π•œ E H} {Ξ± : Type u_8} {l : Filter Ξ±} {g : Ξ± β†’ M} (hg : βˆ€αΆ  (z : Ξ±) in l, g z ∈ f.source) {y : M} (hy : y ∈ f.source) :
                                          Filter.Tendsto (↑(f.extend I) ∘ g) l (nhds (↑(f.extend I) y)) ↔ Filter.Tendsto g l (nhds y)
                                          theorem PartialHomeomorph.continuousWithinAt_writtenInExtend_iff {π•œ : Type u_1} {E : Type u_2} {M : Type u_3} {H : Type u_4} {E' : Type u_5} {M' : Type u_6} {H' : Type u_7} [NontriviallyNormedField π•œ] [NormedAddCommGroup E] [NormedSpace π•œ E] [TopologicalSpace H] [TopologicalSpace M] (f : PartialHomeomorph M H) {I : ModelWithCorners π•œ E H} [NormedAddCommGroup E'] [NormedSpace π•œ E'] [TopologicalSpace H'] [TopologicalSpace M'] {I' : ModelWithCorners π•œ E' H'} {s : Set M} {f' : PartialHomeomorph M' H'} {g : M β†’ M'} {y : M} (hy : y ∈ f.source) (hgy : g y ∈ f'.source) (hmaps : Set.MapsTo g s f'.source) :
                                          ContinuousWithinAt (↑(f'.extend I') ∘ g ∘ ↑(f.extend I).symm) (↑(f.extend I).symm ⁻¹' s ∩ Set.range ↑I) (↑(f.extend I) y) ↔ ContinuousWithinAt g s y
                                          theorem PartialHomeomorph.continuousOn_writtenInExtend_iff {π•œ : Type u_1} {E : Type u_2} {M : Type u_3} {H : Type u_4} {E' : Type u_5} {M' : Type u_6} {H' : Type u_7} [NontriviallyNormedField π•œ] [NormedAddCommGroup E] [NormedSpace π•œ E] [TopologicalSpace H] [TopologicalSpace M] (f : PartialHomeomorph M H) {I : ModelWithCorners π•œ E H} [NormedAddCommGroup E'] [NormedSpace π•œ E'] [TopologicalSpace H'] [TopologicalSpace M'] {I' : ModelWithCorners π•œ E' H'} {s : Set M} {f' : PartialHomeomorph M' H'} {g : M β†’ M'} (hs : s βŠ† f.source) (hmaps : Set.MapsTo g s f'.source) :
                                          ContinuousOn (↑(f'.extend I') ∘ g ∘ ↑(f.extend I).symm) (↑(f.extend I) '' s) ↔ ContinuousOn g s

                                          If s βŠ† f.source and g x ∈ f'.source whenever x ∈ s, then g is continuous on s if and only if g written in charts f.extend I and f'.extend I' is continuous on f.extend I '' s.

                                          theorem PartialHomeomorph.extend_preimage_mem_nhdsWithin {π•œ : Type u_1} {E : Type u_2} {M : Type u_3} {H : Type u_4} [NontriviallyNormedField π•œ] [NormedAddCommGroup E] [NormedSpace π•œ E] [TopologicalSpace H] [TopologicalSpace M] (f : PartialHomeomorph M H) {I : ModelWithCorners π•œ E H} {s t : Set M} {x : M} (h : x ∈ f.source) (ht : t ∈ nhdsWithin x s) :
                                          ↑(f.extend I).symm ⁻¹' t ∈ nhdsWithin (↑(f.extend I) x) (↑(f.extend I).symm ⁻¹' s ∩ Set.range ↑I)

                                          Technical lemma ensuring that the preimage under an extended chart of a neighborhood of a point in the source is a neighborhood of the preimage, within a set.

                                          theorem PartialHomeomorph.extend_preimage_mem_nhds {π•œ : Type u_1} {E : Type u_2} {M : Type u_3} {H : Type u_4} [NontriviallyNormedField π•œ] [NormedAddCommGroup E] [NormedSpace π•œ E] [TopologicalSpace H] [TopologicalSpace M] (f : PartialHomeomorph M H) {I : ModelWithCorners π•œ E H} {t : Set M} {x : M} (h : x ∈ f.source) (ht : t ∈ nhds x) :
                                          ↑(f.extend I).symm ⁻¹' t ∈ nhds (↑(f.extend I) x)
                                          theorem PartialHomeomorph.extend_preimage_inter_eq {π•œ : Type u_1} {E : Type u_2} {M : Type u_3} {H : Type u_4} [NontriviallyNormedField π•œ] [NormedAddCommGroup E] [NormedSpace π•œ E] [TopologicalSpace H] [TopologicalSpace M] (f : PartialHomeomorph M H) {I : ModelWithCorners π•œ E H} {s t : Set M} :
                                          ↑(f.extend I).symm ⁻¹' (s ∩ t) ∩ Set.range ↑I = ↑(f.extend I).symm ⁻¹' s ∩ Set.range ↑I ∩ ↑(f.extend I).symm ⁻¹' t

                                          Technical lemma to rewrite suitably the preimage of an intersection under an extended chart, to bring it into a convenient form to apply derivative lemmas.

                                          theorem PartialHomeomorph.extend_symm_preimage_inter_range_eventuallyEq_aux {π•œ : Type u_1} {E : Type u_2} {M : Type u_3} {H : Type u_4} [NontriviallyNormedField π•œ] [NormedAddCommGroup E] [NormedSpace π•œ E] [TopologicalSpace H] [TopologicalSpace M] (f : PartialHomeomorph M H) {I : ModelWithCorners π•œ E H} {s : Set M} {x : M} (hx : x ∈ f.source) :
                                          ↑(f.extend I).symm ⁻¹' s ∩ Set.range ↑I =αΆ [nhds (↑(f.extend I) x)] (f.extend I).target ∩ ↑(f.extend I).symm ⁻¹' s
                                          theorem PartialHomeomorph.extend_symm_preimage_inter_range_eventuallyEq {π•œ : Type u_1} {E : Type u_2} {M : Type u_3} {H : Type u_4} [NontriviallyNormedField π•œ] [NormedAddCommGroup E] [NormedSpace π•œ E] [TopologicalSpace H] [TopologicalSpace M] (f : PartialHomeomorph M H) {I : ModelWithCorners π•œ E H} {s : Set M} {x : M} (hs : s βŠ† f.source) (hx : x ∈ f.source) :
                                          ↑(f.extend I).symm ⁻¹' s ∩ Set.range ↑I =αΆ [nhds (↑(f.extend I) x)] ↑(f.extend I) '' s

                                          We use the name extend_coord_change for (f'.extend I).symm ≫ f.extend I.

                                          theorem PartialHomeomorph.extend_coord_change_source {π•œ : Type u_1} {E : Type u_2} {M : Type u_3} {H : Type u_4} [NontriviallyNormedField π•œ] [NormedAddCommGroup E] [NormedSpace π•œ E] [TopologicalSpace H] [TopologicalSpace M] (f f' : PartialHomeomorph M H) {I : ModelWithCorners π•œ E H} :
                                          ((f.extend I).symm.trans (f'.extend I)).source = ↑I '' (f.symm.trans f').source
                                          theorem PartialHomeomorph.extend_image_source_inter {π•œ : Type u_1} {E : Type u_2} {M : Type u_3} {H : Type u_4} [NontriviallyNormedField π•œ] [NormedAddCommGroup E] [NormedSpace π•œ E] [TopologicalSpace H] [TopologicalSpace M] (f f' : PartialHomeomorph M H) {I : ModelWithCorners π•œ E H} :
                                          ↑(f.extend I) '' (f.source ∩ f'.source) = ((f.extend I).symm.trans (f'.extend I)).source
                                          theorem PartialHomeomorph.extend_coord_change_source_mem_nhdsWithin {π•œ : Type u_1} {E : Type u_2} {M : Type u_3} {H : Type u_4} [NontriviallyNormedField π•œ] [NormedAddCommGroup E] [NormedSpace π•œ E] [TopologicalSpace H] [TopologicalSpace M] (f f' : PartialHomeomorph M H) {I : ModelWithCorners π•œ E H} {x : E} (hx : x ∈ ((f.extend I).symm.trans (f'.extend I)).source) :
                                          ((f.extend I).symm.trans (f'.extend I)).source ∈ nhdsWithin x (Set.range ↑I)
                                          theorem PartialHomeomorph.extend_coord_change_source_mem_nhdsWithin' {π•œ : Type u_1} {E : Type u_2} {M : Type u_3} {H : Type u_4} [NontriviallyNormedField π•œ] [NormedAddCommGroup E] [NormedSpace π•œ E] [TopologicalSpace H] [TopologicalSpace M] (f f' : PartialHomeomorph M H) {I : ModelWithCorners π•œ E H} {x : M} (hxf : x ∈ f.source) (hxf' : x ∈ f'.source) :
                                          ((f.extend I).symm.trans (f'.extend I)).source ∈ nhdsWithin (↑(f.extend I) x) (Set.range ↑I)
                                          theorem PartialHomeomorph.contDiffOn_extend_coord_change {π•œ : Type u_1} {E : Type u_2} {M : Type u_3} {H : Type u_4} [NontriviallyNormedField π•œ] [NormedAddCommGroup E] [NormedSpace π•œ E] [TopologicalSpace H] [TopologicalSpace M] {n : WithTop β„•βˆž} {f f' : PartialHomeomorph M H} {I : ModelWithCorners π•œ E H} [ChartedSpace H M] (hf : f ∈ IsManifold.maximalAtlas I n M) (hf' : f' ∈ IsManifold.maximalAtlas I n M) :
                                          ContDiffOn π•œ n (↑(f.extend I) ∘ ↑(f'.extend I).symm) ((f'.extend I).symm.trans (f.extend I)).source
                                          theorem PartialHomeomorph.contDiffWithinAt_extend_coord_change {π•œ : Type u_1} {E : Type u_2} {M : Type u_3} {H : Type u_4} [NontriviallyNormedField π•œ] [NormedAddCommGroup E] [NormedSpace π•œ E] [TopologicalSpace H] [TopologicalSpace M] {n : WithTop β„•βˆž} {f f' : PartialHomeomorph M H} {I : ModelWithCorners π•œ E H} [ChartedSpace H M] (hf : f ∈ IsManifold.maximalAtlas I n M) (hf' : f' ∈ IsManifold.maximalAtlas I n M) {x : E} (hx : x ∈ ((f'.extend I).symm.trans (f.extend I)).source) :
                                          ContDiffWithinAt π•œ n (↑(f.extend I) ∘ ↑(f'.extend I).symm) (Set.range ↑I) x
                                          theorem PartialHomeomorph.contDiffWithinAt_extend_coord_change' {π•œ : Type u_1} {E : Type u_2} {M : Type u_3} {H : Type u_4} [NontriviallyNormedField π•œ] [NormedAddCommGroup E] [NormedSpace π•œ E] [TopologicalSpace H] [TopologicalSpace M] {n : WithTop β„•βˆž} {f f' : PartialHomeomorph M H} {I : ModelWithCorners π•œ E H} [ChartedSpace H M] (hf : f ∈ IsManifold.maximalAtlas I n M) (hf' : f' ∈ IsManifold.maximalAtlas I n M) {x : M} (hxf : x ∈ f.source) (hxf' : x ∈ f'.source) :
                                          ContDiffWithinAt π•œ n (↑(f.extend I) ∘ ↑(f'.extend I).symm) (Set.range ↑I) (↑(f'.extend I) x)
                                          def extChartAt {π•œ : Type u_1} {E : Type u_2} {M : Type u_3} {H : Type u_4} [NontriviallyNormedField π•œ] [NormedAddCommGroup E] [NormedSpace π•œ E] [TopologicalSpace H] [TopologicalSpace M] (I : ModelWithCorners π•œ E H) [ChartedSpace H M] (x : M) :

                                          The preferred extended chart on a manifold with corners around a point x, from a neighborhood of x to the model vector space.

                                          Equations
                                          Instances For
                                            theorem extChartAt_coe {π•œ : Type u_1} {E : Type u_2} {M : Type u_3} {H : Type u_4} [NontriviallyNormedField π•œ] [NormedAddCommGroup E] [NormedSpace π•œ E] [TopologicalSpace H] [TopologicalSpace M] {I : ModelWithCorners π•œ E H} [ChartedSpace H M] (x : M) :
                                            ↑(extChartAt I x) = ↑I ∘ ↑(chartAt H x)
                                            theorem extChartAt_coe_symm {π•œ : Type u_1} {E : Type u_2} {M : Type u_3} {H : Type u_4} [NontriviallyNormedField π•œ] [NormedAddCommGroup E] [NormedSpace π•œ E] [TopologicalSpace H] [TopologicalSpace M] {I : ModelWithCorners π•œ E H} [ChartedSpace H M] (x : M) :
                                            ↑(extChartAt I x).symm = ↑(chartAt H x).symm ∘ ↑I.symm
                                            theorem extChartAt_source {π•œ : Type u_1} {E : Type u_2} {M : Type u_3} {H : Type u_4} [NontriviallyNormedField π•œ] [NormedAddCommGroup E] [NormedSpace π•œ E] [TopologicalSpace H] [TopologicalSpace M] (I : ModelWithCorners π•œ E H) [ChartedSpace H M] (x : M) :
                                            (extChartAt I x).source = (chartAt H x).source
                                            theorem isOpen_extChartAt_source {π•œ : Type u_1} {E : Type u_2} {M : Type u_3} {H : Type u_4} [NontriviallyNormedField π•œ] [NormedAddCommGroup E] [NormedSpace π•œ E] [TopologicalSpace H] [TopologicalSpace M] {I : ModelWithCorners π•œ E H} [ChartedSpace H M] (x : M) :
                                            IsOpen (extChartAt I x).source
                                            theorem mem_extChartAt_source {π•œ : Type u_1} {E : Type u_2} {M : Type u_3} {H : Type u_4} [NontriviallyNormedField π•œ] [NormedAddCommGroup E] [NormedSpace π•œ E] [TopologicalSpace H] [TopologicalSpace M] {I : ModelWithCorners π•œ E H} [ChartedSpace H M] (x : M) :
                                            x ∈ (extChartAt I x).source
                                            theorem mem_extChartAt_target {π•œ : Type u_1} {E : Type u_2} {M : Type u_3} {H : Type u_4} [NontriviallyNormedField π•œ] [NormedAddCommGroup E] [NormedSpace π•œ E] [TopologicalSpace H] [TopologicalSpace M] {I : ModelWithCorners π•œ E H} [ChartedSpace H M] (x : M) :
                                            ↑(extChartAt I x) x ∈ (extChartAt I x).target
                                            theorem extChartAt_target {π•œ : Type u_1} {E : Type u_2} {M : Type u_3} {H : Type u_4} [NontriviallyNormedField π•œ] [NormedAddCommGroup E] [NormedSpace π•œ E] [TopologicalSpace H] [TopologicalSpace M] (I : ModelWithCorners π•œ E H) [ChartedSpace H M] (x : M) :
                                            (extChartAt I x).target = ↑I.symm ⁻¹' (chartAt H x).target ∩ Set.range ↑I
                                            theorem uniqueDiffOn_extChartAt_target {π•œ : Type u_1} {E : Type u_2} {M : Type u_3} {H : Type u_4} [NontriviallyNormedField π•œ] [NormedAddCommGroup E] [NormedSpace π•œ E] [TopologicalSpace H] [TopologicalSpace M] {I : ModelWithCorners π•œ E H} [ChartedSpace H M] (x : M) :
                                            UniqueDiffOn π•œ (extChartAt I x).target
                                            theorem uniqueDiffWithinAt_extChartAt_target {π•œ : Type u_1} {E : Type u_2} {M : Type u_3} {H : Type u_4} [NontriviallyNormedField π•œ] [NormedAddCommGroup E] [NormedSpace π•œ E] [TopologicalSpace H] [TopologicalSpace M] {I : ModelWithCorners π•œ E H} [ChartedSpace H M] (x : M) :
                                            UniqueDiffWithinAt π•œ (extChartAt I x).target (↑(extChartAt I x) x)
                                            theorem extChartAt_to_inv {π•œ : Type u_1} {E : Type u_2} {M : Type u_3} {H : Type u_4} [NontriviallyNormedField π•œ] [NormedAddCommGroup E] [NormedSpace π•œ E] [TopologicalSpace H] [TopologicalSpace M] {I : ModelWithCorners π•œ E H} [ChartedSpace H M] (x : M) :
                                            ↑(extChartAt I x).symm (↑(extChartAt I x) x) = x
                                            theorem mapsTo_extChartAt {π•œ : Type u_1} {E : Type u_2} {M : Type u_3} {H : Type u_4} [NontriviallyNormedField π•œ] [NormedAddCommGroup E] [NormedSpace π•œ E] [TopologicalSpace H] [TopologicalSpace M] {I : ModelWithCorners π•œ E H} {s : Set M} [ChartedSpace H M] {x : M} (hs : s βŠ† (chartAt H x).source) :
                                            Set.MapsTo (↑(extChartAt I x)) s (↑(extChartAt I x).symm ⁻¹' s ∩ Set.range ↑I)
                                            theorem extChartAt_source_mem_nhds' {π•œ : Type u_1} {E : Type u_2} {M : Type u_3} {H : Type u_4} [NontriviallyNormedField π•œ] [NormedAddCommGroup E] [NormedSpace π•œ E] [TopologicalSpace H] [TopologicalSpace M] {I : ModelWithCorners π•œ E H} [ChartedSpace H M] {x x' : M} (h : x' ∈ (extChartAt I x).source) :
                                            (extChartAt I x).source ∈ nhds x'
                                            theorem extChartAt_source_mem_nhds {π•œ : Type u_1} {E : Type u_2} {M : Type u_3} {H : Type u_4} [NontriviallyNormedField π•œ] [NormedAddCommGroup E] [NormedSpace π•œ E] [TopologicalSpace H] [TopologicalSpace M] {I : ModelWithCorners π•œ E H} [ChartedSpace H M] (x : M) :
                                            (extChartAt I x).source ∈ nhds x
                                            theorem extChartAt_source_mem_nhdsWithin' {π•œ : Type u_1} {E : Type u_2} {M : Type u_3} {H : Type u_4} [NontriviallyNormedField π•œ] [NormedAddCommGroup E] [NormedSpace π•œ E] [TopologicalSpace H] [TopologicalSpace M] {I : ModelWithCorners π•œ E H} {s : Set M} [ChartedSpace H M] {x x' : M} (h : x' ∈ (extChartAt I x).source) :
                                            (extChartAt I x).source ∈ nhdsWithin x' s
                                            theorem extChartAt_source_mem_nhdsWithin {π•œ : Type u_1} {E : Type u_2} {M : Type u_3} {H : Type u_4} [NontriviallyNormedField π•œ] [NormedAddCommGroup E] [NormedSpace π•œ E] [TopologicalSpace H] [TopologicalSpace M] {I : ModelWithCorners π•œ E H} {s : Set M} [ChartedSpace H M] (x : M) :
                                            (extChartAt I x).source ∈ nhdsWithin x s
                                            theorem continuousOn_extChartAt {π•œ : Type u_1} {E : Type u_2} {M : Type u_3} {H : Type u_4} [NontriviallyNormedField π•œ] [NormedAddCommGroup E] [NormedSpace π•œ E] [TopologicalSpace H] [TopologicalSpace M] {I : ModelWithCorners π•œ E H} [ChartedSpace H M] (x : M) :
                                            ContinuousOn (↑(extChartAt I x)) (extChartAt I x).source
                                            theorem continuousAt_extChartAt' {π•œ : Type u_1} {E : Type u_2} {M : Type u_3} {H : Type u_4} [NontriviallyNormedField π•œ] [NormedAddCommGroup E] [NormedSpace π•œ E] [TopologicalSpace H] [TopologicalSpace M] {I : ModelWithCorners π•œ E H} [ChartedSpace H M] {x x' : M} (h : x' ∈ (extChartAt I x).source) :
                                            ContinuousAt (↑(extChartAt I x)) x'
                                            theorem continuousAt_extChartAt {π•œ : Type u_1} {E : Type u_2} {M : Type u_3} {H : Type u_4} [NontriviallyNormedField π•œ] [NormedAddCommGroup E] [NormedSpace π•œ E] [TopologicalSpace H] [TopologicalSpace M] {I : ModelWithCorners π•œ E H} [ChartedSpace H M] (x : M) :
                                            ContinuousAt (↑(extChartAt I x)) x
                                            theorem map_extChartAt_nhds' {π•œ : Type u_1} {E : Type u_2} {M : Type u_3} {H : Type u_4} [NontriviallyNormedField π•œ] [NormedAddCommGroup E] [NormedSpace π•œ E] [TopologicalSpace H] [TopologicalSpace M] {I : ModelWithCorners π•œ E H} [ChartedSpace H M] {x y : M} (hy : y ∈ (extChartAt I x).source) :
                                            Filter.map (↑(extChartAt I x)) (nhds y) = nhdsWithin (↑(extChartAt I x) y) (Set.range ↑I)
                                            theorem map_extChartAt_nhds {π•œ : Type u_1} {E : Type u_2} {M : Type u_3} {H : Type u_4} [NontriviallyNormedField π•œ] [NormedAddCommGroup E] [NormedSpace π•œ E] [TopologicalSpace H] [TopologicalSpace M] {I : ModelWithCorners π•œ E H} [ChartedSpace H M] (x : M) :
                                            Filter.map (↑(extChartAt I x)) (nhds x) = nhdsWithin (↑(extChartAt I x) x) (Set.range ↑I)
                                            theorem map_extChartAt_nhds_of_boundaryless {π•œ : Type u_1} {E : Type u_2} {M : Type u_3} {H : Type u_4} [NontriviallyNormedField π•œ] [NormedAddCommGroup E] [NormedSpace π•œ E] [TopologicalSpace H] [TopologicalSpace M] {I : ModelWithCorners π•œ E H} [ChartedSpace H M] [I.Boundaryless] (x : M) :
                                            Filter.map (↑(extChartAt I x)) (nhds x) = nhds (↑(extChartAt I x) x)
                                            theorem extChartAt_image_nhd_mem_nhds_of_mem_interior_range {π•œ : Type u_1} {E : Type u_2} {M : Type u_3} {H : Type u_4} [NontriviallyNormedField π•œ] [NormedAddCommGroup E] [NormedSpace π•œ E] [TopologicalSpace H] [TopologicalSpace M] {I : ModelWithCorners π•œ E H} [ChartedSpace H M] {x y : M} (hx : y ∈ (extChartAt I x).source) (h'x : ↑(extChartAt I x) y ∈ interior (Set.range ↑I)) {s : Set M} (h : s ∈ nhds y) :
                                            ↑(extChartAt I x) '' s ∈ nhds (↑(extChartAt I x) y)
                                            theorem extChartAt_image_nhd_mem_nhds_of_boundaryless {π•œ : Type u_1} {E : Type u_2} {M : Type u_3} {H : Type u_4} [NontriviallyNormedField π•œ] [NormedAddCommGroup E] [NormedSpace π•œ E] [TopologicalSpace H] [TopologicalSpace M] {I : ModelWithCorners π•œ E H} {s : Set M} [ChartedSpace H M] [I.Boundaryless] {x : M} (hx : s ∈ nhds x) :
                                            ↑(extChartAt I x) '' s ∈ nhds (↑(extChartAt I x) x)
                                            theorem extChartAt_target_mem_nhdsWithin' {π•œ : Type u_1} {E : Type u_2} {M : Type u_3} {H : Type u_4} [NontriviallyNormedField π•œ] [NormedAddCommGroup E] [NormedSpace π•œ E] [TopologicalSpace H] [TopologicalSpace M] {I : ModelWithCorners π•œ E H} [ChartedSpace H M] {x y : M} (hy : y ∈ (extChartAt I x).source) :
                                            (extChartAt I x).target ∈ nhdsWithin (↑(extChartAt I x) y) (Set.range ↑I)
                                            theorem extChartAt_target_mem_nhdsWithin {π•œ : Type u_1} {E : Type u_2} {M : Type u_3} {H : Type u_4} [NontriviallyNormedField π•œ] [NormedAddCommGroup E] [NormedSpace π•œ E] [TopologicalSpace H] [TopologicalSpace M] {I : ModelWithCorners π•œ E H} [ChartedSpace H M] (x : M) :
                                            (extChartAt I x).target ∈ nhdsWithin (↑(extChartAt I x) x) (Set.range ↑I)
                                            theorem extChartAt_target_mem_nhdsWithin_of_mem {π•œ : Type u_1} {E : Type u_2} {M : Type u_3} {H : Type u_4} [NontriviallyNormedField π•œ] [NormedAddCommGroup E] [NormedSpace π•œ E] [TopologicalSpace H] [TopologicalSpace M] {I : ModelWithCorners π•œ E H} [ChartedSpace H M] {x : M} {y : E} (hy : y ∈ (extChartAt I x).target) :
                                            (extChartAt I x).target ∈ nhdsWithin y (Set.range ↑I)
                                            theorem extChartAt_target_union_compl_range_mem_nhds_of_mem {π•œ : Type u_1} {E : Type u_2} {M : Type u_3} {H : Type u_4} [NontriviallyNormedField π•œ] [NormedAddCommGroup E] [NormedSpace π•œ E] [TopologicalSpace H] [TopologicalSpace M] {I : ModelWithCorners π•œ E H} [ChartedSpace H M] {y : E} {x : M} (hy : y ∈ (extChartAt I x).target) :
                                            (extChartAt I x).target βˆͺ (Set.range ↑I)ᢜ ∈ nhds y
                                            @[deprecated extChartAt_target_union_compl_range_mem_nhds_of_mem (since := "2024-11-27")]
                                            theorem extChartAt_target_union_comp_range_mem_nhds_of_mem {π•œ : Type u_1} {E : Type u_2} {M : Type u_3} {H : Type u_4} [NontriviallyNormedField π•œ] [NormedAddCommGroup E] [NormedSpace π•œ E] [TopologicalSpace H] [TopologicalSpace M] {I : ModelWithCorners π•œ E H} [ChartedSpace H M] {y : E} {x : M} (hy : y ∈ (extChartAt I x).target) :
                                            (extChartAt I x).target βˆͺ (Set.range ↑I)ᢜ ∈ nhds y

                                            Alias of extChartAt_target_union_compl_range_mem_nhds_of_mem.

                                            theorem isOpen_extChartAt_target {π•œ : Type u_1} {E : Type u_2} {M : Type u_3} {H : Type u_4} [NontriviallyNormedField π•œ] [NormedAddCommGroup E] [NormedSpace π•œ E] [TopologicalSpace H] [TopologicalSpace M] {I : ModelWithCorners π•œ E H} [ChartedSpace H M] [I.Boundaryless] (x : M) :
                                            IsOpen (extChartAt I x).target

                                            If we're boundaryless, extChartAt has open target

                                            theorem extChartAt_target_mem_nhds {π•œ : Type u_1} {E : Type u_2} {M : Type u_3} {H : Type u_4} [NontriviallyNormedField π•œ] [NormedAddCommGroup E] [NormedSpace π•œ E] [TopologicalSpace H] [TopologicalSpace M] {I : ModelWithCorners π•œ E H} [ChartedSpace H M] [I.Boundaryless] (x : M) :
                                            (extChartAt I x).target ∈ nhds (↑(extChartAt I x) x)

                                            If we're boundaryless, (extChartAt I x).target is a neighborhood of the key point

                                            theorem extChartAt_target_mem_nhds' {π•œ : Type u_1} {E : Type u_2} {M : Type u_3} {H : Type u_4} [NontriviallyNormedField π•œ] [NormedAddCommGroup E] [NormedSpace π•œ E] [TopologicalSpace H] [TopologicalSpace M] {I : ModelWithCorners π•œ E H} [ChartedSpace H M] [I.Boundaryless] {x : M} {y : E} (m : y ∈ (extChartAt I x).target) :
                                            (extChartAt I x).target ∈ nhds y

                                            If we're boundaryless, (extChartAt I x).target is a neighborhood of any of its points

                                            theorem extChartAt_target_subset_range {π•œ : Type u_1} {E : Type u_2} {M : Type u_3} {H : Type u_4} [NontriviallyNormedField π•œ] [NormedAddCommGroup E] [NormedSpace π•œ E] [TopologicalSpace H] [TopologicalSpace M] {I : ModelWithCorners π•œ E H} [ChartedSpace H M] (x : M) :
                                            (extChartAt I x).target βŠ† Set.range ↑I
                                            theorem nhdsWithin_extChartAt_target_eq' {π•œ : Type u_1} {E : Type u_2} {M : Type u_3} {H : Type u_4} [NontriviallyNormedField π•œ] [NormedAddCommGroup E] [NormedSpace π•œ E] [TopologicalSpace H] [TopologicalSpace M] {I : ModelWithCorners π•œ E H} [ChartedSpace H M] {x y : M} (hy : y ∈ (extChartAt I x).source) :
                                            nhdsWithin (↑(extChartAt I x) y) (extChartAt I x).target = nhdsWithin (↑(extChartAt I x) y) (Set.range ↑I)

                                            Around the image of a point in the source, the neighborhoods are the same within (extChartAt I x).target and within range I.

                                            theorem nhdsWithin_extChartAt_target_eq_of_mem {π•œ : Type u_1} {E : Type u_2} {M : Type u_3} {H : Type u_4} [NontriviallyNormedField π•œ] [NormedAddCommGroup E] [NormedSpace π•œ E] [TopologicalSpace H] [TopologicalSpace M] {I : ModelWithCorners π•œ E H} [ChartedSpace H M] {x : M} {z : E} (hz : z ∈ (extChartAt I x).target) :
                                            nhdsWithin z (extChartAt I x).target = nhdsWithin z (Set.range ↑I)

                                            Around a point in the target, the neighborhoods are the same within (extChartAt I x).target and within range I.

                                            theorem nhdsWithin_extChartAt_target_eq {π•œ : Type u_1} {E : Type u_2} {M : Type u_3} {H : Type u_4} [NontriviallyNormedField π•œ] [NormedAddCommGroup E] [NormedSpace π•œ E] [TopologicalSpace H] [TopologicalSpace M] {I : ModelWithCorners π•œ E H} [ChartedSpace H M] (x : M) :
                                            nhdsWithin (↑(extChartAt I x) x) (extChartAt I x).target = nhdsWithin (↑(extChartAt I x) x) (Set.range ↑I)

                                            Around the image of the base point, the neighborhoods are the same within (extChartAt I x).target and within range I.

                                            theorem extChartAt_target_eventuallyEq' {π•œ : Type u_1} {E : Type u_2} {M : Type u_3} {H : Type u_4} [NontriviallyNormedField π•œ] [NormedAddCommGroup E] [NormedSpace π•œ E] [TopologicalSpace H] [TopologicalSpace M] {I : ModelWithCorners π•œ E H} [ChartedSpace H M] {x y : M} (hy : y ∈ (extChartAt I x).source) :
                                            (extChartAt I x).target =αΆ [nhds (↑(extChartAt I x) y)] Set.range ↑I

                                            Around the image of a point in the source, (extChartAt I x).target and range I coincide locally.

                                            theorem extChartAt_target_eventuallyEq_of_mem {π•œ : Type u_1} {E : Type u_2} {M : Type u_3} {H : Type u_4} [NontriviallyNormedField π•œ] [NormedAddCommGroup E] [NormedSpace π•œ E] [TopologicalSpace H] [TopologicalSpace M] {I : ModelWithCorners π•œ E H} [ChartedSpace H M] {x : M} {z : E} (hz : z ∈ (extChartAt I x).target) :
                                            (extChartAt I x).target =αΆ [nhds z] Set.range ↑I

                                            Around a point in the target, (extChartAt I x).target and range I coincide locally.

                                            theorem extChartAt_target_eventuallyEq {π•œ : Type u_1} {E : Type u_2} {M : Type u_3} {H : Type u_4} [NontriviallyNormedField π•œ] [NormedAddCommGroup E] [NormedSpace π•œ E] [TopologicalSpace H] [TopologicalSpace M] {I : ModelWithCorners π•œ E H} [ChartedSpace H M] {x : M} :
                                            (extChartAt I x).target =αΆ [nhds (↑(extChartAt I x) x)] Set.range ↑I

                                            Around the image of the base point, (extChartAt I x).target and range I coincide locally.

                                            theorem continuousAt_extChartAt_symm'' {π•œ : Type u_1} {E : Type u_2} {M : Type u_3} {H : Type u_4} [NontriviallyNormedField π•œ] [NormedAddCommGroup E] [NormedSpace π•œ E] [TopologicalSpace H] [TopologicalSpace M] {I : ModelWithCorners π•œ E H} [ChartedSpace H M] {x : M} {y : E} (h : y ∈ (extChartAt I x).target) :
                                            ContinuousAt (↑(extChartAt I x).symm) y
                                            theorem continuousAt_extChartAt_symm' {π•œ : Type u_1} {E : Type u_2} {M : Type u_3} {H : Type u_4} [NontriviallyNormedField π•œ] [NormedAddCommGroup E] [NormedSpace π•œ E] [TopologicalSpace H] [TopologicalSpace M] {I : ModelWithCorners π•œ E H} [ChartedSpace H M] {x x' : M} (h : x' ∈ (extChartAt I x).source) :
                                            ContinuousAt (↑(extChartAt I x).symm) (↑(extChartAt I x) x')
                                            theorem continuousAt_extChartAt_symm {π•œ : Type u_1} {E : Type u_2} {M : Type u_3} {H : Type u_4} [NontriviallyNormedField π•œ] [NormedAddCommGroup E] [NormedSpace π•œ E] [TopologicalSpace H] [TopologicalSpace M] {I : ModelWithCorners π•œ E H} [ChartedSpace H M] (x : M) :
                                            ContinuousAt (↑(extChartAt I x).symm) (↑(extChartAt I x) x)
                                            theorem continuousOn_extChartAt_symm {π•œ : Type u_1} {E : Type u_2} {M : Type u_3} {H : Type u_4} [NontriviallyNormedField π•œ] [NormedAddCommGroup E] [NormedSpace π•œ E] [TopologicalSpace H] [TopologicalSpace M] {I : ModelWithCorners π•œ E H} [ChartedSpace H M] (x : M) :
                                            ContinuousOn (↑(extChartAt I x).symm) (extChartAt I x).target
                                            theorem extChartAt_target_subset_closure_interior {π•œ : Type u_1} {E : Type u_2} {M : Type u_3} {H : Type u_4} [NontriviallyNormedField π•œ] [NormedAddCommGroup E] [NormedSpace π•œ E] [TopologicalSpace H] [TopologicalSpace M] {I : ModelWithCorners π•œ E H} [ChartedSpace H M] {x : M} :
                                            (extChartAt I x).target βŠ† closure (interior (extChartAt I x).target)
                                            theorem interior_extChartAt_target_nonempty {π•œ : Type u_1} {E : Type u_2} {M : Type u_3} {H : Type u_4} [NontriviallyNormedField π•œ] [NormedAddCommGroup E] [NormedSpace π•œ E] [TopologicalSpace H] [TopologicalSpace M] (I : ModelWithCorners π•œ E H) [ChartedSpace H M] (x : M) :
                                            (interior (extChartAt I x).target).Nonempty
                                            theorem extChartAt_mem_closure_interior {π•œ : Type u_1} {E : Type u_2} {M : Type u_3} {H : Type u_4} [NontriviallyNormedField π•œ] [NormedAddCommGroup E] [NormedSpace π•œ E] [TopologicalSpace H] [TopologicalSpace M] {I : ModelWithCorners π•œ E H} {s : Set M} [ChartedSpace H M] {xβ‚€ x : M} (hx : x ∈ closure (interior s)) (h'x : x ∈ (extChartAt I xβ‚€).source) :
                                            ↑(extChartAt I xβ‚€) x ∈ closure (interior (↑(extChartAt I xβ‚€).symm ⁻¹' s ∩ (extChartAt I xβ‚€).target))
                                            theorem isOpen_extChartAt_preimage' {π•œ : Type u_1} {E : Type u_2} {M : Type u_3} {H : Type u_4} [NontriviallyNormedField π•œ] [NormedAddCommGroup E] [NormedSpace π•œ E] [TopologicalSpace H] [TopologicalSpace M] {I : ModelWithCorners π•œ E H} [ChartedSpace H M] (x : M) {s : Set E} (hs : IsOpen s) :
                                            IsOpen ((extChartAt I x).source ∩ ↑(extChartAt I x) ⁻¹' s)
                                            theorem isOpen_extChartAt_preimage {π•œ : Type u_1} {E : Type u_2} {M : Type u_3} {H : Type u_4} [NontriviallyNormedField π•œ] [NormedAddCommGroup E] [NormedSpace π•œ E] [TopologicalSpace H] [TopologicalSpace M] {I : ModelWithCorners π•œ E H} [ChartedSpace H M] (x : M) {s : Set E} (hs : IsOpen s) :
                                            IsOpen ((chartAt H x).source ∩ ↑(extChartAt I x) ⁻¹' s)
                                            theorem map_extChartAt_nhdsWithin_eq_image' {π•œ : Type u_1} {E : Type u_2} {M : Type u_3} {H : Type u_4} [NontriviallyNormedField π•œ] [NormedAddCommGroup E] [NormedSpace π•œ E] [TopologicalSpace H] [TopologicalSpace M] {I : ModelWithCorners π•œ E H} {s : Set M} [ChartedSpace H M] {x y : M} (hy : y ∈ (extChartAt I x).source) :
                                            Filter.map (↑(extChartAt I x)) (nhdsWithin y s) = nhdsWithin (↑(extChartAt I x) y) (↑(extChartAt I x) '' ((extChartAt I x).source ∩ s))
                                            theorem map_extChartAt_nhdsWithin_eq_image {π•œ : Type u_1} {E : Type u_2} {M : Type u_3} {H : Type u_4} [NontriviallyNormedField π•œ] [NormedAddCommGroup E] [NormedSpace π•œ E] [TopologicalSpace H] [TopologicalSpace M] {I : ModelWithCorners π•œ E H} {s : Set M} [ChartedSpace H M] (x : M) :
                                            Filter.map (↑(extChartAt I x)) (nhdsWithin x s) = nhdsWithin (↑(extChartAt I x) x) (↑(extChartAt I x) '' ((extChartAt I x).source ∩ s))
                                            theorem map_extChartAt_nhdsWithin' {π•œ : Type u_1} {E : Type u_2} {M : Type u_3} {H : Type u_4} [NontriviallyNormedField π•œ] [NormedAddCommGroup E] [NormedSpace π•œ E] [TopologicalSpace H] [TopologicalSpace M] {I : ModelWithCorners π•œ E H} {s : Set M} [ChartedSpace H M] {x y : M} (hy : y ∈ (extChartAt I x).source) :
                                            Filter.map (↑(extChartAt I x)) (nhdsWithin y s) = nhdsWithin (↑(extChartAt I x) y) (↑(extChartAt I x).symm ⁻¹' s ∩ Set.range ↑I)
                                            theorem map_extChartAt_nhdsWithin {π•œ : Type u_1} {E : Type u_2} {M : Type u_3} {H : Type u_4} [NontriviallyNormedField π•œ] [NormedAddCommGroup E] [NormedSpace π•œ E] [TopologicalSpace H] [TopologicalSpace M] {I : ModelWithCorners π•œ E H} {s : Set M} [ChartedSpace H M] (x : M) :
                                            Filter.map (↑(extChartAt I x)) (nhdsWithin x s) = nhdsWithin (↑(extChartAt I x) x) (↑(extChartAt I x).symm ⁻¹' s ∩ Set.range ↑I)
                                            theorem map_extChartAt_symm_nhdsWithin' {π•œ : Type u_1} {E : Type u_2} {M : Type u_3} {H : Type u_4} [NontriviallyNormedField π•œ] [NormedAddCommGroup E] [NormedSpace π•œ E] [TopologicalSpace H] [TopologicalSpace M] {I : ModelWithCorners π•œ E H} {s : Set M} [ChartedSpace H M] {x y : M} (hy : y ∈ (extChartAt I x).source) :
                                            Filter.map (↑(extChartAt I x).symm) (nhdsWithin (↑(extChartAt I x) y) (↑(extChartAt I x).symm ⁻¹' s ∩ Set.range ↑I)) = nhdsWithin y s
                                            theorem map_extChartAt_symm_nhdsWithin_range' {π•œ : Type u_1} {E : Type u_2} {M : Type u_3} {H : Type u_4} [NontriviallyNormedField π•œ] [NormedAddCommGroup E] [NormedSpace π•œ E] [TopologicalSpace H] [TopologicalSpace M] {I : ModelWithCorners π•œ E H} [ChartedSpace H M] {x y : M} (hy : y ∈ (extChartAt I x).source) :
                                            Filter.map (↑(extChartAt I x).symm) (nhdsWithin (↑(extChartAt I x) y) (Set.range ↑I)) = nhds y
                                            theorem map_extChartAt_symm_nhdsWithin {π•œ : Type u_1} {E : Type u_2} {M : Type u_3} {H : Type u_4} [NontriviallyNormedField π•œ] [NormedAddCommGroup E] [NormedSpace π•œ E] [TopologicalSpace H] [TopologicalSpace M] {I : ModelWithCorners π•œ E H} {s : Set M} [ChartedSpace H M] (x : M) :
                                            Filter.map (↑(extChartAt I x).symm) (nhdsWithin (↑(extChartAt I x) x) (↑(extChartAt I x).symm ⁻¹' s ∩ Set.range ↑I)) = nhdsWithin x s
                                            theorem map_extChartAt_symm_nhdsWithin_range {π•œ : Type u_1} {E : Type u_2} {M : Type u_3} {H : Type u_4} [NontriviallyNormedField π•œ] [NormedAddCommGroup E] [NormedSpace π•œ E] [TopologicalSpace H] [TopologicalSpace M] {I : ModelWithCorners π•œ E H} [ChartedSpace H M] (x : M) :
                                            Filter.map (↑(extChartAt I x).symm) (nhdsWithin (↑(extChartAt I x) x) (Set.range ↑I)) = nhds x
                                            theorem extChartAt_preimage_mem_nhdsWithin' {π•œ : Type u_1} {E : Type u_2} {M : Type u_3} {H : Type u_4} [NontriviallyNormedField π•œ] [NormedAddCommGroup E] [NormedSpace π•œ E] [TopologicalSpace H] [TopologicalSpace M] {I : ModelWithCorners π•œ E H} {s t : Set M} [ChartedSpace H M] {x x' : M} (h : x' ∈ (extChartAt I x).source) (ht : t ∈ nhdsWithin x' s) :
                                            ↑(extChartAt I x).symm ⁻¹' t ∈ nhdsWithin (↑(extChartAt I x) x') (↑(extChartAt I x).symm ⁻¹' s ∩ Set.range ↑I)

                                            Technical lemma ensuring that the preimage under an extended chart of a neighborhood of a point in the source is a neighborhood of the preimage, within a set.

                                            theorem extChartAt_preimage_mem_nhdsWithin {π•œ : Type u_1} {E : Type u_2} {M : Type u_3} {H : Type u_4} [NontriviallyNormedField π•œ] [NormedAddCommGroup E] [NormedSpace π•œ E] [TopologicalSpace H] [TopologicalSpace M] {I : ModelWithCorners π•œ E H} {s t : Set M} [ChartedSpace H M] {x : M} (ht : t ∈ nhdsWithin x s) :
                                            ↑(extChartAt I x).symm ⁻¹' t ∈ nhdsWithin (↑(extChartAt I x) x) (↑(extChartAt I x).symm ⁻¹' s ∩ Set.range ↑I)

                                            Technical lemma ensuring that the preimage under an extended chart of a neighborhood of the base point is a neighborhood of the preimage, within a set.

                                            theorem extChartAt_preimage_mem_nhds' {π•œ : Type u_1} {E : Type u_2} {M : Type u_3} {H : Type u_4} [NontriviallyNormedField π•œ] [NormedAddCommGroup E] [NormedSpace π•œ E] [TopologicalSpace H] [TopologicalSpace M] {I : ModelWithCorners π•œ E H} {t : Set M} [ChartedSpace H M] {x x' : M} (h : x' ∈ (extChartAt I x).source) (ht : t ∈ nhds x') :
                                            ↑(extChartAt I x).symm ⁻¹' t ∈ nhds (↑(extChartAt I x) x')
                                            theorem extChartAt_preimage_mem_nhds {π•œ : Type u_1} {E : Type u_2} {M : Type u_3} {H : Type u_4} [NontriviallyNormedField π•œ] [NormedAddCommGroup E] [NormedSpace π•œ E] [TopologicalSpace H] [TopologicalSpace M] {I : ModelWithCorners π•œ E H} {t : Set M} [ChartedSpace H M] {x : M} (ht : t ∈ nhds x) :
                                            ↑(extChartAt I x).symm ⁻¹' t ∈ nhds (↑(extChartAt I x) x)

                                            Technical lemma ensuring that the preimage under an extended chart of a neighborhood of a point is a neighborhood of the preimage.

                                            theorem extChartAt_preimage_inter_eq {π•œ : Type u_1} {E : Type u_2} {M : Type u_3} {H : Type u_4} [NontriviallyNormedField π•œ] [NormedAddCommGroup E] [NormedSpace π•œ E] [TopologicalSpace H] [TopologicalSpace M] {I : ModelWithCorners π•œ E H} {s t : Set M} [ChartedSpace H M] (x : M) :
                                            ↑(extChartAt I x).symm ⁻¹' (s ∩ t) ∩ Set.range ↑I = ↑(extChartAt I x).symm ⁻¹' s ∩ Set.range ↑I ∩ ↑(extChartAt I x).symm ⁻¹' t

                                            Technical lemma to rewrite suitably the preimage of an intersection under an extended chart, to bring it into a convenient form to apply derivative lemmas.

                                            theorem ContinuousWithinAt.nhdsWithin_extChartAt_symm_preimage_inter_range {π•œ : Type u_1} {E : Type u_2} {M : Type u_3} {H : Type u_4} {E' : Type u_5} {M' : Type u_6} {H' : Type u_7} [NontriviallyNormedField π•œ] [NormedAddCommGroup E] [NormedSpace π•œ E] [TopologicalSpace H] [TopologicalSpace M] {I : ModelWithCorners π•œ E H} [NormedAddCommGroup E'] [NormedSpace π•œ E'] [TopologicalSpace H'] [TopologicalSpace M'] {I' : ModelWithCorners π•œ E' H'} {s : Set M} [ChartedSpace H M] [ChartedSpace H' M'] {f : M β†’ M'} {x : M} (hc : ContinuousWithinAt f s x) :
                                            nhdsWithin (↑(extChartAt I x) x) (↑(extChartAt I x).symm ⁻¹' s ∩ Set.range ↑I) = nhdsWithin (↑(extChartAt I x) x) ((extChartAt I x).target ∩ ↑(extChartAt I x).symm ⁻¹' (s ∩ f ⁻¹' (extChartAt I' (f x)).source))
                                            theorem ContinuousWithinAt.extChartAt_symm_preimage_inter_range_eventuallyEq {π•œ : Type u_1} {E : Type u_2} {M : Type u_3} {H : Type u_4} {E' : Type u_5} {M' : Type u_6} {H' : Type u_7} [NontriviallyNormedField π•œ] [NormedAddCommGroup E] [NormedSpace π•œ E] [TopologicalSpace H] [TopologicalSpace M] {I : ModelWithCorners π•œ E H} [NormedAddCommGroup E'] [NormedSpace π•œ E'] [TopologicalSpace H'] [TopologicalSpace M'] {I' : ModelWithCorners π•œ E' H'} {s : Set M} [ChartedSpace H M] [ChartedSpace H' M'] {f : M β†’ M'} {x : M} (hc : ContinuousWithinAt f s x) :
                                            ↑(extChartAt I x).symm ⁻¹' s ∩ Set.range ↑I =αΆ [nhds (↑(extChartAt I x) x)] (extChartAt I x).target ∩ ↑(extChartAt I x).symm ⁻¹' (s ∩ f ⁻¹' (extChartAt I' (f x)).source)

                                            We use the name ext_coord_change for (extChartAt I x').symm ≫ extChartAt I x.

                                            theorem ext_coord_change_source {π•œ : Type u_1} {E : Type u_2} {M : Type u_3} {H : Type u_4} [NontriviallyNormedField π•œ] [NormedAddCommGroup E] [NormedSpace π•œ E] [TopologicalSpace H] [TopologicalSpace M] {I : ModelWithCorners π•œ E H} [ChartedSpace H M] (x x' : M) :
                                            ((extChartAt I x').symm.trans (extChartAt I x)).source = ↑I '' ((chartAt H x').symm.trans (chartAt H x)).source
                                            theorem contDiffOn_ext_coord_change {π•œ : Type u_1} {E : Type u_2} {M : Type u_3} {H : Type u_4} [NontriviallyNormedField π•œ] [NormedAddCommGroup E] [NormedSpace π•œ E] [TopologicalSpace H] [TopologicalSpace M] {n : WithTop β„•βˆž} {I : ModelWithCorners π•œ E H} [ChartedSpace H M] [IsManifold I n M] (x x' : M) :
                                            ContDiffOn π•œ n (↑(extChartAt I x) ∘ ↑(extChartAt I x').symm) ((extChartAt I x').symm.trans (extChartAt I x)).source
                                            theorem contDiffWithinAt_ext_coord_change {π•œ : Type u_1} {E : Type u_2} {M : Type u_3} {H : Type u_4} [NontriviallyNormedField π•œ] [NormedAddCommGroup E] [NormedSpace π•œ E] [TopologicalSpace H] [TopologicalSpace M] {n : WithTop β„•βˆž} {I : ModelWithCorners π•œ E H} [ChartedSpace H M] [IsManifold I n M] (x x' : M) {y : E} (hy : y ∈ ((extChartAt I x').symm.trans (extChartAt I x)).source) :
                                            ContDiffWithinAt π•œ n (↑(extChartAt I x) ∘ ↑(extChartAt I x').symm) (Set.range ↑I) y
                                            def writtenInExtChartAt {π•œ : Type u_1} {E : Type u_2} {M : Type u_3} {H : Type u_4} {E' : Type u_5} {M' : Type u_6} {H' : Type u_7} [NontriviallyNormedField π•œ] [NormedAddCommGroup E] [NormedSpace π•œ E] [TopologicalSpace H] [TopologicalSpace M] (I : ModelWithCorners π•œ E H) [NormedAddCommGroup E'] [NormedSpace π•œ E'] [TopologicalSpace H'] [TopologicalSpace M'] (I' : ModelWithCorners π•œ E' H') [ChartedSpace H M] [ChartedSpace H' M'] (x : M) (f : M β†’ M') :
                                            E β†’ E'

                                            Conjugating a function to write it in the preferred charts around x. The manifold derivative of f will just be the derivative of this conjugated function.

                                            Equations
                                            Instances For
                                              theorem writtenInExtChartAt_chartAt {π•œ : Type u_1} {E : Type u_2} {M : Type u_3} {H : Type u_4} [NontriviallyNormedField π•œ] [NormedAddCommGroup E] [NormedSpace π•œ E] [TopologicalSpace H] [TopologicalSpace M] {I : ModelWithCorners π•œ E H} [ChartedSpace H M] {x : M} {y : E} (h : y ∈ (extChartAt I x).target) :
                                              writtenInExtChartAt I I x (↑(chartAt H x)) y = y
                                              theorem writtenInExtChartAt_chartAt_symm {π•œ : Type u_1} {E : Type u_2} {M : Type u_3} {H : Type u_4} [NontriviallyNormedField π•œ] [NormedAddCommGroup E] [NormedSpace π•œ E] [TopologicalSpace H] [TopologicalSpace M] {I : ModelWithCorners π•œ E H} [ChartedSpace H M] {x : M} {y : E} (h : y ∈ (extChartAt I x).target) :
                                              writtenInExtChartAt I I (↑(chartAt H x) x) (↑(chartAt H x).symm) y = y
                                              theorem writtenInExtChartAt_extChartAt {π•œ : Type u_1} {E : Type u_2} {M : Type u_3} {H : Type u_4} [NontriviallyNormedField π•œ] [NormedAddCommGroup E] [NormedSpace π•œ E] [TopologicalSpace H] [TopologicalSpace M] {I : ModelWithCorners π•œ E H} [ChartedSpace H M] {x : M} {y : E} (h : y ∈ (extChartAt I x).target) :
                                              writtenInExtChartAt I (modelWithCornersSelf π•œ E) x (↑(extChartAt I x)) y = y
                                              theorem writtenInExtChartAt_extChartAt_symm {π•œ : Type u_1} {E : Type u_2} {M : Type u_3} {H : Type u_4} [NontriviallyNormedField π•œ] [NormedAddCommGroup E] [NormedSpace π•œ E] [TopologicalSpace H] [TopologicalSpace M] {I : ModelWithCorners π•œ E H} [ChartedSpace H M] {x : M} {y : E} (h : y ∈ (extChartAt I x).target) :
                                              writtenInExtChartAt (modelWithCornersSelf π•œ E) I (↑(extChartAt I x) x) (↑(extChartAt I x).symm) y = y
                                              theorem extChartAt_self_eq (π•œ : Type u_1) {E : Type u_2} {H : Type u_4} [NontriviallyNormedField π•œ] [NormedAddCommGroup E] [NormedSpace π•œ E] [TopologicalSpace H] {I : ModelWithCorners π•œ E H} {x : H} :
                                              ↑(extChartAt I x) = ↑I
                                              theorem extChartAt_self_apply (π•œ : Type u_1) {E : Type u_2} {H : Type u_4} [NontriviallyNormedField π•œ] [NormedAddCommGroup E] [NormedSpace π•œ E] [TopologicalSpace H] {I : ModelWithCorners π•œ E H} {x y : H} :
                                              ↑(extChartAt I x) y = ↑I y
                                              theorem extChartAt_model_space_eq_id (π•œ : Type u_1) {E : Type u_2} [NontriviallyNormedField π•œ] [NormedAddCommGroup E] [NormedSpace π•œ E] (x : E) :

                                              In the case of the manifold structure on a vector space, the extended charts are just the identity.

                                              theorem ext_chart_model_space_apply (π•œ : Type u_1) {E : Type u_2} [NontriviallyNormedField π•œ] [NormedAddCommGroup E] [NormedSpace π•œ E] {x y : E} :
                                              ↑(extChartAt (modelWithCornersSelf π•œ E) x) y = y
                                              theorem extChartAt_prod {π•œ : Type u_1} {E : Type u_2} {M : Type u_3} {H : Type u_4} {E' : Type u_5} {M' : Type u_6} {H' : Type u_7} [NontriviallyNormedField π•œ] [NormedAddCommGroup E] [NormedSpace π•œ E] [TopologicalSpace H] [TopologicalSpace M] {I : ModelWithCorners π•œ E H} [NormedAddCommGroup E'] [NormedSpace π•œ E'] [TopologicalSpace H'] [TopologicalSpace M'] {I' : ModelWithCorners π•œ E' H'} [ChartedSpace H M] [ChartedSpace H' M'] (x : M Γ— M') :
                                              extChartAt (I.prod I') x = (extChartAt I x.1).prod (extChartAt I' x.2)
                                              theorem extChartAt_comp {π•œ : Type u_1} {E : Type u_2} {H : Type u_4} {M' : Type u_6} {H' : Type u_7} [NontriviallyNormedField π•œ] [NormedAddCommGroup E] [NormedSpace π•œ E] [TopologicalSpace H] {I : ModelWithCorners π•œ E H} [TopologicalSpace H'] [TopologicalSpace M'] [ChartedSpace H' M'] [ChartedSpace H H'] (x : M') :
                                              extChartAt I x = (chartAt H' x).trans (extChartAt I (↑(chartAt H' x) x))
                                              theorem writtenInExtChartAt_chartAt_comp {π•œ : Type u_1} {E : Type u_2} {H : Type u_4} {M' : Type u_6} {H' : Type u_7} [NontriviallyNormedField π•œ] [NormedAddCommGroup E] [NormedSpace π•œ E] [TopologicalSpace H] {I : ModelWithCorners π•œ E H} [TopologicalSpace H'] [TopologicalSpace M'] [ChartedSpace H' M'] [ChartedSpace H H'] (x : M') {y : E} (hy : y ∈ (extChartAt I x).target) :
                                              writtenInExtChartAt I I x (↑(chartAt H' x)) y = y
                                              theorem writtenInExtChartAt_chartAt_symm_comp {π•œ : Type u_1} {E : Type u_2} {H : Type u_4} {M' : Type u_6} {H' : Type u_7} [NontriviallyNormedField π•œ] [NormedAddCommGroup E] [NormedSpace π•œ E] [TopologicalSpace H] {I : ModelWithCorners π•œ E H} [TopologicalSpace H'] [TopologicalSpace M'] [ChartedSpace H' M'] [ChartedSpace H H'] (x : M') {y : E} (hy : y ∈ (extChartAt I x).target) :
                                              writtenInExtChartAt I I (↑(chartAt H' x) x) (↑(chartAt H' x).symm) y = y
                                              theorem Manifold.locallyCompact_of_finiteDimensional {E : Type u_1} {π•œ : Type u_2} [NontriviallyNormedField π•œ] [NormedAddCommGroup E] [NormedSpace π•œ E] {H : Type u_3} [TopologicalSpace H] {M : Type u_4} [TopologicalSpace M] [ChartedSpace H M] (I : ModelWithCorners π•œ E H) [LocallyCompactSpace π•œ] [FiniteDimensional π•œ E] :

                                              A finite-dimensional manifold modelled on a locally compact field (such as ℝ, β„‚ or the p-adic numbers) is locally compact.

                                              A locally compact manifold must be modelled on a locally compact space.

                                              theorem FiniteDimensional.of_locallyCompact_manifold {E : Type u_1} {π•œ : Type u_2} [NontriviallyNormedField π•œ] [NormedAddCommGroup E] [NormedSpace π•œ E] {H : Type u_3} [TopologicalSpace H] (M : Type u_4) [TopologicalSpace M] [ChartedSpace H M] [CompleteSpace π•œ] (I : ModelWithCorners π•œ E H) [Nonempty M] [LocallyCompactSpace M] :

                                              Riesz's theorem applied to manifolds: a locally compact manifolds must be modelled on a finite-dimensional space. This is the converse to Manifold.locallyCompact_of_finiteDimensional.

                                              def TangentSpace {π•œ : Type u_1} [NontriviallyNormedField π•œ] {E : Type u} [NormedAddCommGroup E] [NormedSpace π•œ E] {H : Type u_2} [TopologicalSpace H] (I : ModelWithCorners π•œ E H) {M : Type u_3} [TopologicalSpace M] [ChartedSpace H M] (_x : M) :

                                              The tangent space at a point of the manifold M. It is just E. We could use instead (tangentBundleCore I M).to_topological_vector_bundle_core.fiber x, but we use E to help the kernel.

                                              Equations
                                              Instances For
                                                instance instAddCommGroupTangentSpace {π•œ : Type u_1} [NontriviallyNormedField π•œ] {E : Type u_2} [NormedAddCommGroup E] [NormedSpace π•œ E] {H : Type u_3} [TopologicalSpace H] (I : ModelWithCorners π•œ E H) {M : Type u_4} [TopologicalSpace M] [ChartedSpace H M] {x : M} :
                                                Equations
                                                instance instTopologicalAddGroupTangentSpace {π•œ : Type u_1} [NontriviallyNormedField π•œ] {E : Type u_2} [NormedAddCommGroup E] [NormedSpace π•œ E] {H : Type u_3} [TopologicalSpace H] (I : ModelWithCorners π•œ E H) {M : Type u_4} [TopologicalSpace M] [ChartedSpace H M] {x : M} :
                                                instance instModuleTangentSpace {π•œ : Type u_1} [NontriviallyNormedField π•œ] {E : Type u_2} [NormedAddCommGroup E] [NormedSpace π•œ E] {H : Type u_3} [TopologicalSpace H] (I : ModelWithCorners π•œ E H) {M : Type u_4} [TopologicalSpace M] [ChartedSpace H M] {x : M} :
                                                Module π•œ (TangentSpace I x)
                                                Equations
                                                instance instInhabitedTangentSpace {π•œ : Type u_1} [NontriviallyNormedField π•œ] {E : Type u_2} [NormedAddCommGroup E] [NormedSpace π•œ E] {H : Type u_3} [TopologicalSpace H] (I : ModelWithCorners π•œ E H) {M : Type u_4} [TopologicalSpace M] [ChartedSpace H M] {x : M} :
                                                Equations
                                                @[reducible, inline]
                                                abbrev TangentBundle {π•œ : Type u_1} [NontriviallyNormedField π•œ] {E : Type u_2} [NormedAddCommGroup E] [NormedSpace π•œ E] {H : Type u_3} [TopologicalSpace H] (I : ModelWithCorners π•œ E H) (M : Type u_4) [TopologicalSpace M] [ChartedSpace H M] :
                                                Type (max u_4 u_2)

                                                The tangent bundle to a manifold, as a Sigma type. Defined in terms of Bundle.TotalSpace to be able to put a suitable topology on it.

                                                Equations
                                                Instances For