Local properties invariant under a groupoid
We study properties of a triple (g, s, x)
where g
is a function between two spaces H
and H'
,
s
is a subset of H
and x
is a point of H
. Our goal is to register how such a property
should behave to make sense in charted spaces modelled on H
and H'
.
The main examples we have in mind are the properties "g
is differentiable at x
within s
", or
"g
is smooth at x
within s
". We want to develop general results that, when applied in these
specific situations, say that the notion of smooth function in a manifold behaves well under
restriction, intersection, is local, and so on.
Main definitions
local_invariant_prop G G' P
says that a propertyP
of a triple(g, s, x)
is local, and invariant under composition by elements of the groupoidsG
andG'
ofH
andH'
respectively.charted_space.lift_prop_within_at
(resp.lift_prop_at
,lift_prop_on
andlift_prop
): given a propertyP
of(g, s, x)
whereg : H → H'
, define the corresponding property for functionsM → M'
whereM
andM'
are charted spaces modelled respectively onH
andH'
. We define these properties within a set at a point, or at a point, or on a set, or in the whole space. This lifting process (obtained by restricting to suitable chart domains) can always be done, but it only behaves well under locality and invariance assumptions.
Given hG : local_invariant_prop G G' P
, we deduce many properties of the lifted property on the
charted spaces. For instance, hG.lift_prop_within_at_inter
says that P g s x
is equivalent to
P g (s ∩ t) x
whenever t
is a neighborhood of x
.
Implementation notes
We do not use dot notation for properties of the lifted property. For instance, we have
hG.lift_prop_within_at_congr
saying that if lift_prop_within_at P g s x
holds, and g
and g'
coincide on s
, then lift_prop_within_at P g' s x
holds. We can't call it
lift_prop_within_at.congr
as it is in the namespace associated to local_invariant_prop
, not
in the one for lift_prop_within_at
.
- is_local : ∀ {s : set H} {x : H} {u : set H} {f : H → H'}, is_open u → x ∈ u → (P f s x ↔ P f (s ∩ u) x)
- right_invariance : ∀ {s : set H} {x : H} {f : H → H'} {e : local_homeomorph H H}, e ∈ G → x ∈ e.to_local_equiv.source → P f s x → P (f ∘ ⇑(e.symm)) (e.to_local_equiv.target ∩ ⇑(e.symm) ⁻¹' s) (⇑e x)
- congr : ∀ {s : set H} {x : H} {f g : H → H'}, (∀ (y : H), y ∈ s → f y = g y) → f x = g x → P f s x → P g s x
- left_invariance : ∀ {s : set H} {x : H} {f : H → H'} {e' : local_homeomorph H' H'}, e' ∈ G' → s ⊆ f ⁻¹' e'.to_local_equiv.source → f x ∈ e'.to_local_equiv.source → P f s x → P (⇑e' ∘ f) s x
Structure recording good behavior of a property of a triple (f, s, x)
where f
is a function,
s
a set and x
a point. Good behavior here means locality and invariance under given groupoids
(both in the source and in the target). Given such a good behavior, the lift of this property
to charted spaces admitting these groupoids will inherit the good behavior.
Given a property of germs of functions and sets in the model space, then one defines
a corresponding property in a charted space, by requiring that it holds at the preferred chart at
this point. (When the property is local and invariant, it will in fact hold using any chart, see
lift_prop_within_at_indep_chart
). We require continuity in the lifted property, as otherwise one
single chart might fail to capture the behavior of the function.
Equations
- charted_space.lift_prop_within_at P f s x = (continuous_within_at f s x ∧ P (⇑(charted_space.chart_at H' (f x)) ∘ f ∘ ⇑((charted_space.chart_at H x).symm)) ((charted_space.chart_at H x).to_local_equiv.target ∩ ⇑((charted_space.chart_at H x).symm) ⁻¹' (s ∩ f ⁻¹' (charted_space.chart_at H' (f x)).to_local_equiv.source)) (⇑(charted_space.chart_at H x) x))
Given a property of germs of functions and sets in the model space, then one defines a corresponding property of functions on sets in a charted space, by requiring that it holds around each point of the set, in the preferred charts.
Equations
- charted_space.lift_prop_on P f s = ∀ (x : M), x ∈ s → charted_space.lift_prop_within_at P f s x
Given a property of germs of functions and sets in the model space, then one defines a corresponding property of a function at a point in a charted space, by requiring that it holds in the preferred chart.
Equations
Given a property of germs of functions and sets in the model space, then one defines a corresponding property of a function in a charted space, by requiring that it holds in the preferred chart around every point.
Equations
- charted_space.lift_prop P f = ∀ (x : M), charted_space.lift_prop_at P f x
If a property of a germ of function g
on a pointed set (s, x)
is invariant under the
structure groupoid (by composition in the source space and in the target space), then
expressing it in charted spaces does not depend on the element of the maximal atlas one uses
both in the source and in the target manifolds, provided they are defined around x
and g x
respectively, and provided g
is continuous within s
at x
(otherwise, the local behavior
of g
at x
can not be captured with a chart in the target).
A function from a model space H
to itself is a local structomorphism, with respect to a
structure groupoid G
for H
, relative to a set s
in H
, if for all points x
in the set, the
function agrees with a G
-structomorphism on s
in a neighbourhood of x
.
Equations
- G.is_local_structomorph_within_at f s x = (x ∈ s → (∃ (e : local_homeomorph H H), e ∈ G ∧ set.eq_on f e.to_local_equiv.to_fun (s ∩ e.to_local_equiv.source) ∧ x ∈ e.to_local_equiv.source))
For a groupoid G
which is closed_under_restriction
, being a local structomorphism is a local
invariant property.