Partially defined linear maps
A linear_pmap R E F
is a linear map from a submodule of E
to F
. We define
a semilattice_inf_bot
instance on this this, and define three operations:
mk_span_singleton
defines a partial linear map defined on the span of a singleton.sup
takes two partial linear mapsf
,g
that agree on the intersection of their domains, and returns the unique partial linear map onf.domain ⊔ g.domain
that extends bothf
andg
.Sup
takes adirected_on (≤)
set of partial linear maps, and returns the unique partial linear map on theSup
of their domains that extends all these maps.
Partially defined maps are currently used in mathlib
to prove Hahn-Banach theorem
and its variations. Namely, linear_pmap.Sup
implies that every chain of linear_pmap
s
is bounded above.
Another possible use (not yet in mathlib
) would be the theory of unbounded linear operators.
A linear_pmap R E F
is a linear map from a submodule of E
to F
.
Equations
- linear_pmap.has_coe_to_fun = {F := λ (f : linear_pmap R E F), ↥(f.domain) → F, coe := λ (f : linear_pmap R E F), ⇑(f.to_fun)}
The unique linear_pmap
on span R {x}
that sends x
to y
. This version works for modules
over rings, and requires a proof of ∀ c, c • x = 0 → c • y = 0
.
Equations
- linear_pmap.mk_span_singleton' x y H = {domain := submodule.span R {x}, to_fun := {to_fun := λ (z : ↥(submodule.span R {x})), classical.some _ • y, map_add' := _, map_smul' := _}}
The unique linear_pmap
on span R {x}
that sends a non-zero vector x
to y
.
This version works for modules over division rings.
Equations
- linear_pmap.mk_span_singleton x y hx = linear_pmap.mk_span_singleton' x y _
Projection to the first coordinate as a linear_pmap
Equations
- linear_pmap.fst p p' = {domain := p.prod p', to_fun := (linear_map.fst R E F).comp (p.prod p').subtype}
Projection to the second coordinate as a linear_pmap
Equations
- linear_pmap.snd p p' = {domain := p.prod p', to_fun := (linear_map.snd R E F).comp (p.prod p').subtype}
Given two partial linear maps f
, g
, the set of points x
such that
both f
and g
are defined at x
and f x = g x
form a submodule.
Equations
- linear_pmap.has_inf = {inf := λ (f g : linear_pmap R E F), {domain := f.eq_locus g, to_fun := f.to_fun.comp (submodule.of_le _)}}
Equations
Equations
- linear_pmap.semilattice_inf_bot = {bot := ⊥, le := has_le.le linear_pmap.has_le, lt := order_bot.lt._default has_le.le, le_refl := _, le_trans := _, lt_iff_le_not_le := _, le_antisymm := _, bot_le := _, inf := has_inf.inf linear_pmap.has_inf, inf_le_left := _, inf_le_right := _, le_inf := _}
Given two partial linear maps that agree on the intersection of their domains,
f.sup g h
is the unique partial linear map on f.domain ⊔ g.domain
that agrees
with f
and g
.
Hypothesis for linear_pmap.sup
holds, if f.domain
is disjoint with g.domain
.
Glue a collection of partially defined linear maps to a linear map defined on Sup
of these submodules.
Equations
- linear_pmap.Sup c hc = {domain := has_Sup.Sup (linear_pmap.domain '' c), to_fun := classical.some _}
Restrict a linear map to a submodule, reinterpreting the result as a linear_pmap
.
Compose a linear map with a linear_pmap
Restrict codomain of a linear_pmap
Equations
- f.cod_restrict p H = {domain := f.domain, to_fun := linear_map.cod_restrict p f.to_fun H}
Compose two linear_pmap
s
f.coprod g
is the partially defined linear map defined on f.domain × g.domain
,
and sending p
to f p.1 + g p.2
.