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 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 #
ModelWithCorners 𝕜 E H
: a structure containing information on the way a spaceH
embeds in a model vector space E over the field𝕜
. This is all that is needed to define a smooth manifold with model spaceH
, and model vector spaceE
.modelWithCornersSelf 𝕜 E
: trivial model with corners structure on the spaceE
embedded in itself by the identity.contDiffGroupoid n I
: whenI
is a model with corners on(𝕜, E, H)
, this is the groupoid of partial homeos ofH
which are of classC^n
over the normed field𝕜
, when read inE
.SmoothManifoldWithCorners I M
: a type class saying that the charted spaceM
, modelled on the spaceH
, hasC^∞
changes of coordinates with respect to the model with cornersI
on(𝕜, E, H)
. This type class is just a shortcut forHasGroupoid M (contDiffGroupoid ∞ I)
.extChartAt I x
: in a smooth manifold with corners with the modelI
on(E, H)
, the charts take values inH
, but often we may want to use theirE
-valued version, obtained by composing the charts withI
. Since the target is in general not open, we can not register them as partial homeomorphisms, but we register them asPartialEquiv
s.extChartAt I x
is the canonical such partial equiv aroundx
.
As specific examples of models with corners, we define (in Geometry.Manifold.Instances.Real
)
modelWithCornersSelf ℝ (EuclideanSpace ℝ (Fin n))
for the model space used to definen
-dimensional real manifolds without boundary (with notation𝓡 n
in the localeManifold
)ModelWithCorners ℝ (EuclideanSpace ℝ (Fin n)) (EuclideanHalfSpace n)
for the model space used to definen
-dimensional real manifolds with boundary (with notation𝓡∂ n
in the localeManifold
)ModelWithCorners ℝ (EuclideanSpace ℝ (Fin n)) (EuclideanQuadrant n)
for the model space used to definen
-dimensional real manifolds with corners
With these definitions at hand, to invoke an n
-dimensional real manifold without boundary,
one could use
variable {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
variable {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 (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.
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
.
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).
The extended natural number ∞
Equations
- Manifold.«term∞» = Lean.ParserDescr.node `Manifold.«term∞» 1024 (Lean.ParserDescr.symbol "∞")
Instances For
Models with corners. #
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 smooth 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.
- toFun : H → E
- invFun : E → H
- source : Set H
- target : Set E
- source_eq : self.source = Set.univ
- uniqueDiffOn' : UniqueDiffOn 𝕜 self.target
- continuous_toFun : Continuous ↑self.toPartialEquiv
- continuous_invFun : Continuous self.invFun
Instances For
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
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
Equations
- ModelWithCorners.instCoeFunForall = { coe := ModelWithCorners.toFun' }
The inverse to a model with corners, only registered as a PartialEquiv
.
Equations
- I.symm = I.symm
Instances For
See Note [custom simps projection]. We need to specify this projection explicitly in this case, because it is a composition of multiple projections.
Equations
- ModelWithCorners.Simps.apply 𝕜 E H I = ↑I
Instances For
See Note [custom simps projection]
Equations
- ModelWithCorners.Simps.symm_apply 𝕜 E H I = ↑I.symm
Instances For
Alias of ModelWithCorners.uniqueDiffOn
.
Alias of ModelWithCorners.isClosedEmbedding
.
Alias of ModelWithCorners.isClosed_range
.
Alias of ModelWithCorners.uniqueDiffOn_preimage
.
Every smooth manifold is a Fréchet space (T1 space) -- regardless of whether it is Hausdorff.
In the trivial model with corners, the associated PartialEquiv
is the identity.
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
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
Special case of product model with corners, which is trivial on the second factor. This shows up as the model to tangent bundles.
Equations
- I.tangent = I.prod (modelWithCornersSelf 𝕜 E)
Instances For
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.
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
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
The trivial model with corners has no boundary
Equations
- ⋯ = ⋯
If two model with corners are boundaryless, their product also is
Equations
- ⋯ = ⋯
Smooth functions on models with corners #
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
- contDiffPregroupoid n I = { property := fun (f : H → H) (s : Set H) => ContDiffOn 𝕜 n (↑I ∘ f ∘ ↑I.symm) (↑I.symm ⁻¹' s ∩ Set.range ↑I), comp := ⋯, id_mem := ⋯, locality := ⋯, congr := ⋯ }
Instances For
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
- contDiffGroupoid n I = (contDiffPregroupoid n I).groupoid
Instances For
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
partial homeomorphisms
An identity partial homeomorphism belongs to the C^n
groupoid.
The composition of a partial homeomorphism from H
to M
and its inverse belongs to
the C^n
groupoid.
The product of two smooth partial homeomorphisms is smooth.
The C^n
groupoid is closed under restriction.
Equations
- ⋯ = ⋯
Smooth manifolds with corners #
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
For any model with corners, the model space is a smooth manifold
Equations
- ⋯ = ⋯
The maximal atlas of M
for the smooth manifold with corners structure corresponding to the
model with corners I
.
Equations
Instances For
The empty set is a smooth manifold w.r.t. any charted space and model.
Equations
- ⋯ = ⋯
The product of two smooth manifolds with corners is naturally a smooth manifold with corners.
Equations
- ⋯ = ⋯
Alias of IsOpenEmbedding.singleton_smoothManifoldWithCorners
.
Equations
- ⋯ = ⋯
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 PartialHomeomorph
as the target is not open in E
in general, but we can still register them
as PartialEquiv
.
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
Variant of f.extend_left_inv I
, stated in terms of images.
If y ∈ f.target
and I y ∈ interior (range I)
,
then I y
is an interior point of (I ∘ f).target
.
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
.
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.
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.
The preferred extended chart on a manifold with corners around a point x
, from a neighborhood
of x
to the model vector space.
Equations
- extChartAt I x = (chartAt H x).extend I
Instances For
If we're boundaryless, extChartAt
has open target
If we're boundaryless, (extChartAt I x).target
is a neighborhood of the key point
If we're boundaryless, (extChartAt I x).target
is a neighborhood of any of its points
Around the image of a point in the source, the neighborhoods are the same
within (extChartAt I x).target
and within range I
.
Around a point in the target, the neighborhoods are the same within (extChartAt I x).target
and within range I
.
Around the image of the base point, the neighborhoods are the same
within (extChartAt I x).target
and within range I
.
Around the image of a point in the source, (extChartAt I x).target
and range I
coincide locally.
Around a point in the target, (extChartAt I x).target
and range I
coincide locally.
Around the image of the base point, (extChartAt I x).target
and range I
coincide locally.
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.
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.
Technical lemma ensuring that the preimage under an extended chart of a neighborhood of a point is a neighborhood of the preimage.
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.
We use the name ext_coord_change
for (extChartAt I x').symm ≫ extChartAt I x
.
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
- writtenInExtChartAt I I' x f = ↑(extChartAt I' (f x)) ∘ f ∘ ↑(extChartAt I x).symm
Instances For
In the case of the manifold structure on a vector space, the extended charts are just the identity.
A finite-dimensional manifold modelled on a locally compact field
(such as ℝ, ℂ or the p
-adic numbers) is locally compact.
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
- TangentSpace I _x = E
Instances For
Equations
Equations
Equations
- ⋯ = ⋯
Equations
- instModuleTangentSpace I = inferInstanceAs (Module 𝕜 E)
Equations
- instInhabitedTangentSpace I = { default := 0 }
The tangent bundle to a smooth manifold, as a Sigma type. Defined in terms of
Bundle.TotalSpace
to be able to put a suitable topology on it.
Equations
- TangentBundle I M = Bundle.TotalSpace E (TangentSpace I)
Instances For
Equations
- ⋯ = ⋯