Documentation

Mathlib.Geometry.Manifold.SmoothManifoldWithCorners

Smooth manifolds (possibly with boundary or corners) #

A smooth 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 smooth 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 smooth when read in E (for any regularity n : ℕ∞). 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). We also introduce a specific type class for C^∞ manifolds as these are the most commonly used.

Some texts assume manifolds to be Hausdorff and secound 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 the file real_instances.lean)

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

variables {n : ℕ} {M : Type*} [TopologicalSpace M] [ChartedSpace (EuclideanSpace (Fin n)) M] [SmoothManifoldWithCorners (𝓡 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

variables {E : Type*} [NormedAddCommGroup E] [NormedSpace ℝ E] [FiniteDimensional ℝ E] {I : ModelWithCorners ℝ E E} [I.Boundaryless] {M : Type*} [TopologicalSpace M] [ChartedSpace E M] [SmoothManifoldWithCorners 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 (λ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.

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.

We concentrate on C^∞ manifolds: all the definitions work equally well for C^n manifolds, but later on it is a pain to carry all over the smoothness parameter, especially when one wants to deal with C^k functions as there would be additional conditions k ≤ n everywhere. Since one deals almost all the time with C^∞ (or analytic) manifolds, this seems to be a reasonable choice that one could revisit later if needed. C^k manifolds are still available, but they should be called using HasGroupoid M (contDiffGroupoid k I) where I is the model with corners.

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

theorem ModelWithCorners.ext {𝕜 : Type u_1} :
∀ {inst : NontriviallyNormedField 𝕜} {E : Type u_2} {inst_1 : NormedAddCommGroup E} {inst_2 : NormedSpace 𝕜 E} {H : Type u_3} {inst_3 : TopologicalSpace H} (x y : ModelWithCorners 𝕜 E H), x.toLocalEquiv = y.toLocalEquivx.invFun = y.invFunx.source = y.sourcex.target = y.targetx = y
theorem ModelWithCorners.ext_iff {𝕜 : Type u_1} :
∀ {inst : NontriviallyNormedField 𝕜} {E : Type u_2} {inst_1 : NormedAddCommGroup E} {inst_2 : NormedSpace 𝕜 E} {H : Type u_3} {inst_3 : TopologicalSpace H} (x y : ModelWithCorners 𝕜 E H), x = y x.toLocalEquiv = y.toLocalEquiv x.invFun = y.invFun x.source = y.source x.target = y.target
structure ModelWithCorners (𝕜 : Type u_1) [NontriviallyNormedField 𝕜] (E : Type u_2) [NormedAddCommGroup E] [NormedSpace 𝕜 E] (H : Type u_3) [TopologicalSpace H] extends LocalEquiv :
Type (max u_2 u_3)
  • toFun : HE
  • invFun : EH
  • source : Set H
  • target : Set E
  • map_source' : ∀ ⦃x : H⦄, x s.sources.toLocalEquiv x s.target
  • map_target' : ∀ ⦃x : E⦄, x s.targetLocalEquiv.invFun s.toLocalEquiv x s.source
  • left_inv' : ∀ ⦃x : H⦄, x s.sourceLocalEquiv.invFun s.toLocalEquiv (s.toLocalEquiv x) = x
  • right_inv' : ∀ ⦃x : E⦄, x s.targets.toLocalEquiv (LocalEquiv.invFun s.toLocalEquiv x) = x
  • source_eq : s.source = Set.univ
  • unique_diff' : UniqueDiffOn 𝕜 s.target
  • continuous_toFun : Continuous s.toLocalEquiv
  • continuous_invFun : Continuous s.invFun

A structure containing informations 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 smooth manifold with model space H, and model vector space E.

Instances For

    A vector space is a model with corners.

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

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

      Instances For
        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 local equiv.

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

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

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

            See Note [custom simps projection]

            Instances For
              @[simp]
              theorem ModelWithCorners.toLocalEquiv_coe {𝕜 : Type u_1} [NontriviallyNormedField 𝕜] {E : Type u_2} [NormedAddCommGroup E] [NormedSpace 𝕜 E] {H : Type u_3} [TopologicalSpace H] (I : ModelWithCorners 𝕜 E H) :
              I.toLocalEquiv = 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 : LocalEquiv H E) (a : e.source = Set.univ) (b : UniqueDiffOn 𝕜 e.target) (c : Continuous e) (d : Continuous e.invFun) :
              ↑(ModelWithCorners.mk e a b) = e
              @[simp]
              theorem ModelWithCorners.toLocalEquiv_coe_symm {𝕜 : Type u_1} [NontriviallyNormedField 𝕜] {E : Type u_2} [NormedAddCommGroup E] [NormedSpace 𝕜 E] {H : Type u_3} [TopologicalSpace H] (I : ModelWithCorners 𝕜 E H) :
              ↑(LocalEquiv.symm I.toLocalEquiv) = ↑(ModelWithCorners.symm I)
              @[simp]
              theorem ModelWithCorners.mk_symm {𝕜 : Type u_1} [NontriviallyNormedField 𝕜] {E : Type u_2} [NormedAddCommGroup E] [NormedSpace 𝕜 E] {H : Type u_3} [TopologicalSpace H] (e : LocalEquiv H E) (a : e.source = Set.univ) (b : UniqueDiffOn 𝕜 e.target) (c : Continuous e) (d : Continuous e.invFun) :
              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) :
              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} :
              @[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.unique_diff {𝕜 : 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) :
              ↑(ModelWithCorners.symm I) (I x) = x
              @[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) :
              ↑(ModelWithCorners.symm I) I = id
              @[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 (↑(ModelWithCorners.symm I) 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) :
              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 (↑(ModelWithCorners.symm I)) (nhdsWithin (I x) (I '' s)) = nhdsWithin x s
              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) :
              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)
              @[simp]

              In the trivial model with corners, the associated local equiv is the identity.

              @[simp]
              theorem modelWithCornersSelf_coe (𝕜 : Type u_1) [NontriviallyNormedField 𝕜] (E : Type u_2) [NormedAddCommGroup E] [NormedSpace 𝕜 E] :
              ↑(modelWithCornersSelf 𝕜 E) = id
              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') :
              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') :
              (ModelWithCorners.prod I I').toLocalEquiv.source = {x | x.fst I.source x.snd I'.source}
              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') :
              ↑(ModelWithCorners.prod I I') x = (I x.fst, I' x.snd)
              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') :
              (ModelWithCorners.prod I I').toLocalEquiv.target = (LocalEquiv.prod I.toLocalEquiv I'.toLocalEquiv).target
              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'.

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

                Instances For
                  @[reducible]
                  def 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.

                  Instances For
                    @[simp]
                    theorem modelWithCorners_prod_toLocalEquiv {𝕜 : Type u_1} [NontriviallyNormedField 𝕜] {E : Type u_2} [NormedAddCommGroup E] [NormedSpace 𝕜 E] {F : Type u_4} [NormedAddCommGroup F] [NormedSpace 𝕜 F] {H : Type u_6} [TopologicalSpace H] {G : Type u_8} [TopologicalSpace G] {I : ModelWithCorners 𝕜 E H} {J : ModelWithCorners 𝕜 F G} :
                    (ModelWithCorners.prod I J).toLocalEquiv = LocalEquiv.prod I.toLocalEquiv J.toLocalEquiv
                    @[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_6} [TopologicalSpace H] {H' : Type u_7} [TopologicalSpace H'] (I : ModelWithCorners 𝕜 E H) (I' : ModelWithCorners 𝕜 E' H') :
                    ↑(ModelWithCorners.prod I I') = Prod.map I I'
                    @[simp]
                    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_6} [TopologicalSpace H] {G : Type u_8} [TopologicalSpace G] {I : ModelWithCorners 𝕜 E H} {J : ModelWithCorners 𝕜 F G} :
                    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.

                    Instances
                      @[simp]

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

                      Instances For

                        The trivial model with corners has no boundary

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

                        Smooth functions on models with corners #

                        def contDiffGroupoid (n : ℕ∞) {𝕜 : 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 C^n transformations of H as the maps that are C^n when read in E through I.

                        Instances For
                          theorem contDiffGroupoid_le {m : ℕ∞} {n : ℕ∞} {𝕜 : 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

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

                          theorem ofSet_mem_contDiffGroupoid (n : ℕ∞) {𝕜 : 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 local homeomorphism belongs to the C^n groupoid.

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

                          theorem contDiffGroupoid_prod {𝕜 : 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 : LocalHomeomorph H H} {e' : LocalHomeomorph H' H'} (he : e contDiffGroupoid I) (he' : e' contDiffGroupoid I') :

                          The product of two smooth local homeomorphisms is smooth.

                          The C^n groupoid is closed under restriction.

                          def analyticGroupoid {𝕜 : 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 analytic transformations of H as the maps that are analytic and map interior to interior when read in E through I. We also explicitly define that they are C^∞ on the whole domain, since we are only requiring analyticity on the interior of the domain.

                          Instances For
                            theorem ofSet_mem_analyticGroupoid {𝕜 : 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 local homeomorphism belongs to the analytic groupoid.

                            The composition of a local homeomorphism from H to M and its inverse belongs to the analytic groupoid.

                            The analytic groupoid is closed under restriction.

                            theorem mem_analyticGroupoid_of_boundaryless {𝕜 : Type u_1} [NontriviallyNormedField 𝕜] {E : Type u_2} [NormedAddCommGroup E] [NormedSpace 𝕜 E] {H : Type u_3} [TopologicalSpace H] (I : ModelWithCorners 𝕜 E H) [CompleteSpace E] [ModelWithCorners.Boundaryless I] (e : LocalHomeomorph H H) :
                            e analyticGroupoid I AnalyticOn 𝕜 (I e ↑(ModelWithCorners.symm I)) (I '' e.source) AnalyticOn 𝕜 (I ↑(LocalHomeomorph.symm e) ↑(ModelWithCorners.symm I)) (I '' e.target)

                            The analytic groupoid on a boundaryless charted space modeled on a complete vector space consists of the local homeomorphisms which are analytic and have analytic inverse.

                            Smooth manifolds with corners #

                            class SmoothManifoldWithCorners {𝕜 : 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] extends HasGroupoid :

                              Typeclass defining smooth manifolds with corners with respect to a model with corners, over a field 𝕜 and with infinite smoothness to simplify typeclass search and statements later on.

                              Instances
                                instance model_space_smooth {𝕜 : Type u_1} [NontriviallyNormedField 𝕜] {E : Type u_2} [NormedAddCommGroup E] [NormedSpace 𝕜 E] {H : Type u_3} [TopologicalSpace H] {I : ModelWithCorners 𝕜 E H} :

                                For any model with corners, the model space is a smooth manifold

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

                                Instances For
                                  instance SmoothManifoldWithCorners.prod {𝕜 : 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] [SmoothManifoldWithCorners I M] (M' : Type u_11) [TopologicalSpace M'] [ChartedSpace H' M'] [SmoothManifoldWithCorners I' M'] :

                                  The product of two smooth manifolds with corners is naturally a smooth manifold with corners.

                                  theorem LocalHomeomorph.singleton_smoothManifoldWithCorners {𝕜 : 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 : LocalHomeomorph M H) (h : e.source = Set.univ) :

                                  Extended charts #

                                  In a smooth 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 LocalHomeomorph as the target is not open in E in general, but we can still register them as LocalEquiv.

                                  def LocalHomeomorph.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 : LocalHomeomorph 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.

                                  Instances For
                                    theorem LocalHomeomorph.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 : LocalHomeomorph M H) (I : ModelWithCorners 𝕜 E H) :
                                    ↑(LocalHomeomorph.extend f I) = I f
                                    theorem LocalHomeomorph.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 : LocalHomeomorph M H) (I : ModelWithCorners 𝕜 E H) :
                                    (LocalHomeomorph.extend f I).source = f.source
                                    theorem LocalHomeomorph.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 : LocalHomeomorph M H) (I : ModelWithCorners 𝕜 E H) :
                                    theorem LocalHomeomorph.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 : LocalHomeomorph M H) (I : ModelWithCorners 𝕜 E H) {s : Set M} (hs : s f.source) :
                                    theorem LocalHomeomorph.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 : LocalHomeomorph M H) (I : ModelWithCorners 𝕜 E H) {x : M} (hxf : x f.source) :
                                    theorem LocalHomeomorph.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 : LocalHomeomorph M H) (I : ModelWithCorners 𝕜 E H) {x : M} (h : x f.source) :
                                    theorem LocalHomeomorph.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 : LocalHomeomorph M H) (I : ModelWithCorners 𝕜 E H) {s : Set M} {x : M} (h : x f.source) :
                                    theorem LocalHomeomorph.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 : LocalHomeomorph M H) (I : ModelWithCorners 𝕜 E H) {x : M} (h : x f.source) :
                                    theorem LocalHomeomorph.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 : LocalHomeomorph M H) (I : ModelWithCorners 𝕜 E H) {x : M} (hy : x f.source) :
                                    theorem LocalHomeomorph.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 : LocalHomeomorph M H) (I : ModelWithCorners 𝕜 E H) {y : M} (hy : y f.source) :
                                    theorem LocalHomeomorph.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 : LocalHomeomorph M H) (I : ModelWithCorners 𝕜 E H) {y : M} (hy : y f.source) :
                                    theorem LocalHomeomorph.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 : LocalHomeomorph M H) (I : ModelWithCorners 𝕜 E H) {x : M} (h : x f.source) :
                                    theorem LocalHomeomorph.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 : LocalHomeomorph M H) (I : ModelWithCorners 𝕜 E H) {s : Set E} (hs : IsOpen s) :
                                    theorem LocalHomeomorph.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 : LocalHomeomorph M H) (I : ModelWithCorners 𝕜 E H) {s : Set E} (hs : IsOpen s) :
                                    IsOpen (f.source ↑(LocalHomeomorph.extend f I) ⁻¹' s)
                                    theorem LocalHomeomorph.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 : LocalHomeomorph M H) (I : ModelWithCorners 𝕜 E H) {s : Set M} {y : M} (hy : y f.source) :
                                    theorem LocalHomeomorph.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 : LocalHomeomorph M H) (I : ModelWithCorners 𝕜 E H) {s : Set M} {y : M} (hy : y f.source) (hs : s f.source) :
                                    theorem LocalHomeomorph.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 : LocalHomeomorph M H) (I : ModelWithCorners 𝕜 E H) {s : Set M} {y : M} (hy : y f.source) :
                                    theorem LocalHomeomorph.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 : LocalHomeomorph M H) (I : ModelWithCorners 𝕜 E H) {y : M} (hy : y f.source) :
                                    theorem LocalHomeomorph.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 : LocalHomeomorph M H) (I : ModelWithCorners 𝕜 E H) {α : Type u_8} {y : M} {l : Filter α} {g : αM} (hg : ∀ᶠ (z : α) in l, g z f.source) (hy : y f.source) :
                                    theorem LocalHomeomorph.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 : LocalHomeomorph M H) (I : ModelWithCorners 𝕜 E H) [NormedAddCommGroup E'] [NormedSpace 𝕜 E'] [TopologicalSpace H'] [TopologicalSpace M'] (I' : ModelWithCorners 𝕜 E' H') {s : Set M} {y : M} {f' : LocalHomeomorph M' H'} {g : MM'} (hy : y f.source) (hgy : g y f'.source) (hmaps : Set.MapsTo g s f'.source) :
                                    theorem LocalHomeomorph.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 : LocalHomeomorph M H) (I : ModelWithCorners 𝕜 E H) [NormedAddCommGroup E'] [NormedSpace 𝕜 E'] [TopologicalSpace H'] [TopologicalSpace M'] (I' : ModelWithCorners 𝕜 E' H') {s : Set M} {f' : LocalHomeomorph M' H'} {g : MM'} (hs : s f.source) (hmaps : Set.MapsTo g s f'.source) :

                                    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 LocalHomeomorph.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 : LocalHomeomorph M H) (I : ModelWithCorners 𝕜 E H) {s : Set M} {t : Set M} {x : M} (h : x f.source) (ht : t nhdsWithin x s) :

                                    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 LocalHomeomorph.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 : LocalHomeomorph M H) (I : ModelWithCorners 𝕜 E H) {t : Set M} {x : M} (h : x f.source) (ht : t nhds x) :

                                    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 LocalHomeomorph.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 : LocalHomeomorph M H) (I : ModelWithCorners 𝕜 E H) {s : Set M} {x : M} (hs : s f.source) (hx : x f.source) :

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

                                    theorem LocalHomeomorph.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 : LocalHomeomorph M H) (f' : LocalHomeomorph M H) (I : ModelWithCorners 𝕜 E H) {x : M} (hxf : x f.source) (hxf' : x f'.source) :
                                    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.

                                    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) (x : M) [ChartedSpace H 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) (x : M) [ChartedSpace H M] :
                                      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) (x : M) [ChartedSpace H M] :
                                      (extChartAt I x).source = (chartAt H x).toLocalEquiv.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) (x : M) [ChartedSpace H 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) (x : M) [ChartedSpace H 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) (x : M) [ChartedSpace H 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 = ↑(ModelWithCorners.symm I) ⁻¹' (chartAt H x).toLocalEquiv.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) (x : M) [ChartedSpace H M] :
                                      ↑(LocalEquiv.symm (extChartAt I x)) (↑(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) (x : M) {s : Set M} [ChartedSpace H M] (hs : s (chartAt H x).toLocalEquiv.source) :
                                      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) (x : M) [ChartedSpace H M] {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) (x : M) [ChartedSpace H 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) (x : M) {s : Set M} [ChartedSpace H M] {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) (x : M) {s : Set M} [ChartedSpace H 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) (x : M) [ChartedSpace H 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) (x : M) [ChartedSpace H M] {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) (x : M) [ChartedSpace H M] :
                                      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} {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) (x : M) [ChartedSpace H M] :
                                      Filter.map (↑(extChartAt I x)) (nhds x) = nhdsWithin (↑(extChartAt I x) x) (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) (x : M) [ChartedSpace H M] {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) (x : M) [ChartedSpace H M] :
                                      (extChartAt I x).target nhdsWithin (↑(extChartAt I x) x) (Set.range I)
                                      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) (x : M) [ChartedSpace H 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) (x : M) [ChartedSpace H M] {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)
                                      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) (x : M) [ChartedSpace H M] :
                                      nhdsWithin (↑(extChartAt I x) x) (extChartAt I x).target = nhdsWithin (↑(extChartAt I x) x) (Set.range I)
                                      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) (x : M) [ChartedSpace H M] {y : E} (h : y (extChartAt I x).target) :
                                      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) (x : M) [ChartedSpace H M] {x' : M} (h : x' (extChartAt I x).source) :
                                      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) (x : M) [ChartedSpace H M] :
                                      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) (x : M) [ChartedSpace H M] :
                                      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) (x : M) [ChartedSpace H 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) (x : M) [ChartedSpace H M] {s : Set E} (hs : IsOpen s) :
                                      IsOpen ((chartAt H x).toLocalEquiv.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) (x : M) {s : Set M} [ChartedSpace H M] {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) (x : M) {s : Set M} [ChartedSpace H 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) (x : M) {s : Set M} [ChartedSpace H M] {y : M} (hy : y (extChartAt I x).source) :
                                      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) (x : M) {s : Set M} [ChartedSpace H M] :
                                      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) (x : M) {s : Set M} [ChartedSpace H M] {y : M} (hy : y (extChartAt I x).source) :
                                      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) (x : M) [ChartedSpace H M] {y : M} (hy : y (extChartAt I x).source) :
                                      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) (x : M) {s : Set M} [ChartedSpace H M] :
                                      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) (x : M) [ChartedSpace H M] :
                                      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) (x : M) {s : Set M} {t : Set M} [ChartedSpace H M] {x' : M} (h : x' (extChartAt I x).source) (ht : t nhdsWithin x' s) :

                                      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) (x : M) {s : Set M} {t : Set M} [ChartedSpace H M] (ht : t nhdsWithin x s) :

                                      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) (x : M) {t : Set M} [ChartedSpace H M] {x' : M} (h : x' (extChartAt I x).source) (ht : t nhds x') :
                                      ↑(LocalEquiv.symm (extChartAt I x)) ⁻¹' 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) (x : M) {t : Set M} [ChartedSpace H M] (ht : t nhds 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) (x : M) {s : Set M} {t : Set M} [ChartedSpace H M] :

                                      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 : MM'} {x : M} (hc : ContinuousWithinAt f s x) :
                                      nhdsWithin (↑(extChartAt I x) x) (↑(LocalEquiv.symm (extChartAt I x)) ⁻¹' s Set.range I) = nhdsWithin (↑(extChartAt I x) x) ((extChartAt I x).target ↑(LocalEquiv.symm (extChartAt I x)) ⁻¹' (s f ⁻¹' (extChartAt I' (f x)).source))

                                      We use the name ext_coord_change for (extChartAt I x').symmextChartAt 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 : M) (x' : M) :
                                      (LocalEquiv.trans (LocalEquiv.symm (extChartAt I x')) (extChartAt I x)).source = I '' (LocalHomeomorph.trans (LocalHomeomorph.symm (chartAt H x')) (chartAt H x)).toLocalEquiv.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] (I : ModelWithCorners 𝕜 E H) [ChartedSpace H M] [SmoothManifoldWithCorners I M] (x : M) (x' : M) :
                                      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] (I : ModelWithCorners 𝕜 E H) [ChartedSpace H M] [SmoothManifoldWithCorners I M] (x : M) (x' : M) {y : E} (hy : y (LocalEquiv.trans (LocalEquiv.symm (extChartAt I x')) (extChartAt I x)).source) :
                                      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 : MM') :
                                      EE'

                                      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.

                                      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) (↑(LocalHomeomorph.symm (chartAt H x))) 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) :
                                        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) :
                                        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 : H} {y : H} :
                                        ↑(extChartAt I x) y = I y

                                        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 : E} {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') :
                                        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 = LocalEquiv.trans (chartAt H' x).toLocalEquiv (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) (↑(LocalHomeomorph.symm (chartAt H' x))) y = y