Documentation

Mathlib.LinearAlgebra.Dual.Defs

Dual vector spaces #

The dual space of an $R$-module $M$ is the $R$-module of $R$-linear maps $M \to R$.

Main definitions #

Main results #

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

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

Equations
    Instances For
      def Module.dualPairing (R : Type u_4) (M : Type u_5) [CommSemiring R] [AddCommMonoid M] [Module R M] :

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

      Equations
        Instances For
          @[simp]
          theorem Module.dualPairing_apply (R : Type u_1) (M : Type u_3) [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 u_1) (M : Type u_3) [CommSemiring R] [AddCommMonoid M] [Module R M] :
          Equations
            def Module.Dual.eval (R : Type u_1) (M : Type u_3) [CommSemiring R] [AddCommMonoid M] [Module R M] :
            M →ₗ[R] Dual R (Dual R M)

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

            Equations
              Instances For
                @[simp]
                theorem Module.Dual.eval_apply (R : Type u_1) (M : Type u_3) [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 u_1} {M : Type u_3} [CommSemiring R] [AddCommMonoid M] [Module R M] {M' : Type u_4} [AddCommMonoid M'] [Module R M'] :
                (M →ₗ[R] M') →ₗ[R] Dual R M' →ₗ[R] Dual R M

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

                Equations
                  Instances For
                    theorem Module.Dual.transpose_apply {R : Type u_1} {M : Type u_3} [CommSemiring R] [AddCommMonoid M] [Module R M] {M' : Type u_4} [AddCommMonoid M'] [Module R M'] (u : M →ₗ[R] M') (l : Dual R M') :
                    (transpose u) l = l ∘ₗ u
                    theorem Module.Dual.transpose_comp {R : Type u_1} {M : Type u_3} [CommSemiring R] [AddCommMonoid M] [Module R M] {M' : Type u_4} [AddCommMonoid M'] [Module R M'] {M'' : Type u_5} [AddCommMonoid M''] [Module R M''] (u : M' →ₗ[R] M'') (v : M →ₗ[R] M') :
                    def LinearMap.dualMap {R : Type u_1} {M₁ : Type u_2} {M₂ : Type u_3} [CommSemiring R] [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
                      Instances For
                        theorem LinearMap.dualMap_eq_lcomp {R : Type u_1} {M₁ : Type u_2} {M₂ : Type u_3} [CommSemiring R] [AddCommMonoid M₁] [Module R M₁] [AddCommMonoid M₂] [Module R M₂] (f : M₁ →ₗ[R] M₂) :
                        f.dualMap = lcomp R R f
                        theorem LinearMap.dualMap_def {R : Type u_1} {M₁ : Type u_2} {M₂ : Type u_3} [CommSemiring R] [AddCommMonoid M₁] [Module R M₁] [AddCommMonoid M₂] [Module R M₂] (f : M₁ →ₗ[R] M₂) :
                        theorem LinearMap.dualMap_apply' {R : Type u_1} {M₁ : Type u_2} {M₂ : Type u_3} [CommSemiring R] [AddCommMonoid M₁] [Module R M₁] [AddCommMonoid M₂] [Module R M₂] (f : M₁ →ₗ[R] M₂) (g : Module.Dual R M₂) :
                        f.dualMap g = g ∘ₗ f
                        @[simp]
                        theorem LinearMap.dualMap_apply {R : Type u_1} {M₁ : Type u_2} {M₂ : Type u_3} [CommSemiring R] [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_1} {M₁ : Type u_2} [CommSemiring R] [AddCommMonoid M₁] [Module R M₁] :
                        theorem LinearMap.dualMap_comp_dualMap {R : Type u_1} {M₁ : Type u_2} {M₂ : Type u_3} [CommSemiring R] [AddCommMonoid M₁] [Module R M₁] [AddCommMonoid M₂] [Module R M₂] {M₃ : Type u_4} [AddCommMonoid M₃] [Module R M₃] (f : M₁ →ₗ[R] M₂) (g : M₂ →ₗ[R] M₃) :
                        theorem LinearMap.dualMap_injective_of_surjective {R : Type u_1} {M₁ : Type u_2} {M₂ : Type u_3} [CommSemiring R] [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_1} {M₁ : Type u_2} {M₂ : Type u_3} [CommSemiring R] [AddCommMonoid M₁] [Module R M₁] [AddCommMonoid M₂] [Module R M₂] (f : M₁ ≃ₗ[R] M₂) :

                        The Linear_equiv version of LinearMap.dualMap.

                        Equations
                          Instances For
                            @[simp]
                            theorem LinearEquiv.dualMap_apply {R : Type u_1} {M₁ : Type u_2} {M₂ : Type u_3} [CommSemiring R] [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_1} {M₁ : Type u_2} [CommSemiring R] [AddCommMonoid M₁] [Module R M₁] :
                            (refl R M₁).dualMap = refl R (Module.Dual R M₁)
                            @[simp]
                            theorem LinearEquiv.dualMap_symm {R : Type u_1} {M₁ : Type u_2} {M₂ : Type u_3} [CommSemiring R] [AddCommMonoid M₁] [Module R M₁] [AddCommMonoid M₂] [Module R M₂] {f : M₁ ≃ₗ[R] M₂} :
                            theorem LinearEquiv.dualMap_trans {R : Type u_1} {M₁ : Type u_2} {M₂ : Type u_3} [CommSemiring R] [AddCommMonoid M₁] [Module R M₁] [AddCommMonoid M₂] [Module R M₂] {M₃ : Type u_4} [AddCommMonoid M₃] [Module R M₃] (f : M₁ ≃ₗ[R] M₂) (g : M₂ ≃ₗ[R] M₃) :
                            theorem Module.Dual.eval_naturality {R : Type u_1} {M₁ : Type u_2} {M₂ : Type u_3} [CommSemiring R] [AddCommMonoid M₁] [Module R M₁] [AddCommMonoid M₂] [Module R M₂] (f : M₁ →ₗ[R] M₂) :
                            f.dualMap.dualMap ∘ₗ eval R M₁ = eval R M₂ ∘ₗ f
                            @[simp]
                            theorem Dual.apply_one_mul_eq {R : Type u_1} [CommSemiring R] (f : Module.Dual R R) (r : R) :
                            f 1 * r = f r
                            @[simp]
                            class Module.IsReflexive (R : Type u_3) (M : Type u_4) [CommSemiring R] [AddCommMonoid 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
                              def Module.evalEquiv (R : Type u_3) (M : Type u_4) [CommSemiring R] [AddCommMonoid M] [Module R M] [IsReflexive R M] :
                              M ≃ₗ[R] Dual R (Dual R M)

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

                              Equations
                                Instances For
                                  @[simp]
                                  theorem Module.evalEquiv_toLinearMap (R : Type u_3) (M : Type u_4) [CommSemiring R] [AddCommMonoid M] [Module R M] [IsReflexive R M] :
                                  (evalEquiv R M) = Dual.eval R M
                                  @[simp]
                                  theorem Module.evalEquiv_apply (R : Type u_3) (M : Type u_4) [CommSemiring R] [AddCommMonoid 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_3) (M : Type u_4) [CommSemiring R] [AddCommMonoid 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]
                                  theorem Module.symm_dualMap_evalEquiv (R : Type u_3) (M : Type u_4) [CommSemiring R] [AddCommMonoid M] [Module R M] [IsReflexive R M] :
                                  @[simp]
                                  theorem Module.Dual.eval_comp_comp_evalEquiv_eq (R : Type u_3) (M : Type u_4) [CommSemiring R] [AddCommMonoid M] [Module R M] [IsReflexive R M] {M' : Type u_6} [AddCommMonoid M'] [Module R M'] {f : M →ₗ[R] M'} :
                                  theorem Module.dualMap_dualMap_eq_iff_of_injective (R : Type u_3) (M : Type u_4) [CommSemiring R] [AddCommMonoid M] [Module R M] [IsReflexive R M] {M' : Type u_6} [AddCommMonoid M'] [Module R M'] {f g : M →ₗ[R] M'} (h : Function.Injective (Dual.eval R M')) :
                                  @[simp]
                                  theorem Module.dualMap_dualMap_eq_iff (R : Type u_3) (M : Type u_4) [CommSemiring R] [AddCommMonoid M] [Module R M] [IsReflexive R M] {M' : Type u_6} [AddCommMonoid M'] [Module R M'] [IsReflexive R M'] {f g : M →ₗ[R] M'} :
                                  instance Module.Dual.instIsReflecive (R : Type u_3) (M : Type u_4) [CommSemiring R] [AddCommMonoid M] [Module R M] [IsReflexive R M] :

                                  The dual of a reflexive module is reflexive.

                                  theorem Module.IsReflexive.of_split {R : Type u_3} {M : Type u_4} {N : Type u_5} [CommSemiring R] [AddCommMonoid M] [AddCommMonoid 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.

                                  def Module.mapEvalEquiv (R : Type u_3) (M : Type u_4) [CommSemiring R] [AddCommMonoid 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
                                    Instances For
                                      @[simp]
                                      theorem Module.mapEvalEquiv_apply (R : Type u_3) (M : Type u_4) [CommSemiring R] [AddCommMonoid M] [Module R M] [IsReflexive R M] (W : Submodule R M) :
                                      @[simp]
                                      theorem Module.mapEvalEquiv_symm_apply (R : Type u_3) (M : Type u_4) [CommSemiring R] [AddCommMonoid M] [Module R M] [IsReflexive R M] (W'' : Submodule R (Dual R (Dual R M))) :
                                      theorem Module.equiv {R : Type u_3} {M : Type u_4} {N : Type u_5} [CommSemiring R] [AddCommMonoid M] [AddCommMonoid N] [Module R M] [Module R N] [IsReflexive R M] (e : M ≃ₗ[R] N) :
                                      @[instance 100]
                                      def Submodule.dualRestrict {R : Type u_1} {M : Type u_2} [CommSemiring R] [AddCommMonoid M] [Module R M] (W : Submodule R M) :

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

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

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

                                              That $\operatorname{ker}(\iota^* : V^* \to W^*) = \operatorname{ann}(W)$. This is the definition of the dual annihilator of the submodule $W$.

                                              def Submodule.dualCoannihilator {R : Type u_1} {M : Type u_2} [CommSemiring R] [AddCommMonoid M] [Module R M] (Φ : Submodule R (Module.Dual R M)) :

                                              The dualAnnihilator of a submodule of the dual space pulled back along the evaluation map Module.Dual.eval.

                                              Equations
                                                Instances For
                                                  @[simp]
                                                  theorem Submodule.mem_dualCoannihilator {R : Type u_1} {M : Type u_2} [CommSemiring R] [AddCommMonoid M] [Module R M] {Φ : Submodule R (Module.Dual R M)} (x : M) :
                                                  x Φ.dualCoannihilator φΦ, φ x = 0
                                                  theorem Submodule.dualAnnihilator_anti {R : Type u_1} {M : Type u_2} [CommSemiring R] [AddCommMonoid M] [Module R M] {U V : Submodule R M} (hUV : U V) :
                                                  theorem Submodule.dualAnnihilator_iSup_eq {R : Type u_1} {M : Type u_2} [CommSemiring R] [AddCommMonoid M] [Module R M] {ι : Sort u_3} (U : ιSubmodule R M) :
                                                  (⨆ (i : ι), U i).dualAnnihilator = ⨅ (i : ι), (U i).dualAnnihilator
                                                  theorem Submodule.dualCoannihilator_iSup_eq {R : Type u_1} {M : Type u_2} [CommSemiring R] [AddCommMonoid M] [Module R M] {ι : Sort u_3} (U : ιSubmodule R (Module.Dual R M)) :
                                                  (⨆ (i : ι), U i).dualCoannihilator = ⨅ (i : ι), (U i).dualCoannihilator

                                                  See also Subspace.dualAnnihilator_inf_eq for vector subspaces.

                                                  theorem Submodule.iSup_dualAnnihilator_le_iInf {R : Type u_1} {M : Type u_2} [CommSemiring R] [AddCommMonoid M] [Module R M] {ι : Sort u_3} (U : ιSubmodule R M) :
                                                  ⨆ (i : ι), (U i).dualAnnihilator (⨅ (i : ι), U i).dualAnnihilator

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

                                                  @[simp]
                                                  theorem Submodule.coe_dualAnnihilator_span {R : Type u_1} {M : Type u_2} [CommSemiring R] [AddCommMonoid M] [Module R M] (s : Set M) :
                                                  @[simp]
                                                  theorem Submodule.coe_dualCoannihilator_span {R : Type u_1} {M : Type u_2} [CommSemiring R] [AddCommMonoid M] [Module R M] (s : Set (Module.Dual R M)) :
                                                  (span R s).dualCoannihilator = {x : M | fs, f x = 0}
                                                  theorem LinearMap.ker_dualMap_eq_dualAnnihilator_range {R : Type u_1} {M₁ : Type u_2} {M₂ : Type u_3} [CommSemiring R] [AddCommMonoid M₁] [Module R M₁] [AddCommMonoid M₂] [Module R M₂] (f : M₁ →ₗ[R] M₂) :
                                                  theorem LinearMap.range_dualMap_le_dualAnnihilator_ker {R : Type u_1} {M₁ : Type u_2} {M₂ : Type u_3} [CommSemiring R] [AddCommMonoid M₁] [Module R M₁] [AddCommMonoid M₂] [Module R M₂] (f : M₁ →ₗ[R] M₂) :
                                                  @[simp]
                                                  theorem LinearMap.dualCoannihilator_range_eq_ker_flip {R : Type u_1} {M : Type u_2} {M' : Type u_3} [CommSemiring R] [AddCommMonoid M] [Module R M] [AddCommMonoid M'] [Module R M'] (B : M →ₗ[R] M' →ₗ[R] R) :