Documentation

Mathlib.LinearAlgebra.Basic

Linear algebra #

This file defines the basics of linear algebra. It sets up the "categorical/lattice structure" of modules over a ring, submodules, and linear maps.

Many of the relevant definitions, including Module, Submodule, and LinearMap, are found in Algebra/Module.

Main definitions #

See LinearAlgebra.Span for the span of a set (as a submodule), and LinearAlgebra.Quotient for quotients by submodules.

Main theorems #

See LinearAlgebra.Isomorphisms for Noether's three isomorphism theorems for modules.

Notations #

Implementation notes #

We note that, when constructing linear maps, it is convenient to use operations defined on bundled maps (LinearMap.prod, LinearMap.coprod, arithmetic operations like +) instead of defining a function and proving it is linear.

TODO #

Tags #

linear algebra, vector space, module

Properties of linear maps #

@[simp]
theorem addMonoidHomLequivNat_apply {A : Type u_20} {B : Type u_21} (R : Type u_22) [Semiring R] [AddCommMonoid A] [AddCommMonoid B] [Module R B] (f : A →+ B) :
(addMonoidHomLequivNat R) f = f.toNatLinearMap
@[simp]
theorem addMonoidHomLequivNat_symm_apply {A : Type u_20} {B : Type u_21} (R : Type u_22) [Semiring R] [AddCommMonoid A] [AddCommMonoid B] [Module R B] (f : A →ₗ[] B) :
(addMonoidHomLequivNat R).symm f = f.toAddMonoidHom
def addMonoidHomLequivNat {A : Type u_20} {B : Type u_21} (R : Type u_22) [Semiring R] [AddCommMonoid A] [AddCommMonoid B] [Module R B] :

The R-linear equivalence between additive morphisms A →+ B and -linear morphisms A →ₗ[ℕ] B.

Equations
  • addMonoidHomLequivNat R = { toFun := AddMonoidHom.toNatLinearMap, map_add' := , map_smul' := , invFun := LinearMap.toAddMonoidHom, left_inv := , right_inv := }
Instances For
    @[simp]
    theorem addMonoidHomLequivInt_apply {A : Type u_20} {B : Type u_21} (R : Type u_22) [Semiring R] [AddCommGroup A] [AddCommGroup B] [Module R B] (f : A →+ B) :
    (addMonoidHomLequivInt R) f = f.toIntLinearMap
    @[simp]
    theorem addMonoidHomLequivInt_symm_apply {A : Type u_20} {B : Type u_21} (R : Type u_22) [Semiring R] [AddCommGroup A] [AddCommGroup B] [Module R B] (f : A →ₗ[] B) :
    (addMonoidHomLequivInt R).symm f = f.toAddMonoidHom
    def addMonoidHomLequivInt {A : Type u_20} {B : Type u_21} (R : Type u_22) [Semiring R] [AddCommGroup A] [AddCommGroup B] [Module R B] :

    The R-linear equivalence between additive morphisms A →+ B and -linear morphisms A →ₗ[ℤ] B.

    Equations
    • addMonoidHomLequivInt R = { toFun := AddMonoidHom.toIntLinearMap, map_add' := , map_smul' := , invFun := LinearMap.toAddMonoidHom, left_inv := , right_inv := }
    Instances For
      @[simp]
      theorem addMonoidEndRingEquivInt_apply (A : Type u_20) [AddCommGroup A] :
      ∀ (a : A →+ A), (addMonoidEndRingEquivInt A) a = ((addMonoidHomLequivInt )).toFun a

      Ring equivalence between additive group endomorphisms of an AddCommGroup A and -module endomorphisms of A.

      Equations
      Instances For

        Properties of linear maps #

        def LinearMap.eqLocus {R : Type u_1} {R₂ : Type u_3} {M : Type u_9} {M₂ : Type u_12} [Semiring R] [Semiring R₂] [AddCommMonoid M] [AddCommMonoid M₂] [Module R M] [Module R₂ M₂] {τ₁₂ : R →+* R₂} {F : Type u_20} [FunLike F M M₂] [SemilinearMapClass F τ₁₂ M M₂] (f : F) (g : F) :

        A linear map version of AddMonoidHom.eqLocusM

        Equations
        • LinearMap.eqLocus f g = let __src := (f).eqLocusM g; { carrier := {x : M | f x = g x}, add_mem' := , zero_mem' := , smul_mem' := }
        Instances For
          @[simp]
          theorem LinearMap.mem_eqLocus {R : Type u_1} {R₂ : Type u_3} {M : Type u_9} {M₂ : Type u_12} [Semiring R] [Semiring R₂] [AddCommMonoid M] [AddCommMonoid M₂] [Module R M] [Module R₂ M₂] {τ₁₂ : R →+* R₂} {F : Type u_20} [FunLike F M M₂] [SemilinearMapClass F τ₁₂ M M₂] {x : M} {f : F} {g : F} :
          x LinearMap.eqLocus f g f x = g x
          theorem LinearMap.eqLocus_toAddSubmonoid {R : Type u_1} {R₂ : Type u_3} {M : Type u_9} {M₂ : Type u_12} [Semiring R] [Semiring R₂] [AddCommMonoid M] [AddCommMonoid M₂] [Module R M] [Module R₂ M₂] {τ₁₂ : R →+* R₂} {F : Type u_20} [FunLike F M M₂] [SemilinearMapClass F τ₁₂ M M₂] (f : F) (g : F) :
          (LinearMap.eqLocus f g).toAddSubmonoid = (f).eqLocusM g
          @[simp]
          theorem LinearMap.eqLocus_eq_top {R : Type u_1} {R₂ : Type u_3} {M : Type u_9} {M₂ : Type u_12} [Semiring R] [Semiring R₂] [AddCommMonoid M] [AddCommMonoid M₂] [Module R M] [Module R₂ M₂] {τ₁₂ : R →+* R₂} {F : Type u_20} [FunLike F M M₂] [SemilinearMapClass F τ₁₂ M M₂] {f : F} {g : F} :
          @[simp]
          theorem LinearMap.eqLocus_same {R : Type u_1} {R₂ : Type u_3} {M : Type u_9} {M₂ : Type u_12} [Semiring R] [Semiring R₂] [AddCommMonoid M] [AddCommMonoid M₂] [Module R M] [Module R₂ M₂] {τ₁₂ : R →+* R₂} {F : Type u_20} [FunLike F M M₂] [SemilinearMapClass F τ₁₂ M M₂] (f : F) :
          theorem LinearMap.le_eqLocus {R : Type u_1} {R₂ : Type u_3} {M : Type u_9} {M₂ : Type u_12} [Semiring R] [Semiring R₂] [AddCommMonoid M] [AddCommMonoid M₂] [Module R M] [Module R₂ M₂] {τ₁₂ : R →+* R₂} {F : Type u_20} [FunLike F M M₂] [SemilinearMapClass F τ₁₂ M M₂] {f : F} {g : F} {S : Submodule R M} :
          S LinearMap.eqLocus f g Set.EqOn f g S
          theorem LinearMap.eqOn_sup {R : Type u_1} {R₂ : Type u_3} {M : Type u_9} {M₂ : Type u_12} [Semiring R] [Semiring R₂] [AddCommMonoid M] [AddCommMonoid M₂] [Module R M] [Module R₂ M₂] {τ₁₂ : R →+* R₂} {F : Type u_20} [FunLike F M M₂] [SemilinearMapClass F τ₁₂ M M₂] {f : F} {g : F} {S : Submodule R M} {T : Submodule R M} (hS : Set.EqOn f g S) (hT : Set.EqOn f g T) :
          Set.EqOn f g (S T)
          theorem LinearMap.ext_on_codisjoint {R : Type u_1} {R₂ : Type u_3} {M : Type u_9} {M₂ : Type u_12} [Semiring R] [Semiring R₂] [AddCommMonoid M] [AddCommMonoid M₂] [Module R M] [Module R₂ M₂] {τ₁₂ : R →+* R₂} {F : Type u_20} [FunLike F M M₂] [SemilinearMapClass F τ₁₂ M M₂] {f : F} {g : F} {S : Submodule R M} {T : Submodule R M} (hST : Codisjoint S T) (hS : Set.EqOn f g S) (hT : Set.EqOn f g T) :
          f = g
          theorem LinearMap.eqLocus_eq_ker_sub {R : Type u_1} {R₂ : Type u_3} {M : Type u_9} {M₂ : Type u_12} [Ring R] [Ring R₂] [AddCommGroup M] [AddCommGroup M₂] [Module R M] [Module R₂ M₂] {τ₁₂ : R →+* R₂} (f : M →ₛₗ[τ₁₂] M₂) (g : M →ₛₗ[τ₁₂] M₂) :
          theorem IsLinearMap.isLinearMap_add {R : Type u_1} {M : Type u_9} [Semiring R] [AddCommMonoid M] [Module R M] :
          IsLinearMap R fun (x : M × M) => x.1 + x.2
          theorem IsLinearMap.isLinearMap_sub {R : Type u_20} {M : Type u_21} [Semiring R] [AddCommGroup M] [Module R M] :
          IsLinearMap R fun (x : M × M) => x.1 - x.2

          Linear equivalences #

          instance LinearEquiv.instZero {R : Type u_1} {R₂ : Type u_3} {M : Type u_9} {M₂ : Type u_12} [Semiring R] [Semiring R₂] [AddCommMonoid M] [AddCommMonoid M₂] [Module R M] [Module R₂ M₂] {σ₁₂ : R →+* R₂} {σ₂₁ : R₂ →+* R} [RingHomInvPair σ₁₂ σ₂₁] [RingHomInvPair σ₂₁ σ₁₂] [Subsingleton M] [Subsingleton M₂] :
          Zero (M ≃ₛₗ[σ₁₂] M₂)

          Between two zero modules, the zero map is an equivalence.

          Equations
          • LinearEquiv.instZero = { zero := let __src := 0; { toFun := 0, map_add' := , map_smul' := , invFun := 0, left_inv := , right_inv := } }
          @[simp]
          theorem LinearEquiv.zero_symm {R : Type u_1} {R₂ : Type u_3} {M : Type u_9} {M₂ : Type u_12} [Semiring R] [Semiring R₂] [AddCommMonoid M] [AddCommMonoid M₂] [Module R M] [Module R₂ M₂] {σ₁₂ : R →+* R₂} {σ₂₁ : R₂ →+* R} [RingHomInvPair σ₁₂ σ₂₁] [RingHomInvPair σ₂₁ σ₁₂] [Subsingleton M] [Subsingleton M₂] :
          @[simp]
          theorem LinearEquiv.coe_zero {R : Type u_1} {R₂ : Type u_3} {M : Type u_9} {M₂ : Type u_12} [Semiring R] [Semiring R₂] [AddCommMonoid M] [AddCommMonoid M₂] [Module R M] [Module R₂ M₂] {σ₁₂ : R →+* R₂} {σ₂₁ : R₂ →+* R} [RingHomInvPair σ₁₂ σ₂₁] [RingHomInvPair σ₂₁ σ₁₂] [Subsingleton M] [Subsingleton M₂] :
          0 = 0
          theorem LinearEquiv.zero_apply {R : Type u_1} {R₂ : Type u_3} {M : Type u_9} {M₂ : Type u_12} [Semiring R] [Semiring R₂] [AddCommMonoid M] [AddCommMonoid M₂] [Module R M] [Module R₂ M₂] {σ₁₂ : R →+* R₂} {σ₂₁ : R₂ →+* R} [RingHomInvPair σ₁₂ σ₂₁] [RingHomInvPair σ₂₁ σ₁₂] [Subsingleton M] [Subsingleton M₂] (x : M) :
          0 x = 0
          instance LinearEquiv.instUnique {R : Type u_1} {R₂ : Type u_3} {M : Type u_9} {M₂ : Type u_12} [Semiring R] [Semiring R₂] [AddCommMonoid M] [AddCommMonoid M₂] [Module R M] [Module R₂ M₂] {σ₁₂ : R →+* R₂} {σ₂₁ : R₂ →+* R} [RingHomInvPair σ₁₂ σ₂₁] [RingHomInvPair σ₂₁ σ₁₂] [Subsingleton M] [Subsingleton M₂] :
          Unique (M ≃ₛₗ[σ₁₂] M₂)

          Between two zero modules, the zero map is the only equivalence.

          Equations
          • LinearEquiv.instUnique = { default := 0, uniq := }
          instance LinearEquiv.uniqueOfSubsingleton {R : Type u_1} {R₂ : Type u_3} {M : Type u_9} {M₂ : Type u_12} [Semiring R] [Semiring R₂] [AddCommMonoid M] [AddCommMonoid M₂] [Module R M] [Module R₂ M₂] {σ₁₂ : R →+* R₂} {σ₂₁ : R₂ →+* R} [RingHomInvPair σ₁₂ σ₂₁] [RingHomInvPair σ₂₁ σ₁₂] [Subsingleton R] [Subsingleton R₂] :
          Unique (M ≃ₛₗ[σ₁₂] M₂)
          Equations
          • LinearEquiv.uniqueOfSubsingleton = inferInstance
          def LinearEquiv.curry (R : Type u_1) (V : Type u_18) (V₂ : Type u_19) [Semiring R] :
          (V × V₂R) ≃ₗ[R] VV₂R

          Linear equivalence between a curried and uncurried function. Differs from TensorProduct.curry.

          Equations
          • LinearEquiv.curry R V V₂ = let __src := Equiv.curry V V₂ R; { toFun := __src.toFun, map_add' := , map_smul' := , invFun := __src.invFun, left_inv := , right_inv := }
          Instances For
            @[simp]
            theorem LinearEquiv.coe_curry (R : Type u_1) (V : Type u_18) (V₂ : Type u_19) [Semiring R] :
            (LinearEquiv.curry R V V₂) = Function.curry
            @[simp]
            theorem LinearEquiv.coe_curry_symm (R : Type u_1) (V : Type u_18) (V₂ : Type u_19) [Semiring R] :
            (LinearEquiv.curry R V V₂).symm = Function.uncurry
            def LinearEquiv.ofEq {R : Type u_1} {M : Type u_9} [Semiring R] [AddCommMonoid M] {module_M : Module R M} (p : Submodule R M) (q : Submodule R M) (h : p = q) :
            p ≃ₗ[R] q

            Linear equivalence between two equal submodules.

            Equations
            • LinearEquiv.ofEq p q h = let __src := Equiv.Set.ofEq ; { toFun := __src.toFun, map_add' := , map_smul' := , invFun := __src.invFun, left_inv := , right_inv := }
            Instances For
              @[simp]
              theorem LinearEquiv.coe_ofEq_apply {R : Type u_1} {M : Type u_9} [Semiring R] [AddCommMonoid M] {module_M : Module R M} {p : Submodule R M} {q : Submodule R M} (h : p = q) (x : p) :
              ((LinearEquiv.ofEq p q h) x) = x
              @[simp]
              theorem LinearEquiv.ofEq_symm {R : Type u_1} {M : Type u_9} [Semiring R] [AddCommMonoid M] {module_M : Module R M} {p : Submodule R M} {q : Submodule R M} (h : p = q) :
              (LinearEquiv.ofEq p q h).symm = LinearEquiv.ofEq q p
              @[simp]
              theorem LinearEquiv.ofEq_rfl {R : Type u_1} {M : Type u_9} [Semiring R] [AddCommMonoid M] {module_M : Module R M} {p : Submodule R M} :
              def LinearEquiv.ofSubmodules {R : Type u_1} {R₂ : Type u_3} {M : Type u_9} {M₂ : Type u_12} [Semiring R] [Semiring R₂] [AddCommMonoid M] [AddCommMonoid M₂] {module_M : Module R M} {module_M₂ : Module R₂ M₂} {σ₁₂ : R →+* R₂} {σ₂₁ : R₂ →+* R} {re₁₂ : RingHomInvPair σ₁₂ σ₂₁} {re₂₁ : RingHomInvPair σ₂₁ σ₁₂} (e : M ≃ₛₗ[σ₁₂] M₂) (p : Submodule R M) (q : Submodule R₂ M₂) (h : Submodule.map (e) p = q) :
              p ≃ₛₗ[σ₁₂] q

              A linear equivalence which maps a submodule of one module onto another, restricts to a linear equivalence of the two submodules.

              Equations
              Instances For
                @[simp]
                theorem LinearEquiv.ofSubmodules_apply {R : Type u_1} {R₂ : Type u_3} {M : Type u_9} {M₂ : Type u_12} [Semiring R] [Semiring R₂] [AddCommMonoid M] [AddCommMonoid M₂] {module_M : Module R M} {module_M₂ : Module R₂ M₂} {σ₁₂ : R →+* R₂} {σ₂₁ : R₂ →+* R} {re₁₂ : RingHomInvPair σ₁₂ σ₂₁} {re₂₁ : RingHomInvPair σ₂₁ σ₁₂} (e : M ≃ₛₗ[σ₁₂] M₂) {p : Submodule R M} {q : Submodule R₂ M₂} (h : Submodule.map e p = q) (x : p) :
                ((e.ofSubmodules p q h) x) = e x
                @[simp]
                theorem LinearEquiv.ofSubmodules_symm_apply {R : Type u_1} {R₂ : Type u_3} {M : Type u_9} {M₂ : Type u_12} [Semiring R] [Semiring R₂] [AddCommMonoid M] [AddCommMonoid M₂] {module_M : Module R M} {module_M₂ : Module R₂ M₂} {σ₁₂ : R →+* R₂} {σ₂₁ : R₂ →+* R} {re₁₂ : RingHomInvPair σ₁₂ σ₂₁} {re₂₁ : RingHomInvPair σ₂₁ σ₁₂} (e : M ≃ₛₗ[σ₁₂] M₂) {p : Submodule R M} {q : Submodule R₂ M₂} (h : Submodule.map e p = q) (x : q) :
                ((e.ofSubmodules p q h).symm x) = e.symm x
                def LinearEquiv.ofSubmodule' {R : Type u_1} {R₂ : Type u_3} {M : Type u_9} {M₂ : Type u_12} [Semiring R] [Semiring R₂] [AddCommMonoid M] [AddCommMonoid M₂] {σ₁₂ : R →+* R₂} {σ₂₁ : R₂ →+* R} {re₁₂ : RingHomInvPair σ₁₂ σ₂₁} {re₂₁ : RingHomInvPair σ₂₁ σ₁₂} [Module R M] [Module R₂ M₂] (f : M ≃ₛₗ[σ₁₂] M₂) (U : Submodule R₂ M₂) :
                (Submodule.comap (f) U) ≃ₛₗ[σ₁₂] U

                A linear equivalence of two modules restricts to a linear equivalence from the preimage of any submodule to that submodule.

                This is LinearEquiv.ofSubmodule but with comap on the left instead of map on the right.

                Equations
                • f.ofSubmodule' U = (f.symm.ofSubmodules U (Submodule.comap (f.symm.symm) U) ).symm
                Instances For
                  theorem LinearEquiv.ofSubmodule'_toLinearMap {R : Type u_1} {R₂ : Type u_3} {M : Type u_9} {M₂ : Type u_12} [Semiring R] [Semiring R₂] [AddCommMonoid M] [AddCommMonoid M₂] {σ₁₂ : R →+* R₂} {σ₂₁ : R₂ →+* R} {re₁₂ : RingHomInvPair σ₁₂ σ₂₁} {re₂₁ : RingHomInvPair σ₂₁ σ₁₂} [Module R M] [Module R₂ M₂] (f : M ≃ₛₗ[σ₁₂] M₂) (U : Submodule R₂ M₂) :
                  (f.ofSubmodule' U) = LinearMap.codRestrict U ((f).domRestrict (Submodule.comap (f) U))
                  @[simp]
                  theorem LinearEquiv.ofSubmodule'_apply {R : Type u_1} {R₂ : Type u_3} {M : Type u_9} {M₂ : Type u_12} [Semiring R] [Semiring R₂] [AddCommMonoid M] [AddCommMonoid M₂] {σ₁₂ : R →+* R₂} {σ₂₁ : R₂ →+* R} {re₁₂ : RingHomInvPair σ₁₂ σ₂₁} {re₂₁ : RingHomInvPair σ₂₁ σ₁₂} [Module R M] [Module R₂ M₂] (f : M ≃ₛₗ[σ₁₂] M₂) (U : Submodule R₂ M₂) (x : (Submodule.comap (f) U)) :
                  ((f.ofSubmodule' U) x) = f x
                  @[simp]
                  theorem LinearEquiv.ofSubmodule'_symm_apply {R : Type u_1} {R₂ : Type u_3} {M : Type u_9} {M₂ : Type u_12} [Semiring R] [Semiring R₂] [AddCommMonoid M] [AddCommMonoid M₂] {σ₁₂ : R →+* R₂} {σ₂₁ : R₂ →+* R} {re₁₂ : RingHomInvPair σ₁₂ σ₂₁} {re₂₁ : RingHomInvPair σ₂₁ σ₁₂} [Module R M] [Module R₂ M₂] (f : M ≃ₛₗ[σ₁₂] M₂) (U : Submodule R₂ M₂) (x : U) :
                  ((f.ofSubmodule' U).symm x) = f.symm x
                  def LinearEquiv.ofTop {R : Type u_1} {M : Type u_9} [Semiring R] [AddCommMonoid M] {module_M : Module R M} (p : Submodule R M) (h : p = ) :
                  p ≃ₗ[R] M

                  The top submodule of M is linearly equivalent to M.

                  Equations
                  • LinearEquiv.ofTop p h = let __src := p.subtype; { toLinearMap := __src, invFun := fun (x : M) => x, , left_inv := , right_inv := }
                  Instances For
                    @[simp]
                    theorem LinearEquiv.ofTop_apply {R : Type u_1} {M : Type u_9} [Semiring R] [AddCommMonoid M] {module_M : Module R M} (p : Submodule R M) {h : p = } (x : p) :
                    (LinearEquiv.ofTop p h) x = x
                    @[simp]
                    theorem LinearEquiv.coe_ofTop_symm_apply {R : Type u_1} {M : Type u_9} [Semiring R] [AddCommMonoid M] {module_M : Module R M} (p : Submodule R M) {h : p = } (x : M) :
                    ((LinearEquiv.ofTop p h).symm x) = x
                    theorem LinearEquiv.ofTop_symm_apply {R : Type u_1} {M : Type u_9} [Semiring R] [AddCommMonoid M] {module_M : Module R M} (p : Submodule R M) {h : p = } (x : M) :
                    (LinearEquiv.ofTop p h).symm x = x,
                    def LinearEquiv.ofLinear {R : Type u_1} {R₂ : Type u_3} {M : Type u_9} {M₂ : Type u_12} [Semiring R] [Semiring R₂] [AddCommMonoid M] [AddCommMonoid M₂] {module_M : Module R M} {module_M₂ : Module R₂ M₂} {σ₁₂ : R →+* R₂} {σ₂₁ : R₂ →+* R} {re₁₂ : RingHomInvPair σ₁₂ σ₂₁} {re₂₁ : RingHomInvPair σ₂₁ σ₁₂} (f : M →ₛₗ[σ₁₂] M₂) (g : M₂ →ₛₗ[σ₂₁] M) (h₁ : f.comp g = LinearMap.id) (h₂ : g.comp f = LinearMap.id) :
                    M ≃ₛₗ[σ₁₂] M₂

                    If a linear map has an inverse, it is a linear equivalence.

                    Equations
                    • LinearEquiv.ofLinear f g h₁ h₂ = { toLinearMap := f, invFun := g, left_inv := , right_inv := }
                    Instances For
                      @[simp]
                      theorem LinearEquiv.ofLinear_apply {R : Type u_1} {R₂ : Type u_3} {M : Type u_9} {M₂ : Type u_12} [Semiring R] [Semiring R₂] [AddCommMonoid M] [AddCommMonoid M₂] {module_M : Module R M} {module_M₂ : Module R₂ M₂} {σ₁₂ : R →+* R₂} {σ₂₁ : R₂ →+* R} {re₁₂ : RingHomInvPair σ₁₂ σ₂₁} {re₂₁ : RingHomInvPair σ₂₁ σ₁₂} (f : M →ₛₗ[σ₁₂] M₂) (g : M₂ →ₛₗ[σ₂₁] M) {h₁ : f.comp g = LinearMap.id} {h₂ : g.comp f = LinearMap.id} (x : M) :
                      (LinearEquiv.ofLinear f g h₁ h₂) x = f x
                      @[simp]
                      theorem LinearEquiv.ofLinear_symm_apply {R : Type u_1} {R₂ : Type u_3} {M : Type u_9} {M₂ : Type u_12} [Semiring R] [Semiring R₂] [AddCommMonoid M] [AddCommMonoid M₂] {module_M : Module R M} {module_M₂ : Module R₂ M₂} {σ₁₂ : R →+* R₂} {σ₂₁ : R₂ →+* R} {re₁₂ : RingHomInvPair σ₁₂ σ₂₁} {re₂₁ : RingHomInvPair σ₂₁ σ₁₂} (f : M →ₛₗ[σ₁₂] M₂) (g : M₂ →ₛₗ[σ₂₁] M) {h₁ : f.comp g = LinearMap.id} {h₂ : g.comp f = LinearMap.id} (x : M₂) :
                      (LinearEquiv.ofLinear f g h₁ h₂).symm x = g x
                      @[simp]
                      theorem LinearEquiv.ofLinear_toLinearMap {R : Type u_1} {R₂ : Type u_3} {M : Type u_9} {M₂ : Type u_12} [Semiring R] [Semiring R₂] [AddCommMonoid M] [AddCommMonoid M₂] {module_M : Module R M} {module_M₂ : Module R₂ M₂} {σ₁₂ : R →+* R₂} {σ₂₁ : R₂ →+* R} {re₁₂ : RingHomInvPair σ₁₂ σ₂₁} {re₂₁ : RingHomInvPair σ₂₁ σ₁₂} (f : M →ₛₗ[σ₁₂] M₂) (g : M₂ →ₛₗ[σ₂₁] M) {h₁ : f.comp g = LinearMap.id} {h₂ : g.comp f = LinearMap.id} :
                      (LinearEquiv.ofLinear f g h₁ h₂) = f
                      @[simp]
                      theorem LinearEquiv.ofLinear_symm_toLinearMap {R : Type u_1} {R₂ : Type u_3} {M : Type u_9} {M₂ : Type u_12} [Semiring R] [Semiring R₂] [AddCommMonoid M] [AddCommMonoid M₂] {module_M : Module R M} {module_M₂ : Module R₂ M₂} {σ₁₂ : R →+* R₂} {σ₂₁ : R₂ →+* R} {re₁₂ : RingHomInvPair σ₁₂ σ₂₁} {re₂₁ : RingHomInvPair σ₂₁ σ₁₂} (f : M →ₛₗ[σ₁₂] M₂) (g : M₂ →ₛₗ[σ₂₁] M) {h₁ : f.comp g = LinearMap.id} {h₂ : g.comp f = LinearMap.id} :
                      (LinearEquiv.ofLinear f g h₁ h₂).symm = g
                      @[simp]
                      theorem LinearEquiv.range {R : Type u_1} {R₂ : Type u_3} {M : Type u_9} {M₂ : Type u_12} [Semiring R] [Semiring R₂] [AddCommMonoid M] [AddCommMonoid M₂] {module_M : Module R M} {module_M₂ : Module R₂ M₂} {σ₁₂ : R →+* R₂} {σ₂₁ : R₂ →+* R} {re₁₂ : RingHomInvPair σ₁₂ σ₂₁} {re₂₁ : RingHomInvPair σ₂₁ σ₁₂} (e : M ≃ₛₗ[σ₁₂] M₂) :
                      @[simp]
                      theorem LinearEquivClass.range {R : Type u_1} {R₂ : Type u_3} {M : Type u_9} {M₂ : Type u_12} [Semiring R] [Semiring R₂] [AddCommMonoid M] [AddCommMonoid M₂] {σ₁₂ : R →+* R₂} {σ₂₁ : R₂ →+* R} {re₁₂ : RingHomInvPair σ₁₂ σ₂₁} {re₂₁ : RingHomInvPair σ₂₁ σ₁₂} [Module R M] [Module R₂ M₂] {F : Type u_20} [EquivLike F M M₂] [SemilinearEquivClass F σ₁₂ M M₂] (e : F) :
                      theorem LinearEquiv.eq_bot_of_equiv {R : Type u_1} {R₂ : Type u_3} {M : Type u_9} {M₂ : Type u_12} [Semiring R] [Semiring R₂] [AddCommMonoid M] [AddCommMonoid M₂] {module_M : Module R M} {σ₁₂ : R →+* R₂} {σ₂₁ : R₂ →+* R} {re₁₂ : RingHomInvPair σ₁₂ σ₂₁} {re₂₁ : RingHomInvPair σ₂₁ σ₁₂} (p : Submodule R M) [Module R₂ M₂] (e : p ≃ₛₗ[σ₁₂] ) :
                      p =
                      @[simp]
                      theorem LinearEquiv.range_comp {R : Type u_1} {R₂ : Type u_3} {R₃ : Type u_4} {M : Type u_9} {M₂ : Type u_12} {M₃ : Type u_13} [Semiring R] [Semiring R₂] [Semiring R₃] [AddCommMonoid M] [AddCommMonoid M₂] [AddCommMonoid M₃] {module_M : Module R M} {module_M₂ : Module R₂ M₂} {module_M₃ : Module R₃ M₃} {σ₁₂ : R →+* R₂} {σ₂₁ : R₂ →+* R} {σ₂₃ : R₂ →+* R₃} {σ₁₃ : R →+* R₃} [RingHomCompTriple σ₁₂ σ₂₃ σ₁₃] {re₁₂ : RingHomInvPair σ₁₂ σ₂₁} {re₂₁ : RingHomInvPair σ₂₁ σ₁₂} (e : M ≃ₛₗ[σ₁₂] M₂) (h : M₂ →ₛₗ[σ₂₃] M₃) [RingHomSurjective σ₂₃] [RingHomSurjective σ₁₃] :
                      def LinearEquiv.ofLeftInverse {R : Type u_1} {R₂ : Type u_3} {M : Type u_9} {M₂ : Type u_12} [Semiring R] [Semiring R₂] [AddCommMonoid M] [AddCommMonoid M₂] {module_M : Module R M} {module_M₂ : Module R₂ M₂} {σ₁₂ : R →+* R₂} {σ₂₁ : R₂ →+* R} {f : M →ₛₗ[σ₁₂] M₂} [RingHomInvPair σ₁₂ σ₂₁] [RingHomInvPair σ₂₁ σ₁₂] {g : M₂M} (h : Function.LeftInverse g f) :
                      M ≃ₛₗ[σ₁₂] (LinearMap.range f)

                      A linear map f : M →ₗ[R] M₂ with a left-inverse g : M₂ →ₗ[R] M defines a linear equivalence between M and f.range.

                      This is a computable alternative to LinearEquiv.ofInjective, and a bidirectional version of LinearMap.rangeRestrict.

                      Equations
                      • One or more equations did not get rendered due to their size.
                      Instances For
                        @[simp]
                        theorem LinearEquiv.ofLeftInverse_apply {R : Type u_1} {R₂ : Type u_3} {M : Type u_9} {M₂ : Type u_12} [Semiring R] [Semiring R₂] [AddCommMonoid M] [AddCommMonoid M₂] {module_M : Module R M} {module_M₂ : Module R₂ M₂} {σ₁₂ : R →+* R₂} {σ₂₁ : R₂ →+* R} {f : M →ₛₗ[σ₁₂] M₂} {g : M₂ →ₛₗ[σ₂₁] M} [RingHomInvPair σ₁₂ σ₂₁] [RingHomInvPair σ₂₁ σ₁₂] (h : Function.LeftInverse g f) (x : M) :
                        @[simp]
                        theorem LinearEquiv.ofLeftInverse_symm_apply {R : Type u_1} {R₂ : Type u_3} {M : Type u_9} {M₂ : Type u_12} [Semiring R] [Semiring R₂] [AddCommMonoid M] [AddCommMonoid M₂] {module_M : Module R M} {module_M₂ : Module R₂ M₂} {σ₁₂ : R →+* R₂} {σ₂₁ : R₂ →+* R} {f : M →ₛₗ[σ₁₂] M₂} {g : M₂ →ₛₗ[σ₂₁] M} [RingHomInvPair σ₁₂ σ₂₁] [RingHomInvPair σ₂₁ σ₁₂] (h : Function.LeftInverse g f) (x : (LinearMap.range f)) :
                        (LinearEquiv.ofLeftInverse h).symm x = g x
                        noncomputable def LinearEquiv.ofInjective {R : Type u_1} {R₂ : Type u_3} {M : Type u_9} {M₂ : Type u_12} [Semiring R] [Semiring R₂] [AddCommMonoid M] [AddCommMonoid M₂] {module_M : Module R M} {module_M₂ : Module R₂ M₂} {σ₁₂ : R →+* R₂} {σ₂₁ : R₂ →+* R} (f : M →ₛₗ[σ₁₂] M₂) [RingHomInvPair σ₁₂ σ₂₁] [RingHomInvPair σ₂₁ σ₁₂] (h : Function.Injective f) :
                        M ≃ₛₗ[σ₁₂] (LinearMap.range f)

                        An Injective linear map f : M →ₗ[R] M₂ defines a linear equivalence between M and f.range. See also LinearMap.ofLeftInverse.

                        Equations
                        Instances For
                          @[simp]
                          theorem LinearEquiv.ofInjective_apply {R : Type u_1} {R₂ : Type u_3} {M : Type u_9} {M₂ : Type u_12} [Semiring R] [Semiring R₂] [AddCommMonoid M] [AddCommMonoid M₂] {module_M : Module R M} {module_M₂ : Module R₂ M₂} {σ₁₂ : R →+* R₂} {σ₂₁ : R₂ →+* R} (f : M →ₛₗ[σ₁₂] M₂) [RingHomInvPair σ₁₂ σ₂₁] [RingHomInvPair σ₂₁ σ₁₂] {h : Function.Injective f} (x : M) :
                          ((LinearEquiv.ofInjective f h) x) = f x
                          noncomputable def LinearEquiv.ofBijective {R : Type u_1} {R₂ : Type u_3} {M : Type u_9} {M₂ : Type u_12} [Semiring R] [Semiring R₂] [AddCommMonoid M] [AddCommMonoid M₂] {module_M : Module R M} {module_M₂ : Module R₂ M₂} {σ₁₂ : R →+* R₂} {σ₂₁ : R₂ →+* R} (f : M →ₛₗ[σ₁₂] M₂) [RingHomInvPair σ₁₂ σ₂₁] [RingHomInvPair σ₂₁ σ₁₂] (hf : Function.Bijective f) :
                          M ≃ₛₗ[σ₁₂] M₂

                          A bijective linear map is a linear equivalence.

                          Equations
                          Instances For
                            @[simp]
                            theorem LinearEquiv.ofBijective_apply {R : Type u_1} {R₂ : Type u_3} {M : Type u_9} {M₂ : Type u_12} [Semiring R] [Semiring R₂] [AddCommMonoid M] [AddCommMonoid M₂] {module_M : Module R M} {module_M₂ : Module R₂ M₂} {σ₁₂ : R →+* R₂} {σ₂₁ : R₂ →+* R} (f : M →ₛₗ[σ₁₂] M₂) [RingHomInvPair σ₁₂ σ₂₁] [RingHomInvPair σ₂₁ σ₁₂] {hf : Function.Bijective f} (x : M) :
                            @[simp]
                            theorem LinearEquiv.ofBijective_symm_apply_apply {R : Type u_1} {R₂ : Type u_3} {M : Type u_9} {M₂ : Type u_12} [Semiring R] [Semiring R₂] [AddCommMonoid M] [AddCommMonoid M₂] {module_M : Module R M} {module_M₂ : Module R₂ M₂} {σ₁₂ : R →+* R₂} {σ₂₁ : R₂ →+* R} (f : M →ₛₗ[σ₁₂] M₂) [RingHomInvPair σ₁₂ σ₂₁] [RingHomInvPair σ₂₁ σ₁₂] {h : Function.Bijective f} (x : M) :
                            (LinearEquiv.ofBijective f h).symm (f x) = x
                            def LinearEquiv.neg (R : Type u_1) {M : Type u_9} [Semiring R] [AddCommGroup M] [Module R M] :

                            x ↦ -x as a LinearEquiv

                            Equations
                            • LinearEquiv.neg R = let __src := Equiv.neg M; let __src_1 := -LinearMap.id; { toFun := __src.toFun, map_add' := , map_smul' := , invFun := __src.invFun, left_inv := , right_inv := }
                            Instances For
                              @[simp]
                              theorem LinearEquiv.coe_neg {R : Type u_1} {M : Type u_9} [Semiring R] [AddCommGroup M] [Module R M] :
                              theorem LinearEquiv.neg_apply {R : Type u_1} {M : Type u_9} [Semiring R] [AddCommGroup M] [Module R M] (x : M) :
                              @[simp]
                              theorem LinearEquiv.symm_neg {R : Type u_1} {M : Type u_9} [Semiring R] [AddCommGroup M] [Module R M] :
                              def LinearEquiv.smulOfUnit {R : Type u_1} {M : Type u_9} [CommSemiring R] [AddCommMonoid M] [Module R M] (a : Rˣ) :

                              Multiplying by a unit a of the ring R is a linear equivalence.

                              Equations
                              Instances For
                                def LinearEquiv.arrowCongr {R : Type u_20} {M₁ : Type u_21} {M₂ : Type u_22} {M₂₁ : Type u_23} {M₂₂ : Type u_24} [CommSemiring R] [AddCommMonoid M₁] [AddCommMonoid M₂] [AddCommMonoid M₂₁] [AddCommMonoid M₂₂] [Module R M₁] [Module R M₂] [Module R M₂₁] [Module R M₂₂] (e₁ : M₁ ≃ₗ[R] M₂) (e₂ : M₂₁ ≃ₗ[R] M₂₂) :
                                (M₁ →ₗ[R] M₂₁) ≃ₗ[R] M₂ →ₗ[R] M₂₂

                                A linear isomorphism between the domains and codomains of two spaces of linear maps gives a linear isomorphism between the two function spaces.

                                Equations
                                • One or more equations did not get rendered due to their size.
                                Instances For
                                  @[simp]
                                  theorem LinearEquiv.arrowCongr_apply {R : Type u_20} {M₁ : Type u_21} {M₂ : Type u_22} {M₂₁ : Type u_23} {M₂₂ : Type u_24} [CommSemiring R] [AddCommMonoid M₁] [AddCommMonoid M₂] [AddCommMonoid M₂₁] [AddCommMonoid M₂₂] [Module R M₁] [Module R M₂] [Module R M₂₁] [Module R M₂₂] (e₁ : M₁ ≃ₗ[R] M₂) (e₂ : M₂₁ ≃ₗ[R] M₂₂) (f : M₁ →ₗ[R] M₂₁) (x : M₂) :
                                  ((e₁.arrowCongr e₂) f) x = e₂ (f (e₁.symm x))
                                  @[simp]
                                  theorem LinearEquiv.arrowCongr_symm_apply {R : Type u_20} {M₁ : Type u_21} {M₂ : Type u_22} {M₂₁ : Type u_23} {M₂₂ : Type u_24} [CommSemiring R] [AddCommMonoid M₁] [AddCommMonoid M₂] [AddCommMonoid M₂₁] [AddCommMonoid M₂₂] [Module R M₁] [Module R M₂] [Module R M₂₁] [Module R M₂₂] (e₁ : M₁ ≃ₗ[R] M₂) (e₂ : M₂₁ ≃ₗ[R] M₂₂) (f : M₂ →ₗ[R] M₂₂) (x : M₁) :
                                  ((e₁.arrowCongr e₂).symm f) x = e₂.symm (f (e₁ x))
                                  theorem LinearEquiv.arrowCongr_comp {R : Type u_1} {M : Type u_9} {M₂ : Type u_12} {M₃ : Type u_13} [CommSemiring R] [AddCommMonoid M] [AddCommMonoid M₂] [AddCommMonoid M₃] [Module R M] [Module R M₂] [Module R M₃] {N : Type u_20} {N₂ : Type u_21} {N₃ : Type u_22} [AddCommMonoid N] [AddCommMonoid N₂] [AddCommMonoid N₃] [Module R N] [Module R N₂] [Module R N₃] (e₁ : M ≃ₗ[R] N) (e₂ : M₂ ≃ₗ[R] N₂) (e₃ : M₃ ≃ₗ[R] N₃) (f : M →ₗ[R] M₂) (g : M₂ →ₗ[R] M₃) :
                                  (e₁.arrowCongr e₃) (g ∘ₗ f) = (e₂.arrowCongr e₃) g ∘ₗ (e₁.arrowCongr e₂) f
                                  theorem LinearEquiv.arrowCongr_trans {R : Type u_1} [CommSemiring R] {M₁ : Type u_20} {M₂ : Type u_21} {M₃ : Type u_22} {N₁ : Type u_23} {N₂ : Type u_24} {N₃ : Type u_25} [AddCommMonoid M₁] [Module R M₁] [AddCommMonoid M₂] [Module R M₂] [AddCommMonoid M₃] [Module R M₃] [AddCommMonoid N₁] [Module R N₁] [AddCommMonoid N₂] [Module R N₂] [AddCommMonoid N₃] [Module R N₃] (e₁ : M₁ ≃ₗ[R] M₂) (e₂ : N₁ ≃ₗ[R] N₂) (e₃ : M₂ ≃ₗ[R] M₃) (e₄ : N₂ ≃ₗ[R] N₃) :
                                  e₁.arrowCongr e₂ ≪≫ₗ e₃.arrowCongr e₄ = (e₁ ≪≫ₗ e₃).arrowCongr (e₂ ≪≫ₗ e₄)
                                  def LinearEquiv.congrRight {R : Type u_1} {M : Type u_9} {M₂ : Type u_12} {M₃ : Type u_13} [CommSemiring R] [AddCommMonoid M] [AddCommMonoid M₂] [AddCommMonoid M₃] [Module R M] [Module R M₂] [Module R M₃] (f : M₂ ≃ₗ[R] M₃) :
                                  (M →ₗ[R] M₂) ≃ₗ[R] M →ₗ[R] M₃

                                  If M₂ and M₃ are linearly isomorphic then the two spaces of linear maps from M into M₂ and M into M₃ are linearly isomorphic.

                                  Equations
                                  Instances For
                                    def LinearEquiv.conj {R : Type u_1} {M : Type u_9} {M₂ : Type u_12} [CommSemiring R] [AddCommMonoid M] [AddCommMonoid M₂] [Module R M] [Module R M₂] (e : M ≃ₗ[R] M₂) :

                                    If M and M₂ are linearly isomorphic then the two spaces of linear maps from M and M₂ to themselves are linearly isomorphic.

                                    Equations
                                    • e.conj = e.arrowCongr e
                                    Instances For
                                      theorem LinearEquiv.conj_apply {R : Type u_1} {M : Type u_9} {M₂ : Type u_12} [CommSemiring R] [AddCommMonoid M] [AddCommMonoid M₂] [Module R M] [Module R M₂] (e : M ≃ₗ[R] M₂) (f : Module.End R M) :
                                      e.conj f = (e ∘ₗ f) ∘ₗ e.symm
                                      theorem LinearEquiv.conj_apply_apply {R : Type u_1} {M : Type u_9} {M₂ : Type u_12} [CommSemiring R] [AddCommMonoid M] [AddCommMonoid M₂] [Module R M] [Module R M₂] (e : M ≃ₗ[R] M₂) (f : Module.End R M) (x : M₂) :
                                      (e.conj f) x = e (f (e.symm x))
                                      theorem LinearEquiv.symm_conj_apply {R : Type u_1} {M : Type u_9} {M₂ : Type u_12} [CommSemiring R] [AddCommMonoid M] [AddCommMonoid M₂] [Module R M] [Module R M₂] (e : M ≃ₗ[R] M₂) (f : Module.End R M₂) :
                                      e.symm.conj f = (e.symm ∘ₗ f) ∘ₗ e
                                      theorem LinearEquiv.conj_comp {R : Type u_1} {M : Type u_9} {M₂ : Type u_12} [CommSemiring R] [AddCommMonoid M] [AddCommMonoid M₂] [Module R M] [Module R M₂] (e : M ≃ₗ[R] M₂) (f : Module.End R M) (g : Module.End R M) :
                                      e.conj (g ∘ₗ f) = e.conj g ∘ₗ e.conj f
                                      theorem LinearEquiv.conj_trans {R : Type u_1} {M : Type u_9} {M₂ : Type u_12} {M₃ : Type u_13} [CommSemiring R] [AddCommMonoid M] [AddCommMonoid M₂] [AddCommMonoid M₃] [Module R M] [Module R M₂] [Module R M₃] (e₁ : M ≃ₗ[R] M₂) (e₂ : M₂ ≃ₗ[R] M₃) :
                                      e₁.conj ≪≫ₗ e₂.conj = (e₁ ≪≫ₗ e₂).conj
                                      @[simp]
                                      theorem LinearEquiv.conj_id {R : Type u_1} {M : Type u_9} {M₂ : Type u_12} [CommSemiring R] [AddCommMonoid M] [AddCommMonoid M₂] [Module R M] [Module R M₂] (e : M ≃ₗ[R] M₂) :
                                      e.conj LinearMap.id = LinearMap.id
                                      def LinearEquiv.congrLeft (M : Type u_9) {M₂ : Type u_12} {M₃ : Type u_13} [AddCommMonoid M] [AddCommMonoid M₂] [AddCommMonoid M₃] {R : Type u_20} (S : Type u_21) [Semiring R] [Semiring S] [Module R M₂] [Module R M₃] [Module R M] [Module S M] [SMulCommClass R S M] (e : M₂ ≃ₗ[R] M₃) :
                                      (M₂ →ₗ[R] M) ≃ₗ[S] M₃ →ₗ[R] M

                                      An R-linear isomorphism between two R-modules M₂ and M₃ induces an S-linear isomorphism between M₂ →ₗ[R] M and M₃ →ₗ[R] M, if M is both an R-module and an S-module and their actions commute.

                                      Equations
                                      • LinearEquiv.congrLeft M S e = { toFun := fun (f : M₂ →ₗ[R] M) => f ∘ₗ e.symm, map_add' := , map_smul' := , invFun := fun (f : M₃ →ₗ[R] M) => f ∘ₗ e, left_inv := , right_inv := }
                                      Instances For
                                        @[simp]
                                        theorem LinearEquiv.smulOfNeZero_symm_apply (K : Type u_7) (M : Type u_9) [Field K] [AddCommGroup M] [Module K M] (a : K) (ha : a 0) :
                                        ∀ (a_1 : M), (LinearEquiv.smulOfNeZero K M a ha).symm a_1 = (Units.mk0 a ha)⁻¹ a_1
                                        @[simp]
                                        theorem LinearEquiv.smulOfNeZero_apply (K : Type u_7) (M : Type u_9) [Field K] [AddCommGroup M] [Module K M] (a : K) (ha : a 0) :
                                        ∀ (a_1 : M), (LinearEquiv.smulOfNeZero K M a ha) a_1 = a a_1
                                        def LinearEquiv.smulOfNeZero (K : Type u_7) (M : Type u_9) [Field K] [AddCommGroup M] [Module K M] (a : K) (ha : a 0) :

                                        Multiplying by a nonzero element a of the field K is a linear equivalence.

                                        Equations
                                        Instances For
                                          def Submodule.equivSubtypeMap {R : Type u_1} {M : Type u_9} [Semiring R] [AddCommMonoid M] [Module R M] (p : Submodule R M) (q : Submodule R p) :
                                          q ≃ₗ[R] (Submodule.map p.subtype q)

                                          Given p a submodule of the module M and q a submodule of p, p.equivSubtypeMap q is the natural LinearEquiv between q and q.map p.subtype.

                                          Equations
                                          • One or more equations did not get rendered due to their size.
                                          Instances For
                                            @[simp]
                                            theorem Submodule.equivSubtypeMap_apply {R : Type u_1} {M : Type u_9} [Semiring R] [AddCommMonoid M] [Module R M] {p : Submodule R M} {q : Submodule R p} (x : q) :
                                            ((p.equivSubtypeMap q) x) = (p.subtype.domRestrict q) x
                                            @[simp]
                                            theorem Submodule.equivSubtypeMap_symm_apply {R : Type u_1} {M : Type u_9} [Semiring R] [AddCommMonoid M] [Module R M] {p : Submodule R M} {q : Submodule R p} (x : (Submodule.map p.subtype q)) :
                                            ((p.equivSubtypeMap q).symm x) = x
                                            @[simp]
                                            theorem Submodule.comap_equiv_self_of_inj_of_le_apply {R : Type u_1} {M : Type u_9} {N : Type u_15} [Semiring R] [AddCommMonoid M] [Module R M] [AddCommMonoid N] [Module R N] {f : M →ₗ[R] N} {p : Submodule R N} (hf : Function.Injective f) (h : p LinearMap.range f) :
                                            ∀ (a : (Submodule.comap f p)), (Submodule.comap_equiv_self_of_inj_of_le hf h) a = (LinearMap.codRestrict p (f ∘ₗ (Submodule.comap f p).subtype) ) a
                                            noncomputable def Submodule.comap_equiv_self_of_inj_of_le {R : Type u_1} {M : Type u_9} {N : Type u_15} [Semiring R] [AddCommMonoid M] [Module R M] [AddCommMonoid N] [Module R N] {f : M →ₗ[R] N} {p : Submodule R N} (hf : Function.Injective f) (h : p LinearMap.range f) :
                                            (Submodule.comap f p) ≃ₗ[R] p

                                            A linear injection M ↪ N restricts to an equivalence f⁻¹ p ≃ p for any submodule p contained in its range.

                                            Equations
                                            Instances For
                                              def Equiv.toLinearEquiv {R : Type u_1} {M : Type u_9} {M₂ : Type u_12} [Semiring R] [AddCommMonoid M] [Module R M] [AddCommMonoid M₂] [Module R M₂] (e : M M₂) (h : IsLinearMap R e) :
                                              M ≃ₗ[R] M₂

                                              An equivalence whose underlying function is linear is a linear equivalence.

                                              Equations
                                              • e.toLinearEquiv h = let __src := IsLinearMap.mk' (e) h; { toFun := e.toFun, map_add' := , map_smul' := , invFun := e.invFun, left_inv := , right_inv := }
                                              Instances For
                                                def LinearMap.funLeft (R : Type u_1) (M : Type u_9) [Semiring R] [AddCommMonoid M] [Module R M] {m : Type u_20} {n : Type u_21} (f : mn) :
                                                (nM) →ₗ[R] mM

                                                Given an R-module M and a function m → n between arbitrary types, construct a linear map (n → M) →ₗ[R] (m → M)

                                                Equations
                                                Instances For
                                                  @[simp]
                                                  theorem LinearMap.funLeft_apply (R : Type u_1) (M : Type u_9) [Semiring R] [AddCommMonoid M] [Module R M] {m : Type u_20} {n : Type u_21} (f : mn) (g : nM) (i : m) :
                                                  (LinearMap.funLeft R M f) g i = g (f i)
                                                  @[simp]
                                                  theorem LinearMap.funLeft_id (R : Type u_1) (M : Type u_9) [Semiring R] [AddCommMonoid M] [Module R M] {n : Type u_21} (g : nM) :
                                                  (LinearMap.funLeft R M id) g = g
                                                  theorem LinearMap.funLeft_comp (R : Type u_1) (M : Type u_9) [Semiring R] [AddCommMonoid M] [Module R M] {m : Type u_20} {n : Type u_21} {p : Type u_22} (f₁ : np) (f₂ : mn) :
                                                  LinearMap.funLeft R M (f₁ f₂) = LinearMap.funLeft R M f₂ ∘ₗ LinearMap.funLeft R M f₁
                                                  theorem LinearMap.funLeft_surjective_of_injective (R : Type u_1) (M : Type u_9) [Semiring R] [AddCommMonoid M] [Module R M] {m : Type u_20} {n : Type u_21} (f : mn) (hf : Function.Injective f) :
                                                  theorem LinearMap.funLeft_injective_of_surjective (R : Type u_1) (M : Type u_9) [Semiring R] [AddCommMonoid M] [Module R M] {m : Type u_20} {n : Type u_21} (f : mn) (hf : Function.Surjective f) :
                                                  def LinearEquiv.funCongrLeft (R : Type u_1) (M : Type u_9) [Semiring R] [AddCommMonoid M] [Module R M] {m : Type u_20} {n : Type u_21} (e : m n) :
                                                  (nM) ≃ₗ[R] mM

                                                  Given an R-module M and an equivalence m ≃ n between arbitrary types, construct a linear equivalence (n → M) ≃ₗ[R] (m → M)

                                                  Equations
                                                  Instances For
                                                    @[simp]
                                                    theorem LinearEquiv.funCongrLeft_apply (R : Type u_1) (M : Type u_9) [Semiring R] [AddCommMonoid M] [Module R M] {m : Type u_20} {n : Type u_21} (e : m n) (x : nM) :
                                                    @[simp]
                                                    theorem LinearEquiv.funCongrLeft_id (R : Type u_1) (M : Type u_9) [Semiring R] [AddCommMonoid M] [Module R M] {n : Type u_21} :
                                                    @[simp]
                                                    theorem LinearEquiv.funCongrLeft_comp (R : Type u_1) (M : Type u_9) [Semiring R] [AddCommMonoid M] [Module R M] {m : Type u_20} {n : Type u_21} {p : Type u_22} (e₁ : m n) (e₂ : n p) :
                                                    LinearEquiv.funCongrLeft R M (e₁.trans e₂) = LinearEquiv.funCongrLeft R M e₂ ≪≫ₗ LinearEquiv.funCongrLeft R M e₁
                                                    @[simp]
                                                    theorem LinearEquiv.funCongrLeft_symm (R : Type u_1) (M : Type u_9) [Semiring R] [AddCommMonoid M] [Module R M] {m : Type u_20} {n : Type u_21} (e : m n) :