Documentation

Mathlib.LinearAlgebra.Dual.Lemmas

Dual vector spaces #

The dual space of an $R$-module $M$ is the $R$-module of $R$-linear maps $M \to R$. This file contains basic results on dual vector spaces.

Main definitions #

Main results #

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
Instances For
    @[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') :

    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.

    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] :
    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) :
    (∀ (φ : Dual K V), φ v = 0) v = 0
    @[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] :
    @[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.

    @[instance 900]
    instance Prod.instModuleIsReflexive (R : Type u_1) (M : Type u_2) (N : Type u_3) [CommRing R] [AddCommGroup M] [AddCommGroup N] [Module R M] [Module R N] [Module.IsReflexive R M] [Module.IsReflexive R N] :
    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 $L_1, ..., L_n, K$ over a vector space $E$, if $\bigcap_{i=1}^n \mathrm{ker}(L_i) \subseteq \mathrm{ker}(K)$, then $K$ is in the space generated by $L_1, ..., L_n$.

    theorem Subspace.forall_mem_dualAnnihilator_apply_eq_zero_iff {K : Type u} {V : Type v} [Field K] [AddCommGroup V] [Module K V] (W : Subspace K V) (v : V) :
    (∀ φSubmodule.dualAnnihilator W, φ v = 0) v W

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

    Equations
    • One or more equations did not get rendered due to their size.
    Instances For
      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
      Instances For
        @[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
        Instances For
          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
          Instances For
            @[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]

            The quotient by the dual is isomorphic to its dual annihilator.

            Equations
            Instances For
              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.
              Instances For
                def Submodule.dualCopairing {R : Type u_1} {M : Type u_2} [CommRing R] [AddCommGroup M] [Module R M] (W : Submodule R M) :

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

                See Subspace.dualCopairing_nondegenerate.

                Equations
                Instances For
                  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) :
                  (W.dualCopairing φ) (Quotient.mk x) = φ x
                  def Submodule.dualPairing {R : Type u_1} {M : Type u_2} [CommRing R] [AddCommGroup M] [Module R M] (W : Submodule R 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
                  Instances For
                    @[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 $\operatorname{im}(q^* : (V/W)^* \to V^*) = \operatorname{ann}(W)$.

                    Equivalence $(M/W)^* \cong \operatorname{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
                    Instances For
                      @[simp]
                      theorem Submodule.dualQuotEquivDualAnnihilator_apply {R : Type u_1} {M : Type u_2} [CommRing R] [AddCommGroup M] [Module R M] (W : Submodule R M) (φ : Module.Dual R (M W)) (x : M) :
                      @[simp]
                      @[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
                      Instances For
                        @[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
                        Instances For
                          @[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.
                          Instances For

                            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.
                            Instances For

                              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
                              Instances For
                                @[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) :
                                ((dualDistrib R M N) (f ⊗ₜ[R] g)) (m ⊗ₜ[R] n) = f m * g 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) :
                                ((dualDistrib R A M N) (f ⊗ₜ[R] g)) (m ⊗ₜ[R] n) = g n f m
                                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.
                                Instances For
                                  @[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
                                  Instances For
                                    @[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
                                    Instances For