Documentation

Mathlib.LinearAlgebra.Dual

Dual vector spaces #

The dual space of an R-module M is the R-module of R-linear maps MR.

Main definitions #

Main results #

@[reducible, inline]
abbrev Module.Dual (R : Type uR) (M : Type uM) [CommSemiring R] [AddCommMonoid M] [Module R M] :
Type (max uM uR)

The dual space of an R-module M is the R-module of linear maps M → R.

Equations
def Module.dualPairing (R : Type u_1) (M : Type u_2) [CommSemiring R] [AddCommMonoid M] [Module R M] :

The canonical pairing of a vector space and its algebraic dual.

Equations
@[simp]
theorem Module.dualPairing_apply (R : Type uR) (M : Type uM) [CommSemiring R] [AddCommMonoid M] [Module R M] (v : Dual R M) (x : M) :
((dualPairing R M) v) x = v x
instance Module.Dual.instInhabited (R : Type uR) (M : Type uM) [CommSemiring R] [AddCommMonoid M] [Module R M] :
Equations
def Module.Dual.eval (R : Type uR) (M : Type uM) [CommSemiring R] [AddCommMonoid M] [Module R M] :

Maps a module M to the dual of the dual of M. See Module.erange_coe and Module.evalEquiv.

Equations
@[simp]
theorem Module.Dual.eval_apply (R : Type uR) (M : Type uM) [CommSemiring R] [AddCommMonoid M] [Module R M] (v : M) (a : Dual R M) :
((eval R M) v) a = a v
def Module.Dual.transpose {R : Type uR} {M : Type uM} [CommSemiring R] [AddCommMonoid M] [Module R M] {M' : Type uM'} [AddCommMonoid M'] [Module R M'] :

The transposition of linear maps, as a linear map from M →ₗ[R] M' to Dual R M' →ₗ[R] Dual R M.

Equations
theorem Module.Dual.transpose_apply {R : Type uR} {M : Type uM} [CommSemiring R] [AddCommMonoid M] [Module R M] {M' : Type uM'} [AddCommMonoid M'] [Module R M'] (u : M →ₗ[R] M') (l : Dual R M') :
theorem Module.Dual.transpose_comp {R : Type uR} {M : Type uM} [CommSemiring R] [AddCommMonoid M] [Module R M] {M' : Type uM'} [AddCommMonoid M'] [Module R M'] {M'' : Type uM''} [AddCommMonoid M''] [Module R M''] (u : M' →ₗ[R] M'') (v : M →ₗ[R] M') :
def Module.dualProdDualEquivDual (R : Type uR) (M : Type uM) [CommSemiring R] [AddCommMonoid M] [Module R M] (M' : Type uM') [AddCommMonoid M'] [Module R M'] :
(Dual R M × Dual R M') ≃ₗ[R] Dual R (M × M')

Taking duals distributes over products.

Equations
@[simp]
theorem Module.dualProdDualEquivDual_apply_apply (R : Type uR) (M : Type uM) [CommSemiring R] [AddCommMonoid M] [Module R M] (M' : Type uM') [AddCommMonoid M'] [Module R M'] (f : (M →ₗ[R] R) × (M' →ₗ[R] R)) (a : M × M') :
((dualProdDualEquivDual R M M') f) a = f.1 a.1 + f.2 a.2
@[simp]
theorem Module.dualProdDualEquivDual_symm_apply (R : Type uR) (M : Type uM) [CommSemiring R] [AddCommMonoid M] [Module R M] (M' : Type uM') [AddCommMonoid M'] [Module R M'] (f : M × M' →ₗ[R] R) :
@[simp]
theorem Module.dualProdDualEquivDual_apply (R : Type uR) (M : Type uM) [CommSemiring R] [AddCommMonoid M] [Module R M] (M' : Type uM') [AddCommMonoid M'] [Module R M'] (φ : Dual R M) (ψ : Dual R M') :
def LinearMap.dualMap {R : Type u} [CommSemiring R] {M₁ : Type v} {M₂ : Type v'} [AddCommMonoid M₁] [Module R M₁] [AddCommMonoid M₂] [Module R M₂] (f : M₁ →ₗ[R] M₂) :

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
theorem LinearMap.dualMap_eq_lcomp {R : Type u} [CommSemiring R] {M₁ : Type v} {M₂ : Type v'} [AddCommMonoid M₁] [Module R M₁] [AddCommMonoid M₂] [Module R M₂] (f : M₁ →ₗ[R] M₂) :
theorem LinearMap.dualMap_def {R : Type u} [CommSemiring R] {M₁ : Type v} {M₂ : Type v'} [AddCommMonoid M₁] [Module R M₁] [AddCommMonoid M₂] [Module R M₂] (f : M₁ →ₗ[R] M₂) :
theorem LinearMap.dualMap_apply' {R : Type u} [CommSemiring R] {M₁ : Type v} {M₂ : Type v'} [AddCommMonoid M₁] [Module R M₁] [AddCommMonoid M₂] [Module R M₂] (f : M₁ →ₗ[R] M₂) (g : Module.Dual R M₂) :
@[simp]
theorem LinearMap.dualMap_apply {R : Type u} [CommSemiring R] {M₁ : Type v} {M₂ : Type v'} [AddCommMonoid M₁] [Module R M₁] [AddCommMonoid M₂] [Module R M₂] (f : M₁ →ₗ[R] M₂) (g : Module.Dual R M₂) (x : M₁) :
(f.dualMap g) x = g (f x)
@[simp]
theorem LinearMap.dualMap_id {R : Type u} [CommSemiring R] {M₁ : Type v} [AddCommMonoid M₁] [Module R M₁] :
theorem LinearMap.dualMap_comp_dualMap {R : Type u} [CommSemiring R] {M₁ : Type v} {M₂ : Type v'} [AddCommMonoid M₁] [Module R M₁] [AddCommMonoid M₂] [Module R M₂] {M₃ : Type u_1} [AddCommGroup M₃] [Module R M₃] (f : M₁ →ₗ[R] M₂) (g : M₂ →ₗ[R] M₃) :
theorem LinearMap.dualMap_injective_of_surjective {R : Type u} [CommSemiring R] {M₁ : Type v} {M₂ : Type v'} [AddCommMonoid M₁] [Module R M₁] [AddCommMonoid M₂] [Module R M₂] {f : M₁ →ₗ[R] M₂} (hf : Function.Surjective f) :

If a linear map is surjective, then its dual is injective.

def LinearEquiv.dualMap {R : Type u} [CommSemiring R] {M₁ : Type v} {M₂ : Type v'} [AddCommMonoid M₁] [Module R M₁] [AddCommMonoid M₂] [Module R M₂] (f : M₁ ≃ₗ[R] M₂) :

The Linear_equiv version of LinearMap.dualMap.

Equations
@[simp]
theorem LinearEquiv.dualMap_apply {R : Type u} [CommSemiring R] {M₁ : Type v} {M₂ : Type v'} [AddCommMonoid M₁] [Module R M₁] [AddCommMonoid M₂] [Module R M₂] (f : M₁ ≃ₗ[R] M₂) (g : Module.Dual R M₂) (x : M₁) :
(f.dualMap g) x = g (f x)
@[simp]
theorem LinearEquiv.dualMap_refl {R : Type u} [CommSemiring R] {M₁ : Type v} [AddCommMonoid M₁] [Module R M₁] :
(refl R M₁).dualMap = refl R (Module.Dual R M₁)
@[simp]
theorem LinearEquiv.dualMap_symm {R : Type u} [CommSemiring R] {M₁ : Type v} {M₂ : Type v'} [AddCommMonoid M₁] [Module R M₁] [AddCommMonoid M₂] [Module R M₂] {f : M₁ ≃ₗ[R] M₂} :
theorem LinearEquiv.dualMap_trans {R : Type u} [CommSemiring R] {M₁ : Type v} {M₂ : Type v'} [AddCommMonoid M₁] [Module R M₁] [AddCommMonoid M₂] [Module R M₂] {M₃ : Type u_1} [AddCommGroup M₃] [Module R M₃] (f : M₁ ≃ₗ[R] M₂) (g : M₂ ≃ₗ[R] M₃) :
theorem Module.Dual.eval_naturality {R : Type u} [CommSemiring R] {M₁ : Type v} {M₂ : Type v'} [AddCommMonoid M₁] [Module R M₁] [AddCommMonoid M₂] [Module R M₂] (f : M₁ →ₗ[R] M₂) :
@[simp]
theorem Dual.apply_one_mul_eq {R : Type u} [CommSemiring R] (f : Module.Dual R R) (r : R) :
f 1 * r = f r
def Basis.toDual {R : Type uR} {M : Type uM} {ι : Type uι} [CommSemiring R] [AddCommMonoid M] [Module R M] [DecidableEq ι] (b : Basis ι R M) :

The linear map from a vector space equipped with basis to its dual vector space, taking basis elements to corresponding dual basis elements.

Equations
theorem Basis.toDual_apply {R : Type uR} {M : Type uM} {ι : Type uι} [CommSemiring R] [AddCommMonoid M] [Module R M] [DecidableEq ι] (b : Basis ι R M) (i j : ι) :
(b.toDual (b i)) (b j) = if i = j then 1 else 0
@[simp]
theorem Basis.toDual_linearCombination_left {R : Type uR} {M : Type uM} {ι : Type uι} [CommSemiring R] [AddCommMonoid M] [Module R M] [DecidableEq ι] (b : Basis ι R M) (f : ι →₀ R) (i : ι) :
(b.toDual ((Finsupp.linearCombination R b) f)) (b i) = f i
@[deprecated Basis.toDual_linearCombination_left (since := "2024-08-29")]
theorem Basis.toDual_total_left {R : Type uR} {M : Type uM} {ι : Type uι} [CommSemiring R] [AddCommMonoid M] [Module R M] [DecidableEq ι] (b : Basis ι R M) (f : ι →₀ R) (i : ι) :
(b.toDual ((Finsupp.linearCombination R b) f)) (b i) = f i

Alias of Basis.toDual_linearCombination_left.

@[simp]
theorem Basis.toDual_linearCombination_right {R : Type uR} {M : Type uM} {ι : Type uι} [CommSemiring R] [AddCommMonoid M] [Module R M] [DecidableEq ι] (b : Basis ι R M) (f : ι →₀ R) (i : ι) :
(b.toDual (b i)) ((Finsupp.linearCombination R b) f) = f i
@[deprecated Basis.toDual_linearCombination_right (since := "2024-08-29")]
theorem Basis.toDual_total_right {R : Type uR} {M : Type uM} {ι : Type uι} [CommSemiring R] [AddCommMonoid M] [Module R M] [DecidableEq ι] (b : Basis ι R M) (f : ι →₀ R) (i : ι) :
(b.toDual (b i)) ((Finsupp.linearCombination R b) f) = f i

Alias of Basis.toDual_linearCombination_right.

theorem Basis.toDual_apply_left {R : Type uR} {M : Type uM} {ι : Type uι} [CommSemiring R] [AddCommMonoid M] [Module R M] [DecidableEq ι] (b : Basis ι R M) (m : M) (i : ι) :
(b.toDual m) (b i) = (b.repr m) i
theorem Basis.toDual_apply_right {R : Type uR} {M : Type uM} {ι : Type uι} [CommSemiring R] [AddCommMonoid M] [Module R M] [DecidableEq ι] (b : Basis ι R M) (i : ι) (m : M) :
(b.toDual (b i)) m = (b.repr m) i
theorem Basis.coe_toDual_self {R : Type uR} {M : Type uM} {ι : Type uι} [CommSemiring R] [AddCommMonoid M] [Module R M] [DecidableEq ι] (b : Basis ι R M) (i : ι) :
b.toDual (b i) = b.coord i
def Basis.toDualFlip {R : Type uR} {M : Type uM} {ι : Type uι} [CommSemiring R] [AddCommMonoid M] [Module R M] [DecidableEq ι] (b : Basis ι R M) (m : M) :

h.toDual_flip v is the linear map sending w to h.toDual w v.

Equations
theorem Basis.toDualFlip_apply {R : Type uR} {M : Type uM} {ι : Type uι} [CommSemiring R] [AddCommMonoid M] [Module R M] [DecidableEq ι] (b : Basis ι R M) (m₁ m₂ : M) :
(b.toDualFlip m₁) m₂ = (b.toDual m₂) m₁
theorem Basis.toDual_eq_repr {R : Type uR} {M : Type uM} {ι : Type uι} [CommSemiring R] [AddCommMonoid M] [Module R M] [DecidableEq ι] (b : Basis ι R M) (m : M) (i : ι) :
(b.toDual m) (b i) = (b.repr m) i
theorem Basis.toDual_eq_equivFun {R : Type uR} {M : Type uM} {ι : Type uι} [CommSemiring R] [AddCommMonoid M] [Module R M] [DecidableEq ι] (b : Basis ι R M) [Finite ι] (m : M) (i : ι) :
(b.toDual m) (b i) = b.equivFun m i
theorem Basis.toDual_injective {R : Type uR} {M : Type uM} {ι : Type uι} [CommSemiring R] [AddCommMonoid M] [Module R M] [DecidableEq ι] (b : Basis ι R M) :
theorem Basis.toDual_inj {R : Type uR} {M : Type uM} {ι : Type uι} [CommSemiring R] [AddCommMonoid M] [Module R M] [DecidableEq ι] (b : Basis ι R M) (m : M) (a : b.toDual m = 0) :
m = 0
theorem Basis.toDual_ker {R : Type uR} {M : Type uM} {ι : Type uι} [CommSemiring R] [AddCommMonoid M] [Module R M] [DecidableEq ι] (b : Basis ι R M) :
theorem Basis.toDual_range {R : Type uR} {M : Type uM} {ι : Type uι} [CommSemiring R] [AddCommMonoid M] [Module R M] [DecidableEq ι] (b : Basis ι R M) [Finite ι] :
@[simp]
theorem Basis.sum_dual_apply_smul_coord {R : Type uR} {M : Type uM} {ι : Type uι} [CommSemiring R] [AddCommMonoid M] [Module R M] [Fintype ι] (b : Basis ι R M) (f : Module.Dual R M) :
def Basis.toDualEquiv {R : Type uR} {M : Type uM} {ι : Type uι} [CommRing R] [AddCommGroup M] [Module R M] [DecidableEq ι] (b : Basis ι R M) [Finite ι] :

A vector space is linearly equivalent to its dual space.

Equations
@[simp]
theorem Basis.toDualEquiv_apply {R : Type uR} {M : Type uM} {ι : Type uι} [CommRing R] [AddCommGroup M] [Module R M] [DecidableEq ι] (b : Basis ι R M) [Finite ι] (m : M) :

A vector space over a field is isomorphic to its dual if and only if it is finite-dimensional: a consequence of the Erdős-Kaplansky theorem.

def Basis.dualBasis {R : Type uR} {M : Type uM} {ι : Type uι} [CommRing R] [AddCommGroup M] [Module R M] [DecidableEq ι] (b : Basis ι R M) [Finite ι] :
Basis ι R (Module.Dual R M)

Maps a basis for V to a basis for the dual space.

Equations
theorem Basis.dualBasis_apply_self {R : Type uR} {M : Type uM} {ι : Type uι} [CommRing R] [AddCommGroup M] [Module R M] [DecidableEq ι] (b : Basis ι R M) [Finite ι] (i j : ι) :
(b.dualBasis i) (b j) = if j = i then 1 else 0
theorem Basis.linearCombination_dualBasis {R : Type uR} {M : Type uM} {ι : Type uι} [CommRing R] [AddCommGroup M] [Module R M] [DecidableEq ι] (b : Basis ι R M) [Finite ι] (f : ι →₀ R) (i : ι) :
@[deprecated Basis.linearCombination_dualBasis (since := "2024-08-29")]
theorem Basis.total_dualBasis {R : Type uR} {M : Type uM} {ι : Type uι} [CommRing R] [AddCommGroup M] [Module R M] [DecidableEq ι] (b : Basis ι R M) [Finite ι] (f : ι →₀ R) (i : ι) :

Alias of Basis.linearCombination_dualBasis.

@[simp]
theorem Basis.dualBasis_repr {R : Type uR} {M : Type uM} {ι : Type uι} [CommRing R] [AddCommGroup M] [Module R M] [DecidableEq ι] (b : Basis ι R M) [Finite ι] (l : Module.Dual R M) (i : ι) :
(b.dualBasis.repr l) i = l (b i)
theorem Basis.dualBasis_apply {R : Type uR} {M : Type uM} {ι : Type uι} [CommRing R] [AddCommGroup M] [Module R M] [DecidableEq ι] (b : Basis ι R M) [Finite ι] (i : ι) (m : M) :
(b.dualBasis i) m = (b.repr m) i
@[simp]
theorem Basis.coe_dualBasis {R : Type uR} {M : Type uM} {ι : Type uι} [CommRing R] [AddCommGroup M] [Module R M] [DecidableEq ι] (b : Basis ι R M) [Finite ι] :
@[simp]
theorem Basis.toDual_toDual {R : Type uR} {M : Type uM} {ι : Type uι} [CommRing R] [AddCommGroup M] [Module R M] [DecidableEq ι] (b : Basis ι R M) [Finite ι] :
theorem Basis.dualBasis_equivFun {R : Type uR} {M : Type uM} {ι : Type uι} [CommRing R] [AddCommGroup M] [Module R M] [DecidableEq ι] (b : Basis ι R M) [Finite ι] (l : Module.Dual R M) (i : ι) :
b.dualBasis.equivFun l i = l (b i)
theorem Basis.eval_ker {R : Type uR} {M : Type uM} [CommRing R] [AddCommGroup M] [Module R M] {ι : Type u_1} (b : Basis ι R M) :
theorem Basis.eval_range {R : Type uR} {M : Type uM} [CommRing R] [AddCommGroup M] [Module R M] {ι : Type u_1} [Finite ι] (b : Basis ι R M) :
instance Basis.dual_free {R : Type uR} {M : Type uM} [CommRing R] [AddCommGroup M] [Module R M] [Module.Finite R M] [Module.Free R M] :
@[simp]
theorem Basis.linearCombination_coord {R : Type uR} {M : Type uM} {ι : Type uι} [CommRing R] [AddCommGroup M] [Module R M] [Finite ι] (b : Basis ι R M) (f : ι →₀ R) (i : ι) :

simp normal form version of linearCombination_dualBasis

@[deprecated Basis.linearCombination_coord (since := "2024-08-29")]
theorem Basis.total_coord {R : Type uR} {M : Type uM} {ι : Type uι} [CommRing R] [AddCommGroup M] [Module R M] [Finite ι] (b : Basis ι R M) (f : ι →₀ R) (i : ι) :

Alias of Basis.linearCombination_coord.


simp normal form version of linearCombination_dualBasis

theorem Basis.dual_rank_eq {K : Type uK} {V : Type uV} {ι : Type uι} [CommRing K] [AddCommGroup V] [Module K V] [Finite ι] (b : Basis ι K V) :
theorem Module.eval_ker (K : Type uK) (V : Type uV) [CommRing K] [AddCommGroup V] [Module K V] [Projective K V] :
theorem Module.eval_apply_eq_zero_iff (K : Type uK) {V : Type uV} [CommRing K] [AddCommGroup V] [Module K V] [Projective K V] (v : V) :
(Dual.eval K V) v = 0 v = 0
theorem Module.forall_dual_apply_eq_zero_iff (K : Type uK) {V : Type uV} [CommRing K] [AddCommGroup V] [Module K V] [Projective K V] (v : V) :
@[simp]
@[simp]
theorem Module.nontrivial_dual_iff (K : Type uK) {V : Type uV} [CommRing K] [AddCommGroup V] [Module K V] [Projective K V] :
instance Module.instNontrivialDual (K : Type uK) {V : Type uV} [CommRing K] [AddCommGroup V] [Module K V] [Projective K V] [Nontrivial V] :
theorem Module.finite_dual_iff (K : Type uK) {V : Type uV} [CommRing K] [AddCommGroup V] [Module K V] [Free K V] :
class Module.IsReflexive (R : Type u_1) (M : Type u_2) [CommRing R] [AddCommGroup M] [Module R M] :

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
    @[instance 900]
    instance Module.IsReflexive.of_finite_of_free (R : Type u_1) (M : Type u_2) [CommRing R] [AddCommGroup M] [Module R M] [Module.Finite R M] [Free R M] :

    See also Module.instFiniteDimensionalOfIsReflexive for the converse over a field.

    theorem Module.erange_coe (R : Type u_1) (M : Type u_2) [CommRing R] [AddCommGroup M] [Module R M] [IsReflexive R M] :
    def Module.evalEquiv (R : Type u_1) (M : Type u_2) [CommRing R] [AddCommGroup M] [Module R M] [IsReflexive R M] :

    The bijection between a reflexive module and its double dual, bundled as a LinearEquiv.

    Equations
    @[simp]
    theorem Module.evalEquiv_toLinearMap (R : Type u_1) (M : Type u_2) [CommRing R] [AddCommGroup M] [Module R M] [IsReflexive R M] :
    @[simp]
    theorem Module.evalEquiv_apply (R : Type u_1) (M : Type u_2) [CommRing R] [AddCommGroup M] [Module R M] [IsReflexive R M] (m : M) :
    (evalEquiv R M) m = (Dual.eval R M) m
    @[simp]
    theorem Module.apply_evalEquiv_symm_apply (R : Type u_1) (M : Type u_2) [CommRing R] [AddCommGroup M] [Module R M] [IsReflexive R M] (f : Dual R M) (g : Dual R (Dual R M)) :
    f ((evalEquiv R M).symm g) = g f
    @[simp]
    @[simp]
    theorem Module.Dual.eval_comp_comp_evalEquiv_eq (R : Type u_1) (M : Type u_2) [CommRing R] [AddCommGroup M] [Module R M] [IsReflexive R M] {M' : Type u_4} [AddCommGroup M'] [Module R M'] {f : M →ₗ[R] M'} :
    @[simp]
    theorem Module.dualMap_dualMap_eq_iff (R : Type u_1) (M : Type u_2) [CommRing R] [AddCommGroup M] [Module R M] [IsReflexive R M] {M' : Type u_4} [AddCommGroup M'] [Module R M'] [IsReflexive R M'] {f g : M →ₗ[R] M'} :
    instance Module.Dual.instIsReflecive (R : Type u_1) (M : Type u_2) [CommRing R] [AddCommGroup M] [Module R M] [IsReflexive R M] :

    The dual of a reflexive module is reflexive.

    theorem Module.IsReflexive.of_split {R : Type u_1} {M : Type u_2} {N : Type u_3} [CommRing R] [AddCommGroup M] [AddCommGroup N] [Module R M] [Module R N] [IsReflexive R M] (i : N →ₗ[R] M) (s : M →ₗ[R] N) (H : s ∘ₗ i = LinearMap.id) :

    A direct summand of a reflexive module is reflexive.

    @[instance 900]
    def Module.mapEvalEquiv (R : Type u_1) (M : Type u_2) [CommRing R] [AddCommGroup M] [Module R M] [IsReflexive R M] :
    Submodule R M ≃o Submodule R (Dual R (Dual R M))

    The isomorphism Module.evalEquiv induces an order isomorphism on subspaces.

    Equations
    @[simp]
    theorem Module.mapEvalEquiv_apply (R : Type u_1) (M : Type u_2) [CommRing R] [AddCommGroup M] [Module R M] [IsReflexive R M] (W : Submodule R M) :
    @[simp]
    theorem Module.mapEvalEquiv_symm_apply (R : Type u_1) (M : Type u_2) [CommRing R] [AddCommGroup M] [Module R M] [IsReflexive R M] (W'' : Submodule R (Dual R (Dual R M))) :
    theorem Module.equiv {R : Type u_1} {M : Type u_2} {N : Type u_3} [CommRing R] [AddCommGroup M] [AddCommGroup N] [Module R M] [Module R N] [IsReflexive R M] (e : M ≃ₗ[R] N) :
    @[instance 100]
    theorem Submodule.exists_dual_map_eq_bot_of_nmem {R : Type u_1} {M : Type u_2} [CommRing R] [AddCommGroup M] [Module R M] {p : Submodule R M} {x : M} (hx : xp) (hp' : Module.Free R (M p)) :
    ∃ (f : Module.Dual R M), f x 0 map f p =
    theorem Submodule.exists_dual_map_eq_bot_of_lt_top {R : Type u_1} {M : Type u_2} [CommRing R] [AddCommGroup M] [Module R M] {p : Submodule R M} (hp : p < ) (hp' : Module.Free R (M p)) :
    ∃ (f : Module.Dual R M), f 0 map f p =
    theorem Submodule.span_eq_top_of_ne_zero {R : Type u_1} {M : Type u_2} [CommRing R] [AddCommGroup M] [Module R M] [Module.IsReflexive R M] {s : Set (M →ₗ[R] R)} [Module.Free R ((M →ₗ[R] R) span R s)] (h : ∀ (z : M), z 0fs, f z 0) :
    span R s =

    Consider a reflexive module and a set s of linear forms. If for any z ≠ 0 there exists f ∈ s such that f z ≠ 0, then s spans the whole dual space.

    theorem FiniteDimensional.mem_span_of_iInf_ker_le_ker {ι : Type u_3} {𝕜 : Type u_4} {E : Type u_5} [Field 𝕜] [AddCommGroup E] [Module 𝕜 E] [FiniteDimensional 𝕜 E] {L : ιE →ₗ[𝕜] 𝕜} {K : E →ₗ[𝕜] 𝕜} (h : ⨅ (i : ι), LinearMap.ker (L i) LinearMap.ker K) :
    theorem mem_span_of_iInf_ker_le_ker {ι : Type u_3} {𝕜 : Type u_4} {E : Type u_5} [Field 𝕜] [AddCommGroup E] [Module 𝕜 E] [Finite ι] {L : ιE →ₗ[𝕜] 𝕜} {K : E →ₗ[𝕜] 𝕜} (h : ⨅ (i : ι), LinearMap.ker (L i) LinearMap.ker K) :

    Given some linear forms L1,...,Ln,K over a vector space E, if i=1nker(Li)ker(K), then K is in the space generated by L1,...,Ln.

    Try using Set.toFinite to dispatch a Set.Finite goal.

    Equations
    • One or more equations did not get rendered due to their size.
    structure Module.DualBases {R : Type u_1} {M : Type u_2} {ι : Type u_3} [CommSemiring R] [AddCommMonoid M] [Module R M] (e : ιM) (ε : ιDual R M) :

    e and ε have characteristic properties of a basis and its dual

    def Module.DualBases.coeffs {R : Type u_1} {M : Type u_2} {ι : Type u_3} [CommRing R] [AddCommGroup M] [Module R M] {e : ιM} {ε : ιDual R M} (h : DualBases e ε) (m : M) :

    The coefficients of v on the basis e

    Equations
    • h.coeffs m = { support := .toFinset, toFun := fun (i : ι) => (ε i) m, mem_support_toFun := }
    @[simp]
    theorem Module.DualBases.coeffs_apply {R : Type u_1} {M : Type u_2} {ι : Type u_3} [CommRing R] [AddCommGroup M] [Module R M] {e : ιM} {ε : ιDual R M} (h : DualBases e ε) (m : M) (i : ι) :
    (h.coeffs m) i = (ε i) m
    def Module.DualBases.lc {R : Type u_1} {M : Type u_2} [CommRing R] [AddCommGroup M] [Module R M] {ι : Type u_4} (e : ιM) (l : ι →₀ R) :
    M

    linear combinations of elements of e. This is a convenient abbreviation for Finsupp.linearCombination R e l

    Equations
    theorem Module.DualBases.lc_def {R : Type u_1} {M : Type u_2} {ι : Type u_3} [CommRing R] [AddCommGroup M] [Module R M] (e : ιM) (l : ι →₀ R) :
    theorem Module.DualBases.dual_lc {R : Type u_1} {M : Type u_2} {ι : Type u_3} [CommRing R] [AddCommGroup M] [Module R M] {e : ιM} {ε : ιDual R M} (h : DualBases e ε) (l : ι →₀ R) (i : ι) :
    (ε i) (lc e l) = l i
    @[simp]
    theorem Module.DualBases.coeffs_lc {R : Type u_1} {M : Type u_2} {ι : Type u_3} [CommRing R] [AddCommGroup M] [Module R M] {e : ιM} {ε : ιDual R M} (h : DualBases e ε) (l : ι →₀ R) :
    h.coeffs (lc e l) = l
    @[simp]
    theorem Module.DualBases.lc_coeffs {R : Type u_1} {M : Type u_2} {ι : Type u_3} [CommRing R] [AddCommGroup M] [Module R M] {e : ιM} {ε : ιDual R M} (h : DualBases e ε) (m : M) :
    lc e (h.coeffs m) = m

    For any m : M n, \sum_{p ∈ Q n} (ε p m) • e p = m

    def Module.DualBases.basis {R : Type u_1} {M : Type u_2} {ι : Type u_3} [CommRing R] [AddCommGroup M] [Module R M] {e : ιM} {ε : ιDual R M} (h : DualBases e ε) :
    Basis ι R M

    (h : DualBases e ε).basis shows the family of vectors e forms a basis.

    Equations
    @[simp]
    theorem Module.DualBases.basis_repr_apply {R : Type u_1} {M : Type u_2} {ι : Type u_3} [CommRing R] [AddCommGroup M] [Module R M] {e : ιM} {ε : ιDual R M} (h : DualBases e ε) (m : M) :
    theorem Module.DualBases.basis_repr_symm_apply {R : Type u_1} {M : Type u_2} {ι : Type u_3} [CommRing R] [AddCommGroup M] [Module R M] {e : ιM} {ε : ιDual R M} (h : DualBases e ε) (l : ι →₀ R) :
    @[simp]
    theorem Module.DualBases.coe_basis {R : Type u_1} {M : Type u_2} {ι : Type u_3} [CommRing R] [AddCommGroup M] [Module R M] {e : ιM} {ε : ιDual R M} (h : DualBases e ε) :
    theorem Module.DualBases.mem_of_mem_span {R : Type u_1} {M : Type u_2} {ι : Type u_3} [CommRing R] [AddCommGroup M] [Module R M] {e : ιM} {ε : ιDual R M} (h : DualBases e ε) {H : Set ι} {x : M} (hmem : x Submodule.span R (e '' H)) (i : ι) :
    theorem Module.DualBases.coe_dualBasis {R : Type u_1} {M : Type u_2} {ι : Type u_3} [CommRing R] [AddCommGroup M] [Module R M] {e : ιM} {ε : ιDual R M} (h : DualBases e ε) [DecidableEq ι] [Finite ι] :

    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
    @[simp]
    theorem Submodule.dualRestrict_apply {R : Type u} {M : Type v} [CommSemiring R] [AddCommMonoid M] [Module R M] (W : Submodule R M) (φ : Module.Dual R M) (x : W) :
    (W.dualRestrict φ) x = φ x

    The dualAnnihilator of a submodule W is the set of linear maps φ such that φ w = 0 for all w ∈ W.

    Equations
    @[simp]
    theorem Submodule.mem_dualAnnihilator {R : Type u} {M : Type v} [CommSemiring R] [AddCommMonoid M] [Module R M] {W : Submodule R M} (φ : Module.Dual R M) :
    φ W.dualAnnihilator wW, φ w = 0

    That ker(ι:VW)=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
    @[simp]
    theorem Submodule.mem_dualCoannihilator {R : Type u} {M : Type v} [CommSemiring R] [AddCommMonoid M] [Module R M] {Φ : Submodule R (Module.Dual R M)} (x : M) :
    x Φ.dualCoannihilator φΦ, φ x = 0
    theorem Submodule.dualAnnihilator_iSup_eq {R : Type u} {M : Type v} [CommSemiring R] [AddCommMonoid M] [Module R M] {ι : Sort u_1} (U : ιSubmodule R M) :
    theorem Submodule.iSup_dualAnnihilator_le_iInf {R : Type u} {M : Type v} [CommSemiring R] [AddCommMonoid M] [Module R M] {ι : Sort u_1} (U : ιSubmodule R M) :

    See also Subspace.dualAnnihilator_iInf_eq for vector subspaces when ι is finite.

    @[simp]
    theorem Submodule.coe_dualCoannihilator_span {R : Type u} {M : Type v} [CommSemiring R] [AddCommMonoid M] [Module R M] (s : Set (Module.Dual R M)) :
    (span R s).dualCoannihilator = {x : M | fs, f x = 0}

    Submodule.dualAnnihilator and Submodule.dualCoannihilator form a Galois coinsertion.

    Equations
    • One or more equations did not get rendered due to their size.
    noncomputable def Subspace.dualLift {K : Type u} {V : Type v} [Field K] [AddCommGroup V] [Module K V] (W : Subspace K V) :

    Given a subspace W of V and an element of its dual φ, dualLift W φ is an arbitrary extension of φ to an element of the dual of V. That is, dualLift W φ sends w ∈ W to φ x and x in a chosen complement of W to 0.

    Equations
    @[simp]
    theorem Subspace.dualLift_of_subtype {K : Type u} {V : Type v} [Field K] [AddCommGroup V] [Module K V] {W : Subspace K V} {φ : Module.Dual K W} (w : W) :
    (W.dualLift φ) w = φ w
    theorem Subspace.dualLift_of_mem {K : Type u} {V : Type v} [Field K] [AddCommGroup V] [Module K V] {W : Subspace K V} {φ : Module.Dual K W} {w : V} (hw : w W) :
    (W.dualLift φ) w = φ w, hw
    noncomputable def Subspace.quotAnnihilatorEquiv {K : Type u} {V : Type v} [Field K] [AddCommGroup V] [Module K V] (W : Subspace K V) :

    The quotient by the dualAnnihilator of a subspace is isomorphic to the dual of that subspace.

    Equations
    noncomputable def Subspace.dualEquivDual {K : Type u} {V : Type v} [Field K] [AddCommGroup V] [Module K V] (W : Subspace K V) :

    The natural isomorphism from the dual of a subspace W to W.dualLift.range.

    Equations
    @[simp]
    theorem Subspace.dualEquivDual_apply {K : Type u} {V : Type v} [Field K] [AddCommGroup V] [Module K V] {W : Subspace K V} (φ : Module.Dual K W) :
    W.dualEquivDual φ = W.dualLift φ,
    @[simp]
    noncomputable def Subspace.quotEquivAnnihilator {K : Type u} {V : Type v} [Field K] [AddCommGroup V] [Module K V] [FiniteDimensional K V] (W : Subspace K V) :

    The quotient by a subspace is isomorphic to its dual annihilator.

    Equations
    • One or more equations did not get rendered due to their size.
    theorem LinearMap.ker_dualMap_eq_dualAnnihilator_range {R : Type uR} [CommSemiring R] {M₁ : Type uM₁} {M₂ : Type uM₂} [AddCommMonoid M₁] [Module R M₁] [AddCommMonoid M₂] [Module R M₂] (f : M₁ →ₗ[R] M₂) :
    theorem LinearMap.range_dualMap_le_dualAnnihilator_ker {R : Type uR} [CommSemiring R] {M₁ : Type uM₁} {M₂ : Type uM₂} [AddCommMonoid M₁] [Module R M₁] [AddCommMonoid M₂] [Module R M₂] (f : M₁ →ₗ[R] M₂) :

    Given a submodule, corestrict to the pairing on M ⧸ W by simultaneously restricting to W.dualAnnihilator.

    See Subspace.dualCopairing_nondegenerate.

    Equations
    Equations
    @[simp]
    theorem Submodule.dualCopairing_apply {R : Type u_1} {M : Type u_2} [CommRing R] [AddCommGroup M] [Module R M] {W : Submodule R M} (φ : W.dualAnnihilator) (x : M) :

    Given a submodule, restrict to the pairing on W by simultaneously corestricting to Module.Dual R M ⧸ W.dualAnnihilator. This is Submodule.dualRestrict factored through the quotient by its kernel (which is W.dualAnnihilator by definition).

    See Subspace.dualPairing_nondegenerate.

    Equations
    @[simp]
    theorem Submodule.dualPairing_apply {R : Type u_1} {M : Type u_2} [CommRing R] [AddCommGroup M] [Module R M] {W : Submodule R M} (φ : Module.Dual R M) (x : W) :
    (W.dualPairing (Quotient.mk φ)) x = φ x

    That im(q:(V/W)V)=ann(W).

    Equivalence (M/W)ann(W). That is, there is a one-to-one correspondence between the dual of M ⧸ W and those elements of the dual of M that vanish on W.

    The inverse of this is Submodule.dualCopairing.

    Equations
    @[simp]

    The pairing between a submodule W of a dual module Dual R M and the quotient of M by the coannihilator of W, which is always nondegenerate.

    Equations
    @[simp]
    theorem Submodule.quotDualCoannihilatorToDual_apply {R : Type u_1} {M : Type u_2} [CommRing R] [AddCommGroup M] [Module R M] (W : Submodule R (Module.Dual R M)) (m : M) (w : W) :
    theorem Module.Dual.range_eq_top_of_ne_zero {K : Type uK} [Field K] {V₁ : Type uV₁} [AddCommGroup V₁] [Module K V₁] {f : Dual K V₁} (hf : f 0) :
    theorem Module.Dual.finrank_ker_add_one_of_ne_zero {K : Type uK} [Field K] {V₁ : Type uV₁} [AddCommGroup V₁] [Module K V₁] {f : Dual K V₁} (hf : f 0) [FiniteDimensional K V₁] :
    finrank K (LinearMap.ker f) + 1 = finrank K V₁
    theorem Module.Dual.isCompl_ker_of_disjoint_of_ne_bot {K : Type uK} [Field K] {V₁ : Type uV₁} [AddCommGroup V₁] [Module K V₁] {f : Dual K V₁} (hf : f 0) [FiniteDimensional K V₁] {p : Submodule K V₁} (hpf : Disjoint (LinearMap.ker f) p) (hp : p ) :
    theorem Module.Dual.eq_of_ker_eq_of_apply_eq {K : Type uK} [Field K] {V₁ : Type uV₁} [AddCommGroup V₁] [Module K V₁] [FiniteDimensional K V₁] {f g : Dual K V₁} (x : V₁) (h : LinearMap.ker f = LinearMap.ker g) (h' : f x = g x) (hx : f x 0) :
    f = g
    theorem LinearMap.dualPairing_nondegenerate {K : Type uK} [Field K] {V₁ : Type uV₁} [AddCommGroup V₁] [Module K V₁] :
    theorem LinearMap.dualMap_surjective_of_injective {K : Type uK} [Field K] {V₁ : Type uV₁} {V₂ : Type uV₂} [AddCommGroup V₁] [Module K V₁] [AddCommGroup V₂] [Module K V₂] {f : V₁ →ₗ[K] V₂} (hf : Function.Injective f) :
    theorem LinearMap.range_dualMap_eq_dualAnnihilator_ker {K : Type uK} [Field K] {V₁ : Type uV₁} {V₂ : Type uV₂} [AddCommGroup V₁] [Module K V₁] [AddCommGroup V₂] [Module K V₂] (f : V₁ →ₗ[K] V₂) :
    @[simp]
    theorem LinearMap.dualMap_surjective_iff {K : Type uK} [Field K] {V₁ : Type uV₁} {V₂ : Type uV₂} [AddCommGroup V₁] [Module K V₁] [AddCommGroup V₂] [Module K V₂] {f : V₁ →ₗ[K] V₂} :

    For vector spaces, f.dualMap is surjective if and only if f is injective

    theorem Subspace.dualPairing_eq {K : Type uK} [Field K] {V₁ : Type uV₁} [AddCommGroup V₁] [Module K V₁] (W : Subspace K V₁) :
    theorem Subspace.dualPairing_nondegenerate {K : Type uK} [Field K] {V₁ : Type uV₁} [AddCommGroup V₁] [Module K V₁] (W : Subspace K V₁) :
    theorem Subspace.dualCopairing_nondegenerate {K : Type uK} [Field K] {V₁ : Type uV₁} [AddCommGroup V₁] [Module K V₁] (W : Subspace K V₁) :
    theorem Subspace.dualAnnihilator_iInf_eq {K : Type uK} [Field K] {V₁ : Type uV₁} [AddCommGroup V₁] [Module K V₁] {ι : Type u_1} [Finite ι] (W : ιSubspace K V₁) :
    Submodule.dualAnnihilator (⨅ (i : ι), W i) = ⨆ (i : ι), Submodule.dualAnnihilator (W i)
    theorem Subspace.isCompl_dualAnnihilator {K : Type uK} [Field K] {V₁ : Type uV₁} [AddCommGroup V₁] [Module K V₁] {W W' : Subspace K V₁} (h : IsCompl W W') :

    For vector spaces, dual annihilators carry direct sum decompositions to direct sum decompositions.

    def Subspace.dualQuotDistrib {K : Type uK} [Field K] {V₁ : Type uV₁} [AddCommGroup V₁] [Module K V₁] [FiniteDimensional K V₁] (W : Subspace K V₁) :

    For finite-dimensional vector spaces, one can distribute duals over quotients by identifying W.dualLift.range with W. Note that this depends on a choice of splitting of V₁.

    Equations
    @[simp]
    theorem LinearMap.finrank_range_dualMap_eq_finrank_range {K : Type uK} [Field K] {V₁ : Type uV₁} {V₂ : Type uV₂} [AddCommGroup V₁] [Module K V₁] [AddCommGroup V₂] [Module K V₂] (f : V₁ →ₗ[K] V₂) :
    @[simp]
    theorem LinearMap.dualMap_injective_iff {K : Type uK} [Field K] {V₁ : Type uV₁} {V₂ : Type uV₂} [AddCommGroup V₁] [Module K V₁] [AddCommGroup V₂] [Module K V₂] {f : V₁ →ₗ[K] V₂} :

    f.dualMap is injective if and only if f is surjective

    @[simp]
    theorem LinearMap.dualMap_bijective_iff {K : Type uK} [Field K] {V₁ : Type uV₁} {V₂ : Type uV₂} [AddCommGroup V₁] [Module K V₁] [AddCommGroup V₂] [Module K V₂] {f : V₁ →ₗ[K] V₂} :

    f.dualMap is bijective if and only if f is

    @[simp]
    theorem LinearMap.dualAnnihilator_ker_eq_range_flip {K : Type uK} [Field K] {V₁ : Type uV₁} {V₂ : Type uV₂} [AddCommGroup V₁] [Module K V₁] [AddCommGroup V₂] [Module K V₂] {B : V₁ →ₗ[K] V₂ →ₗ[K] K} [Module.IsReflexive K V₂] :
    theorem LinearMap.flip_injective_iff₁ {K : Type uK} [Field K] {V₁ : Type uV₁} {V₂ : Type uV₂} [AddCommGroup V₁] [Module K V₁] [AddCommGroup V₂] [Module K V₂] {B : V₁ →ₗ[K] V₂ →ₗ[K] K} [FiniteDimensional K V₁] :
    theorem LinearMap.flip_injective_iff₂ {K : Type uK} [Field K] {V₁ : Type uV₁} {V₂ : Type uV₂} [AddCommGroup V₁] [Module K V₁] [AddCommGroup V₂] [Module K V₂] {B : V₁ →ₗ[K] V₂ →ₗ[K] K} [FiniteDimensional K V₂] :
    theorem LinearMap.flip_surjective_iff₁ {K : Type uK} [Field K] {V₁ : Type uV₁} {V₂ : Type uV₂} [AddCommGroup V₁] [Module K V₁] [AddCommGroup V₂] [Module K V₂] {B : V₁ →ₗ[K] V₂ →ₗ[K] K} [FiniteDimensional K V₁] :
    theorem LinearMap.flip_surjective_iff₂ {K : Type uK} [Field K] {V₁ : Type uV₁} {V₂ : Type uV₂} [AddCommGroup V₁] [Module K V₁] [AddCommGroup V₂] [Module K V₂] {B : V₁ →ₗ[K] V₂ →ₗ[K] K} [FiniteDimensional K V₂] :
    theorem LinearMap.flip_bijective_iff₁ {K : Type uK} [Field K] {V₁ : Type uV₁} {V₂ : Type uV₂} [AddCommGroup V₁] [Module K V₁] [AddCommGroup V₂] [Module K V₂] {B : V₁ →ₗ[K] V₂ →ₗ[K] K} [FiniteDimensional K V₁] :
    theorem LinearMap.flip_bijective_iff₂ {K : Type uK} [Field K] {V₁ : Type uV₁} {V₂ : Type uV₂} [AddCommGroup V₁] [Module K V₁] [AddCommGroup V₂] [Module K V₂] {B : V₁ →ₗ[K] V₂ →ₗ[K] K} [FiniteDimensional K V₂] :

    For any vector space, dualAnnihilator and dualCoannihilator gives an antitone order isomorphism between the finite-codimensional subspaces in the vector space and the finite-dimensional subspaces in its dual.

    Equations
    • One or more equations did not get rendered due to their size.

    For any finite-dimensional vector space, dualAnnihilator and dualCoannihilator give an antitone order isomorphism between the subspaces in the vector space and the subspaces in its dual.

    Equations
    • One or more equations did not get rendered due to their size.

    The canonical linear map from Dual M ⊗ Dual N to Dual (M ⊗ N), sending f ⊗ g to the composition of TensorProduct.map f g with the natural isomorphism R ⊗ R ≃ R.

    Equations
    @[simp]
    theorem TensorProduct.dualDistrib_apply {R : Type u_1} {M : Type u_3} {N : Type u_4} [CommSemiring R] [AddCommMonoid M] [AddCommMonoid N] [Module R M] [Module R N] (f : Module.Dual R M) (g : Module.Dual R N) (m : M) (n : N) :
    @[simp]
    theorem TensorProduct.AlgebraTensorModule.dualDistrib_apply {R : Type u_1} (A : Type u_2) {M : Type u_3} {N : Type u_4} [CommSemiring R] [CommSemiring A] [Algebra R A] [AddCommMonoid M] [AddCommMonoid N] [Module R M] [Module A M] [Module R N] [IsScalarTower R A M] (f : Module.Dual A M) (g : Module.Dual R N) (m : M) (n : N) :
    noncomputable def TensorProduct.dualDistribInvOfBasis {R : Type u_1} {M : Type u_3} {N : Type u_4} {ι : Type u_5} {κ : Type u_6} [DecidableEq ι] [DecidableEq κ] [Fintype ι] [Fintype κ] [CommRing R] [AddCommGroup M] [AddCommGroup N] [Module R M] [Module R N] (b : Basis ι R M) (c : Basis κ R N) :

    An inverse to TensorProduct.dualDistrib given bases.

    Equations
    • One or more equations did not get rendered due to their size.
    @[simp]
    theorem TensorProduct.dualDistribInvOfBasis_apply {R : Type u_1} {M : Type u_3} {N : Type u_4} {ι : Type u_5} {κ : Type u_6} [DecidableEq ι] [DecidableEq κ] [Fintype ι] [Fintype κ] [CommRing R] [AddCommGroup M] [AddCommGroup N] [Module R M] [Module R N] (b : Basis ι R M) (c : Basis κ R N) (f : Module.Dual R (TensorProduct R M N)) :
    (dualDistribInvOfBasis b c) f = i : ι, j : κ, f (b i ⊗ₜ[R] c j) b.dualBasis i ⊗ₜ[R] c.dualBasis j
    theorem TensorProduct.dualDistrib_dualDistribInvOfBasis_left_inverse {R : Type u_1} {M : Type u_3} {N : Type u_4} {ι : Type u_5} {κ : Type u_6} [DecidableEq ι] [DecidableEq κ] [Fintype ι] [Fintype κ] [CommRing R] [AddCommGroup M] [AddCommGroup N] [Module R M] [Module R N] (b : Basis ι R M) (c : Basis κ R N) :
    theorem TensorProduct.dualDistrib_dualDistribInvOfBasis_right_inverse {R : Type u_1} {M : Type u_3} {N : Type u_4} {ι : Type u_5} {κ : Type u_6} [DecidableEq ι] [DecidableEq κ] [Fintype ι] [Fintype κ] [CommRing R] [AddCommGroup M] [AddCommGroup N] [Module R M] [Module R N] (b : Basis ι R M) (c : Basis κ R N) :
    noncomputable def TensorProduct.dualDistribEquivOfBasis {R : Type u_1} {M : Type u_3} {N : Type u_4} {ι : Type u_5} {κ : Type u_6} [DecidableEq ι] [DecidableEq κ] [Fintype ι] [Fintype κ] [CommRing R] [AddCommGroup M] [AddCommGroup N] [Module R M] [Module R N] (b : Basis ι R M) (c : Basis κ R N) :

    A linear equivalence between Dual M ⊗ Dual N and Dual (M ⊗ N) given bases for M and N. It sends f ⊗ g to the composition of TensorProduct.map f g with the natural isomorphism R ⊗ R ≃ R.

    Equations
    @[simp]
    theorem TensorProduct.dualDistribEquivOfBasis_symm_apply {R : Type u_1} {M : Type u_3} {N : Type u_4} {ι : Type u_5} {κ : Type u_6} [DecidableEq ι] [DecidableEq κ] [Fintype ι] [Fintype κ] [CommRing R] [AddCommGroup M] [AddCommGroup N] [Module R M] [Module R N] (b : Basis ι R M) (c : Basis κ R N) (a : Module.Dual R (TensorProduct R M N)) :
    (dualDistribEquivOfBasis b c).symm a = x : ι, x_1 : κ, a (b x ⊗ₜ[R] c x_1) b.coord x ⊗ₜ[R] c.coord x_1
    @[simp]
    theorem TensorProduct.dualDistribEquivOfBasis_apply_apply {R : Type u_1} {M : Type u_3} {N : Type u_4} {ι : Type u_5} {κ : Type u_6} [DecidableEq ι] [DecidableEq κ] [Fintype ι] [Fintype κ] [CommRing R] [AddCommGroup M] [AddCommGroup N] [Module R M] [Module R N] (b : Basis ι R M) (c : Basis κ R N) (a✝ : TensorProduct R (Module.Dual R M) (Module.Dual R N)) (a✝¹ : TensorProduct R M N) :
    ((dualDistribEquivOfBasis b c) a✝) a✝¹ = (TensorProduct.lid R R) (((homTensorHomMap R M N R R) a✝) a✝¹)
    noncomputable def TensorProduct.dualDistribEquiv (R : Type u_1) (M : Type u_3) (N : Type u_4) [CommRing R] [AddCommGroup M] [AddCommGroup N] [Module R M] [Module R N] [Module.Finite R M] [Module.Finite R N] [Module.Free R M] [Module.Free R N] :

    A linear equivalence between Dual M ⊗ Dual N and Dual (M ⊗ N) when M and N are finite free modules. It sends f ⊗ g to the composition of TensorProduct.map f g with the natural isomorphism R ⊗ R ≃ R.

    Equations