Dual vector spaces #
The dual space of an $R$-module $M$ is the $R$-module of $R$-linear maps $M \to R$.
Main definitions #
- Duals and transposes:
Module.Dual R M
defines the dual space of theR
-moduleM
, asM →ₗ[R] R
.Module.dualPairing R M
is the canonical pairing betweenDual R M
andM
.Module.Dual.eval R M : M →ₗ[R] Dual R (Dual R)
is the canonical map to the double dual.Module.Dual.transpose
is the linear map fromM →ₗ[R] M'
toDual R M' →ₗ[R] Dual R M
.LinearMap.dualMap
isModule.Dual.transpose
of a given linear map, for dot notation.LinearEquiv.dualMap
is for the dual of an equivalence.
- Submodules:
Submodule.dualRestrict W
is the transposeDual R M →ₗ[R] Dual R W
of the inclusion map.Submodule.dualAnnihilator W
is the kernel ofW.dualRestrict
. That is, it is the submodule ofdual R M
whose elements all annihilateW
.Submodule.dualPairing W
is the canonical pairing betweenDual R M ⧸ W.dualAnnihilator
andW
. It is nondegenerate for vector spaces (Subspace.dualPairing_nondegenerate
).
Main results #
- Annihilators:
Module.dualAnnihilator_gc R M
is the antitone Galois correspondence betweenSubmodule.dualAnnihilator
andSubmodule.dualConnihilator
.
- Finite-dimensional vector spaces:
Module.evalEquiv
is the equivalenceV ≃ₗ[K] Dual K (Dual K V)
Module.mapEvalEquiv
is the order isomorphism between subspaces ofV
and subspaces ofDual K (Dual K V)
.
The dual space of an R-module M is the R-module of linear maps M → R
.
Equations
- Module.Dual R M = (M →ₗ[R] R)
Instances For
The canonical pairing of a vector space and its algebraic dual.
Equations
Instances For
Equations
- Module.Dual.instInhabited R M = { default := 0 }
Maps a module M to the dual of the dual of M. See Module.erange_coe
and
Module.evalEquiv
.
Equations
Instances For
The transposition of linear maps, as a linear map from M →ₗ[R] M'
to
Dual R M' →ₗ[R] Dual R M
.
Equations
- Module.Dual.transpose = (LinearMap.llcomp R M M' R).flip
Instances For
Given a linear map f : M₁ →ₗ[R] M₂
, f.dualMap
is the linear map between the dual of
M₂
and M₁
such that it maps the functional φ
to φ ∘ f
.
Equations
Instances For
If a linear map is surjective, then its dual is injective.
The Linear_equiv
version of LinearMap.dualMap
.
Equations
Instances For
A reflexive module is one for which the natural map to its double dual is a bijection.
Any finitely-generated projective module (and thus any finite-dimensional vector space)
is reflexive. See Module.instIsReflexiveOfFiniteOfProjective
.
- bijective_dual_eval' : Function.Bijective ⇑(Dual.eval R M)
A reflexive module is one for which the natural map to its double dual is a bijection.
Instances
The bijection between a reflexive module and its double dual, bundled as a LinearEquiv
.
Equations
- Module.evalEquiv R M = LinearEquiv.ofBijective (Module.Dual.eval R M) ⋯
Instances For
The dual of a reflexive module is reflexive.
A direct summand of a reflexive module is reflexive.
The isomorphism Module.evalEquiv
induces an order isomorphism on subspaces.
Equations
Instances For
The dualRestrict
of a submodule W
of M
is the linear map from the
dual of M
to the dual of W
such that the domain of each linear map is
restricted to W
.
Equations
Instances For
The dualAnnihilator
of a submodule W
is the set of linear maps φ
such
that φ w = 0
for all w ∈ W
.
Equations
Instances For
That $\operatorname{ker}(\iota^* : V^* \to W^*) = \operatorname{ann}(W)$. This is the definition of the dual annihilator of the submodule $W$.
The dualAnnihilator
of a submodule of the dual space pulled back along the evaluation map
Module.Dual.eval
.
Equations
Instances For
See also Subspace.dualAnnihilator_inf_eq
for vector subspaces.
See also Subspace.dualAnnihilator_iInf_eq
for vector subspaces when ι
is finite.