mathlib documentation

linear_algebra.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. If p and q are submodules of a module, p ≤ q means that p ⊆ q.

Many of the relevant definitions, including module, submodule, and linear_map, are found in src/algebra/module.lean.

Main definitions

Main statements

Notations

Implementation notes

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

Tags

linear algebra, vector space, module

theorem finsupp.smul_sum {α : Type u} {β : Type v} {R : Type w} {M : Type y} [has_zero β] [semiring R] [add_comm_monoid M] [semimodule R M] {v : α →₀ β} {c : R} {h : α → β → M} :
c v.sum h = v.sum (λ (a : α) (b : β), c h a b)

theorem pi_eq_sum_univ {ι : Type u} [fintype ι] {R : Type v} [semiring R] (x : ι → R) :
x = ∑ (i : ι), x i λ (j : ι), ite (i = j) 1 0

decomposing x : ι → R as a sum along the canonical basis

Properties of linear maps

@[simp]
theorem linear_map.comp_id {R : Type u} {M : Type v} {M₂ : Type w} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] [semimodule R M] [semimodule R M₂] (f : M →ₗ[R] M₂) :

@[simp]
theorem linear_map.id_comp {R : Type u} {M : Type v} {M₂ : Type w} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] [semimodule R M] [semimodule R M₂] (f : M →ₗ[R] M₂) :

theorem linear_map.comp_assoc {R : Type u} {M : Type v} {M₂ : Type w} {M₃ : Type y} {M₄ : Type z} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] [add_comm_monoid M₃] [add_comm_monoid M₄] [semimodule R M] [semimodule R M₂] [semimodule R M₃] [semimodule R M₄] (f : M →ₗ[R] M₂) (g : M₂ →ₗ[R] M₃) (h : M₃ →ₗ[R] M₄) :
(h.comp g).comp f = h.comp (g.comp f)

def linear_map.dom_restrict {R : Type u} {M : Type v} {M₂ : Type w} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] [semimodule R M] [semimodule R M₂] (f : M →ₗ[R] M₂) (p : submodule R M) :
p →ₗ[R] M₂

The restriction of a linear map f : M → M₂ to a submodule p ⊆ M gives a linear map p → M₂.

Equations
@[simp]
theorem linear_map.dom_restrict_apply {R : Type u} {M : Type v} {M₂ : Type w} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] [semimodule R M] [semimodule R M₂] (f : M →ₗ[R] M₂) (p : submodule R M) (x : p) :

def linear_map.cod_restrict {R : Type u} {M : Type v} {M₂ : Type w} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] [semimodule R M] [semimodule R M₂] (p : submodule R M) (f : M₂ →ₗ[R] M) (h : ∀ (c : M₂), f c p) :
M₂ →ₗ[R] p

A linear map f : M₂ → M whose values lie in a submodule p ⊆ M can be restricted to a linear map M₂ → p.

Equations
@[simp]
theorem linear_map.cod_restrict_apply {R : Type u} {M : Type v} {M₂ : Type w} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] [semimodule R M] [semimodule R M₂] (p : submodule R M) (f : M₂ →ₗ[R] M) {h : ∀ (c : M₂), f c p} (x : M₂) :

@[simp]
theorem linear_map.comp_cod_restrict {R : Type u} {M : Type v} {M₂ : Type w} {M₃ : Type y} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] [add_comm_monoid M₃] [semimodule R M] [semimodule R M₂] [semimodule R M₃] (f : M →ₗ[R] M₂) (p : submodule R M₂) (h : ∀ (b : M), f b p) (g : M₃ →ₗ[R] M) :

@[simp]
theorem linear_map.subtype_comp_cod_restrict {R : Type u} {M : Type v} {M₂ : Type w} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] [semimodule R M] [semimodule R M₂] (f : M →ₗ[R] M₂) (p : submodule R M₂) (h : ∀ (b : M), f b p) :

def linear_map.restrict {R : Type u} {M : Type v} [semiring R] [add_comm_monoid M] [semimodule R M] (f : M →ₗ[R] M) {p : submodule R M} (hf : ∀ (x : M), x pf x p) :

Restrict domain and codomain of an endomorphism.

Equations
theorem linear_map.restrict_apply {R : Type u} {M : Type v} [semiring R] [add_comm_monoid M] [semimodule R M] {f : M →ₗ[R] M} {p : submodule R M} (hf : ∀ (x : M), x pf x p) (x : p) :
(f.restrict hf) x = f x, _⟩

theorem linear_map.subtype_comp_restrict {R : Type u} {M : Type v} [semiring R] [add_comm_monoid M] [semimodule R M] {f : M →ₗ[R] M} {p : submodule R M} (hf : ∀ (x : M), x pf x p) :

theorem linear_map.restrict_eq_cod_restrict_dom_restrict {R : Type u} {M : Type v} [semiring R] [add_comm_monoid M] [semimodule R M] {f : M →ₗ[R] M} {p : submodule R M} (hf : ∀ (x : M), x pf x p) :

theorem linear_map.restrict_eq_dom_restrict_cod_restrict {R : Type u} {M : Type v} [semiring R] [add_comm_monoid M] [semimodule R M] {f : M →ₗ[R] M} {p : submodule R M} (hf : ∀ (x : M), f x p) :

def linear_map.inverse {R : Type u} {M : Type v} {M₂ : Type w} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] [semimodule R M] [semimodule R M₂] (f : M →ₗ[R] M₂) (g : M₂ → M) (h₁ : function.left_inverse g f) (h₂ : function.right_inverse g f) :
M₂ →ₗ[R] M

If a function g is a left and right inverse of a linear map f, then g is linear itself.

Equations
@[instance]
def linear_map.has_zero {R : Type u} {M : Type v} {M₂ : Type w} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] [semimodule R M] [semimodule R M₂] :
has_zero (M →ₗ[R] M₂)

The constant 0 map is linear.

Equations
@[instance]
def linear_map.inhabited {R : Type u} {M : Type v} {M₂ : Type w} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] [semimodule R M] [semimodule R M₂] :

Equations
@[simp]
theorem linear_map.zero_apply {R : Type u} {M : Type v} {M₂ : Type w} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] [semimodule R M] [semimodule R M₂] (x : M) :
0 x = 0

@[instance]
def linear_map.has_add {R : Type u} {M : Type v} {M₂ : Type w} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] [semimodule R M] [semimodule R M₂] :
has_add (M →ₗ[R] M₂)

The sum of two linear maps is linear.

Equations
@[simp]
theorem linear_map.add_apply {R : Type u} {M : Type v} {M₂ : Type w} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] [semimodule R M] [semimodule R M₂] (f g : M →ₗ[R] M₂) (x : M) :
(f + g) x = f x + g x

@[instance]
def linear_map.add_comm_monoid {R : Type u} {M : Type v} {M₂ : Type w} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] [semimodule R M] [semimodule R M₂] :

The type of linear maps is an additive monoid.

Equations
@[instance]
def linear_map.linear_map_apply_is_add_monoid_hom {R : Type u} {M : Type v} {M₂ : Type w} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] [semimodule R M] [semimodule R M₂] (a : M) :
is_add_monoid_hom (λ (f : M →ₗ[R] M₂), f a)

theorem linear_map.add_comp {R : Type u} {M : Type v} {M₂ : Type w} {M₃ : Type y} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] [add_comm_monoid M₃] [semimodule R M] [semimodule R M₂] [semimodule R M₃] (f : M →ₗ[R] M₂) (g h : M₂ →ₗ[R] M₃) :
(h + g).comp f = h.comp f + g.comp f

theorem linear_map.comp_add {R : Type u} {M : Type v} {M₂ : Type w} {M₃ : Type y} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] [add_comm_monoid M₃] [semimodule R M] [semimodule R M₂] [semimodule R M₃] (f g : M →ₗ[R] M₂) (h : M₂ →ₗ[R] M₃) :
h.comp (f + g) = h.comp f + h.comp g

theorem linear_map.sum_apply {R : Type u} {M : Type v} {M₂ : Type w} {ι : Type x} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] [semimodule R M] [semimodule R M₂] (t : finset ι) (f : ι → (M →ₗ[R] M₂)) (b : M) :
(∑ (d : ι) in t, f d) b = ∑ (d : ι) in t, (f d) b

def linear_map.smul_right {R : Type u} {M : Type v} {M₂ : Type w} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] [semimodule R M] [semimodule R M₂] (f : M₂ →ₗ[R] R) (x : M) :
M₂ →ₗ[R] M

λb, f b • x is a linear map.

Equations
@[simp]
theorem linear_map.smul_right_apply {R : Type u} {M : Type v} {M₂ : Type w} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] [semimodule R M] [semimodule R M₂] (f : M₂ →ₗ[R] R) (x : M) (c : M₂) :
(f.smul_right x) c = f c x

@[instance]
def linear_map.has_one {R : Type u} {M : Type v} [semiring R] [add_comm_monoid M] [semimodule R M] :

Equations
@[instance]
def linear_map.has_mul {R : Type u} {M : Type v} [semiring R] [add_comm_monoid M] [semimodule R M] :

Equations
@[simp]
theorem linear_map.one_app {R : Type u} {M : Type v} [semiring R] [add_comm_monoid M] [semimodule R M] (x : M) :
1 x = x

@[simp]
theorem linear_map.mul_app {R : Type u} {M : Type v} [semiring R] [add_comm_monoid M] [semimodule R M] (A B : M →ₗ[R] M) (x : M) :
(A * B) x = A (B x)

@[simp]
theorem linear_map.comp_zero {R : Type u} {M : Type v} {M₂ : Type w} {M₃ : Type y} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] [add_comm_monoid M₃] [semimodule R M] [semimodule R M₂] [semimodule R M₃] (f : M →ₗ[R] M₂) :
f.comp 0 = 0

@[simp]
theorem linear_map.zero_comp {R : Type u} {M : Type v} {M₂ : Type w} {M₃ : Type y} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] [add_comm_monoid M₃] [semimodule R M] [semimodule R M₂] [semimodule R M₃] (f : M →ₗ[R] M₂) :
0.comp f = 0

theorem linear_map.coe_fn_sum {R : Type u} {M : Type v} {M₂ : Type w} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] [semimodule R M] [semimodule R M₂] {ι : Type u_1} (t : finset ι) (f : ι → (M →ₗ[R] M₂)) :
∑ (i : ι) in t, f i = ∑ (i : ι) in t, (f i)

@[instance]
def linear_map.monoid {R : Type u} {M : Type v} [semiring R] [add_comm_monoid M] [semimodule R M] :

Equations
theorem linear_map.pi_apply_eq_sum_univ {R : Type u} {M : Type v} {ι : Type x} [semiring R] [add_comm_monoid M] [semimodule R M] [fintype ι] (f : (ι → R) →ₗ[R] M) (x : ι → R) :
f x = ∑ (i : ι), x i f (λ (j : ι), ite (i = j) 1 0)

A linear map f applied to x : ι → R can be computed using the image under f of elements of the canonical basis.

def linear_map.fst (R : Type u) (M : Type v) (M₂ : Type w) [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] [semimodule R M] [semimodule R M₂] :
M × M₂ →ₗ[R] M

The first projection of a product is a linear map.

Equations
def linear_map.snd (R : Type u) (M : Type v) (M₂ : Type w) [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] [semimodule R M] [semimodule R M₂] :
M × M₂ →ₗ[R] M₂

The second projection of a product is a linear map.

Equations
@[simp]
theorem linear_map.fst_apply {R : Type u} {M : Type v} {M₂ : Type w} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] [semimodule R M] [semimodule R M₂] (x : M × M₂) :
(linear_map.fst R M M₂) x = x.fst

@[simp]
theorem linear_map.snd_apply {R : Type u} {M : Type v} {M₂ : Type w} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] [semimodule R M] [semimodule R M₂] (x : M × M₂) :
(linear_map.snd R M M₂) x = x.snd

def linear_map.prod {R : Type u} {M : Type v} {M₂ : Type w} {M₃ : Type y} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] [add_comm_monoid M₃] [semimodule R M] [semimodule R M₂] [semimodule R M₃] (f : M →ₗ[R] M₂) (g : M →ₗ[R] M₃) :
M →ₗ[R] M₂ × M₃

The prod of two linear maps is a linear map.

Equations
@[simp]
theorem linear_map.prod_apply {R : Type u} {M : Type v} {M₂ : Type w} {M₃ : Type y} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] [add_comm_monoid M₃] [semimodule R M] [semimodule R M₂] [semimodule R M₃] (f : M →ₗ[R] M₂) (g : M →ₗ[R] M₃) (x : M) :
(f.prod g) x = (f x, g x)

@[simp]
theorem linear_map.fst_prod {R : Type u} {M : Type v} {M₂ : Type w} {M₃ : Type y} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] [add_comm_monoid M₃] [semimodule R M] [semimodule R M₂] [semimodule R M₃] (f : M →ₗ[R] M₂) (g : M →ₗ[R] M₃) :
(linear_map.fst R M₂ M₃).comp (f.prod g) = f

@[simp]
theorem linear_map.snd_prod {R : Type u} {M : Type v} {M₂ : Type w} {M₃ : Type y} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] [add_comm_monoid M₃] [semimodule R M] [semimodule R M₂] [semimodule R M₃] (f : M →ₗ[R] M₂) (g : M →ₗ[R] M₃) :
(linear_map.snd R M₂ M₃).comp (f.prod g) = g

@[simp]
theorem linear_map.pair_fst_snd {R : Type u} {M : Type v} {M₂ : Type w} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] [semimodule R M] [semimodule R M₂] :

def linear_map.inl (R : Type u) (M : Type v) (M₂ : Type w) [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] [semimodule R M] [semimodule R M₂] :
M →ₗ[R] M × M₂

The left injection into a product is a linear map.

Equations
def linear_map.inr (R : Type u) (M : Type v) (M₂ : Type w) [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] [semimodule R M] [semimodule R M₂] :
M₂ →ₗ[R] M × M₂

The right injection into a product is a linear map.

Equations
@[simp]
theorem linear_map.inl_apply {R : Type u} {M : Type v} {M₂ : Type w} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] [semimodule R M] [semimodule R M₂] (x : M) :
(linear_map.inl R M M₂) x = (x, 0)

@[simp]
theorem linear_map.inr_apply {R : Type u} {M : Type v} {M₂ : Type w} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] [semimodule R M] [semimodule R M₂] (x : M₂) :
(linear_map.inr R M M₂) x = (0, x)

theorem linear_map.inl_injective {R : Type u} {M : Type v} {M₂ : Type w} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] [semimodule R M] [semimodule R M₂] :

theorem linear_map.inr_injective {R : Type u} {M : Type v} {M₂ : Type w} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] [semimodule R M] [semimodule R M₂] :

def linear_map.coprod {R : Type u} {M : Type v} {M₂ : Type w} {M₃ : Type y} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] [add_comm_monoid M₃] [semimodule R M] [semimodule R M₂] [semimodule R M₃] (f : M →ₗ[R] M₃) (g : M₂ →ₗ[R] M₃) :
M × M₂ →ₗ[R] M₃

The coprod function λ x : M × M₂, f x.1 + g x.2 is a linear map.

Equations
@[simp]
theorem linear_map.coprod_apply {R : Type u} {M : Type v} {M₂ : Type w} {M₃ : Type y} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] [add_comm_monoid M₃] [semimodule R M] [semimodule R M₂] [semimodule R M₃] (f : M →ₗ[R] M₃) (g : M₂ →ₗ[R] M₃) (x : M) (y : M₂) :
(f.coprod g) (x, y) = f x + g y

@[simp]
theorem linear_map.coprod_inl {R : Type u} {M : Type v} {M₂ : Type w} {M₃ : Type y} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] [add_comm_monoid M₃] [semimodule R M] [semimodule R M₂] [semimodule R M₃] (f : M →ₗ[R] M₃) (g : M₂ →ₗ[R] M₃) :
(f.coprod g).comp (linear_map.inl R M M₂) = f

@[simp]
theorem linear_map.coprod_inr {R : Type u} {M : Type v} {M₂ : Type w} {M₃ : Type y} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] [add_comm_monoid M₃] [semimodule R M] [semimodule R M₂] [semimodule R M₃] (f : M →ₗ[R] M₃) (g : M₂ →ₗ[R] M₃) :
(f.coprod g).comp (linear_map.inr R M M₂) = g

@[simp]
theorem linear_map.coprod_inl_inr {R : Type u} {M : Type v} {M₂ : Type w} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] [semimodule R M] [semimodule R M₂] :

theorem linear_map.fst_eq_coprod {R : Type u} {M : Type v} {M₂ : Type w} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] [semimodule R M] [semimodule R M₂] :

theorem linear_map.snd_eq_coprod {R : Type u} {M : Type v} {M₂ : Type w} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] [semimodule R M] [semimodule R M₂] :

theorem linear_map.inl_eq_prod {R : Type u} {M : Type v} {M₂ : Type w} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] [semimodule R M] [semimodule R M₂] :

theorem linear_map.inr_eq_prod {R : Type u} {M : Type v} {M₂ : Type w} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] [semimodule R M] [semimodule R M₂] :

def linear_map.prod_map {R : Type u} {M : Type v} {M₂ : Type w} {M₃ : Type y} {M₄ : Type z} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] [add_comm_monoid M₃] [add_comm_monoid M₄] [semimodule R M] [semimodule R M₂] [semimodule R M₃] [semimodule R M₄] (f : M →ₗ[R] M₃) (g : M₂ →ₗ[R] M₄) :
M × M₂ →ₗ[R] M₃ × M₄

prod.map of two linear maps.

Equations
@[simp]
theorem linear_map.prod_map_apply {R : Type u} {M : Type v} {M₂ : Type w} {M₃ : Type y} {M₄ : Type z} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] [add_comm_monoid M₃] [add_comm_monoid M₄] [semimodule R M] [semimodule R M₂] [semimodule R M₃] [semimodule R M₄] (f : M →ₗ[R] M₃) (g : M₂ →ₗ[R] M₄) (x : M × M₂) :
(f.prod_map g) x = (f x.fst, g x.snd)

@[instance]
def linear_map.has_neg {R : Type u} {M : Type v} {M₂ : Type w} [semiring R] [add_comm_group M] [add_comm_group M₂] [semimodule R M] [semimodule R M₂] :
has_neg (M →ₗ[R] M₂)

The negation of a linear map is linear.

Equations
@[simp]
theorem linear_map.neg_apply {R : Type u} {M : Type v} {M₂ : Type w} [semiring R] [add_comm_group M] [add_comm_group M₂] [semimodule R M] [semimodule R M₂] (f : M →ₗ[R] M₂) (x : M) :
(-f) x = -f x

@[simp]
theorem linear_map.comp_neg {R : Type u} {M : Type v} {M₂ : Type w} {M₃ : Type y} [semiring R] [add_comm_group M] [add_comm_group M₂] [add_comm_group M₃] [semimodule R M] [semimodule R M₂] [semimodule R M₃] (f : M →ₗ[R] M₂) (g : M₂ →ₗ[R] M₃) :
g.comp (-f) = -g.comp f

@[instance]
def linear_map.add_comm_group {R : Type u} {M : Type v} {M₂ : Type w} [semiring R] [add_comm_group M] [add_comm_group M₂] [semimodule R M] [semimodule R M₂] :

The type of linear maps is an additive group.

Equations
@[instance]
def linear_map.linear_map_apply_is_add_group_hom {R : Type u} {M : Type v} {M₂ : Type w} [semiring R] [add_comm_group M] [add_comm_group M₂] [semimodule R M] [semimodule R M₂] (a : M) :
is_add_group_hom (λ (f : M →ₗ[R] M₂), f a)

@[simp]
theorem linear_map.sub_apply {R : Type u} {M : Type v} {M₂ : Type w} [semiring R] [add_comm_group M] [add_comm_group M₂] [semimodule R M] [semimodule R M₂] (f g : M →ₗ[R] M₂) (x : M) :
(f - g) x = f x - g x

theorem linear_map.sub_comp {R : Type u} {M : Type v} {M₂ : Type w} {M₃ : Type y} [semiring R] [add_comm_group M] [add_comm_group M₂] [add_comm_group M₃] [semimodule R M] [semimodule R M₂] [semimodule R M₃] (f : M →ₗ[R] M₂) (g h : M₂ →ₗ[R] M₃) :
(g - h).comp f = g.comp f - h.comp f

theorem linear_map.comp_sub {R : Type u} {M : Type v} {M₂ : Type w} {M₃ : Type y} [semiring R] [add_comm_group M] [add_comm_group M₂] [add_comm_group M₃] [semimodule R M] [semimodule R M₂] [semimodule R M₃] (f g : M →ₗ[R] M₂) (h : M₂ →ₗ[R] M₃) :
h.comp (g - f) = h.comp g - h.comp f

@[instance]
def linear_map.has_scalar {R : Type u} {M : Type v} {M₂ : Type w} [comm_semiring R] [add_comm_monoid M] [add_comm_monoid M₂] [semimodule R M] [semimodule R M₂] :
has_scalar R (M →ₗ[R] M₂)

Equations
@[simp]
theorem linear_map.smul_apply {R : Type u} {M : Type v} {M₂ : Type w} [comm_semiring R] [add_comm_monoid M] [add_comm_monoid M₂] [semimodule R M] [semimodule R M₂] (f : M →ₗ[R] M₂) (a : R) (x : M) :
(a f) x = a f x

@[instance]
def linear_map.semimodule {R : Type u} {M : Type v} {M₂ : Type w} [comm_semiring R] [add_comm_monoid M] [add_comm_monoid M₂] [semimodule R M] [semimodule R M₂] :
semimodule R (M →ₗ[R] M₂)

Equations
def linear_map.comp_right {R : Type u} {M : Type v} {M₂ : Type w} {M₃ : Type y} [comm_semiring R] [add_comm_monoid M] [add_comm_monoid M₂] [add_comm_monoid M₃] [semimodule R M] [semimodule R M₂] [semimodule R M₃] (f : M₂ →ₗ[R] M₃) :
(M →ₗ[R] M₂) →ₗ[R] M →ₗ[R] M₃

Composition by f : M₂ → M₃ is a linear map from the space of linear maps M → M₂ to the space of linear maps M₂ → M₃.

Equations
theorem linear_map.smul_comp {R : Type u} {M : Type v} {M₂ : Type w} {M₃ : Type y} [comm_semiring R] [add_comm_monoid M] [add_comm_monoid M₂] [add_comm_monoid M₃] [semimodule R M] [semimodule R M₂] [semimodule R M₃] (f : M →ₗ[R] M₂) (g : M₂ →ₗ[R] M₃) (a : R) :
(a g).comp f = a g.comp f

theorem linear_map.comp_smul {R : Type u} {M : Type v} {M₂ : Type w} {M₃ : Type y} [comm_semiring R] [add_comm_monoid M] [add_comm_monoid M₂] [add_comm_monoid M₃] [semimodule R M] [semimodule R M₂] [semimodule R M₃] (f : M →ₗ[R] M₂) (g : M₂ →ₗ[R] M₃) (a : R) :
g.comp (a f) = a g.comp f

def linear_map.applyₗ {R : Type u} {M : Type v} {M₂ : Type w} [comm_semiring R] [add_comm_monoid M] [add_comm_monoid M₂] [semimodule R M] [semimodule R M₂] (v : M) :
(M →ₗ[R] M₂) →ₗ[R] M₂

Applying a linear map at v : M, seen as a linear map from M →ₗ[R] M₂ to M₂.

Equations
@[simp]
theorem linear_map.mul_apply {R : Type u} {M : Type v} [ring R] [add_comm_group M] [semimodule R M] (f g : M →ₗ[R] M) (x : M) :
(f * g) x = f (g x)

def linear_map.smul_rightₗ {R : Type u} {M : Type v} {M₂ : Type w} [comm_ring R] [add_comm_group M] [add_comm_group M₂] [semimodule R M] [semimodule R M₂] :
(M₂ →ₗ[R] R) →ₗ[R] M →ₗ[R] M₂ →ₗ[R] M

The family of linear maps M₂ → M parameterised by f ∈ M₂ → R, x ∈ M, is linear in f, x.

Equations
@[simp]
theorem linear_map.smul_rightₗ_apply {R : Type u} {M : Type v} {M₂ : Type w} [comm_ring R] [add_comm_group M] [add_comm_group M₂] [semimodule R M] [semimodule R M₂] (f : M₂ →ₗ[R] R) (x : M) (c : M₂) :

Properties of submodules

@[instance]
def submodule.partial_order {R : Type u} {M : Type v} [semiring R] [add_comm_monoid M] [semimodule R M] :

Equations
theorem submodule.le_def {R : Type u} {M : Type v} [semiring R] [add_comm_monoid M] [semimodule R M] {p p' : submodule R M} :
p p' p p'

theorem submodule.le_def' {R : Type u} {M : Type v} [semiring R] [add_comm_monoid M] [semimodule R M] {p p' : submodule R M} :
p p' ∀ (x : M), x px p'

theorem submodule.lt_def {R : Type u} {M : Type v} [semiring R] [add_comm_monoid M] [semimodule R M] {p p' : submodule R M} :
p < p' p p'

theorem submodule.not_le_iff_exists {R : Type u} {M : Type v} [semiring R] [add_comm_monoid M] [semimodule R M] {p p' : submodule R M} :
¬p p' ∃ (x : M) (H : x p), x p'

theorem submodule.exists_of_lt {R : Type u} {M : Type v} [semiring R] [add_comm_monoid M] [semimodule R M] {p p' : submodule R M} (a : p < p') :
∃ (x : M) (H : x p'), x p

theorem submodule.lt_iff_le_and_exists {R : Type u} {M : Type v} [semiring R] [add_comm_monoid M] [semimodule R M] {p p' : submodule R M} :
p < p' p p' ∃ (x : M) (H : x p'), x p

def submodule.of_le {R : Type u} {M : Type v} [semiring R] [add_comm_monoid M] [semimodule R M] {p p' : submodule R M} (h : p p') :

If two submodules p and p' satisfy p ⊆ p', then of_le p p' is the linear map version of this inclusion.

Equations
@[simp]
theorem submodule.coe_of_le {R : Type u} {M : Type v} [semiring R] [add_comm_monoid M] [semimodule R M] {p p' : submodule R M} (h : p p') (x : p) :

theorem submodule.of_le_apply {R : Type u} {M : Type v} [semiring R] [add_comm_monoid M] [semimodule R M] {p p' : submodule R M} (h : p p') (x : p) :

theorem submodule.subtype_comp_of_le {R : Type u} {M : Type v} [semiring R] [add_comm_monoid M] [semimodule R M] (p q : submodule R M) (h : p q) :

@[instance]
def submodule.has_bot {R : Type u} {M : Type v} [semiring R] [add_comm_monoid M] [semimodule R M] :

The set {0} is the bottom element of the lattice of submodules.

Equations
@[instance]
def submodule.inhabited' {R : Type u} {M : Type v} [semiring R] [add_comm_monoid M] [semimodule R M] :

Equations
@[simp]
theorem submodule.bot_coe {R : Type u} {M : Type v} [semiring R] [add_comm_monoid M] [semimodule R M] :
= {0}

@[simp]
theorem submodule.mem_bot (R : Type u) {M : Type v} [semiring R] [add_comm_monoid M] [semimodule R M] {x : M} :
x x = 0

theorem submodule.nonzero_mem_of_bot_lt {R : Type u} {M : Type v} [semiring R] [add_comm_monoid M] [semimodule R M] {I : submodule R M} (bot_lt : < I) :
∃ (a : I), a 0

theorem submodule.eq_bot_iff {R : Type u} {M : Type v} [semiring R] [add_comm_monoid M] [semimodule R M] (p : submodule R M) :
p = ∀ (x : M), x px = 0

theorem submodule.ne_bot_iff {R : Type u} {M : Type v} [semiring R] [add_comm_monoid M] [semimodule R M] (p : submodule R M) :
p ∃ (x : M) (H : x p), x 0

@[instance]
def submodule.has_top {R : Type u} {M : Type v} [semiring R] [add_comm_monoid M] [semimodule R M] :

The universal set is the top element of the lattice of submodules.

Equations
@[simp]
theorem submodule.top_coe {R : Type u} {M : Type v} [semiring R] [add_comm_monoid M] [semimodule R M] :

@[simp]
theorem submodule.mem_top {R : Type u} {M : Type v} [semiring R] [add_comm_monoid M] [semimodule R M] {x : M} :

theorem submodule.eq_bot_of_zero_eq_one {R : Type u} {M : Type v} [semiring R] [add_comm_monoid M] [semimodule R M] (p : submodule R M) (zero_eq_one : 0 = 1) :
p =

@[instance]
def submodule.has_Inf {R : Type u} {M : Type v} [semiring R] [add_comm_monoid M] [semimodule R M] :

Equations
@[instance]
def submodule.has_inf {R : Type u} {M : Type v} [semiring R] [add_comm_monoid M] [semimodule R M] :

Equations
@[instance]
def submodule.complete_lattice {R : Type u} {M : Type v} [semiring R] [add_comm_monoid M] [semimodule R M] :

Equations
@[simp]
theorem submodule.add_eq_sup {R : Type u} {M : Type v} [semiring R] [add_comm_monoid M] [semimodule R M] (p q : submodule R M) :
p + q = p q

@[simp]
theorem submodule.zero_eq_bot {R : Type u} {M : Type v} [semiring R] [add_comm_monoid M] [semimodule R M] :
0 =

theorem submodule.eq_top_iff' {R : Type u} {M : Type v} [semiring R] [add_comm_monoid M] [semimodule R M] {p : submodule R M} :
p = ∀ (x : M), x p

theorem submodule.bot_ne_top {R : Type u} {M : Type v} [semiring R] [add_comm_monoid M] [semimodule R M] [nontrivial M] :

@[simp]
theorem submodule.inf_coe {R : Type u} {M : Type v} [semiring R] [add_comm_monoid M] [semimodule R M] (p p' : submodule R M) :

@[simp]
theorem submodule.mem_inf {R : Type u} {M : Type v} [semiring R] [add_comm_monoid M] [semimodule R M] {x : M} {p p' : submodule R M} :
x p p' x p x p'

@[simp]
theorem submodule.Inf_coe {R : Type u} {M : Type v} [semiring R] [add_comm_monoid M] [semimodule R M] (P : set (submodule R M)) :
(Inf P) = ⋂ (p : submodule R M) (H : p P), p

@[simp]
theorem submodule.infi_coe {R : Type u} {M : Type v} [semiring R] [add_comm_monoid M] [semimodule R M] {ι : Sort u_1} (p : ι → submodule R M) :
(⨅ (i : ι), p i) = ⋂ (i : ι), (p i)

@[simp]
theorem submodule.mem_Inf {R : Type u} {M : Type v} [semiring R] [add_comm_monoid M] [semimodule R M] {S : set (submodule R M)} {x : M} :
x Inf S ∀ (p : submodule R M), p Sx p

@[simp]
theorem submodule.mem_infi {R : Type u} {M : Type v} [semiring R] [add_comm_monoid M] [semimodule R M] {x : M} {ι : Sort u_1} (p : ι → submodule R M) :
(x ⨅ (i : ι), p i) ∀ (i : ι), x p i

theorem submodule.disjoint_def {R : Type u} {M : Type v} [semiring R] [add_comm_monoid M] [semimodule R M] {p p' : submodule R M} :
disjoint p p' ∀ (x : M), x px p'x = 0

theorem submodule.mem_right_iff_eq_zero_of_disjoint {R : Type u} {M : Type v} [semiring R] [add_comm_monoid M] [semimodule R M] {p p' : submodule R M} (h : disjoint p p') {x : p} :
x p' x = 0

theorem submodule.mem_left_iff_eq_zero_of_disjoint {R : Type u} {M : Type v} [semiring R] [add_comm_monoid M] [semimodule R M] {p p' : submodule R M} (h : disjoint p p') {x : p'} :
x p x = 0

def submodule.map {R : Type u} {M : Type v} {M₂ : Type w} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] [semimodule R M] [semimodule R M₂] (f : M →ₗ[R] M₂) (p : submodule R M) :
submodule R M₂

The pushforward of a submodule p ⊆ M by f : M → M₂

Equations
@[simp]
theorem submodule.map_coe {R : Type u} {M : Type v} {M₂ : Type w} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] [semimodule R M] [semimodule R M₂] (f : M →ₗ[R] M₂) (p : submodule R M) :

@[simp]
theorem submodule.mem_map {R : Type u} {M : Type v} {M₂ : Type w} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] [semimodule R M] [semimodule R M₂] {f : M →ₗ[R] M₂} {p : submodule R M} {x : M₂} :
x submodule.map f p ∃ (y : M), y p f y = x

theorem submodule.mem_map_of_mem {R : Type u} {M : Type v} {M₂ : Type w} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] [semimodule R M] [semimodule R M₂] {f : M →ₗ[R] M₂} {p : submodule R M} {r : M} (h : r p) :

theorem submodule.map_id {R : Type u} {M : Type v} [semiring R] [add_comm_monoid M] [semimodule R M] (p : submodule R M) :

theorem submodule.map_comp {R : Type u} {M : Type v} {M₂ : Type w} {M₃ : Type y} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] [add_comm_monoid M₃] [semimodule R M] [semimodule R M₂] [semimodule R M₃] (f : M →ₗ[R] M₂) (g : M₂ →ₗ[R] M₃) (p : submodule R M) :

theorem submodule.map_mono {R : Type u} {M : Type v} {M₂ : Type w} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] [semimodule R M] [semimodule R M₂] {f : M →ₗ[R] M₂} {p p' : submodule R M} (a : p p') :

@[simp]
theorem submodule.map_zero {R : Type u} {M : Type v} {M₂ : Type w} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] [semimodule R M] [semimodule R M₂] (p : submodule R M) :

def submodule.comap {R : Type u} {M : Type v} {M₂ : Type w} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] [semimodule R M] [semimodule R M₂] (f : M →ₗ[R] M₂) (p : submodule R M₂) :

The pullback of a submodule p ⊆ M₂ along f : M → M₂

Equations
@[simp]
theorem submodule.comap_coe {R : Type u} {M : Type v} {M₂ : Type w} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] [semimodule R M] [semimodule R M₂] (f : M →ₗ[R] M₂) (p : submodule R M₂) :

@[simp]
theorem submodule.mem_comap {R : Type u} {M : Type v} {M₂ : Type w} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] [semimodule R M] [semimodule R M₂] {x : M} {f : M →ₗ[R] M₂} {p : submodule R M₂} :

theorem submodule.comap_id {R : Type u} {M : Type v} [semiring R] [add_comm_monoid M] [semimodule R M] (p : submodule R M) :

theorem submodule.comap_comp {R : Type u} {M : Type v} {M₂ : Type w} {M₃ : Type y} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] [add_comm_monoid M₃] [semimodule R M] [semimodule R M₂] [semimodule R M₃] (f : M →ₗ[R] M₂) (g : M₂ →ₗ[R] M₃) (p : submodule R M₃) :

theorem submodule.comap_mono {R : Type u} {M : Type v} {M₂ : Type w} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] [semimodule R M] [semimodule R M₂] {f : M →ₗ[R] M₂} {q q' : submodule R M₂} (a : q q') :

theorem submodule.map_le_iff_le_comap {R : Type u} {M : Type v} {M₂ : Type w} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] [semimodule R M] [semimodule R M₂] {f : M →ₗ[R] M₂} {p : submodule R M} {q : submodule R M₂} :

theorem submodule.gc_map_comap {R : Type u} {M : Type v} {M₂ : Type w} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] [semimodule R M] [semimodule R M₂] (f : M →ₗ[R] M₂) :

@[simp]
theorem submodule.map_bot {R : Type u} {M : Type v} {M₂ : Type w} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] [semimodule R M] [semimodule R M₂] (f : M →ₗ[R] M₂) :

@[simp]
theorem submodule.map_sup {R : Type u} {M : Type v} {M₂ : Type w} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] [semimodule R M] [semimodule R M₂] (p p' : submodule R M) (f : M →ₗ[R] M₂) :

@[simp]
theorem submodule.map_supr {R : Type u} {M : Type v} {M₂ : Type w} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] [semimodule R M] [semimodule R M₂] {ι : Sort u_1} (f : M →ₗ[R] M₂) (p : ι → submodule R M) :
submodule.map f (⨆ (i : ι), p i) = ⨆ (i : ι), submodule.map f (p i)

@[simp]
theorem submodule.comap_top {R : Type u} {M : Type v} {M₂ : Type w} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] [semimodule R M] [semimodule R M₂] (f : M →ₗ[R] M₂) :

@[simp]
theorem submodule.comap_inf {R : Type u} {M : Type v} {M₂ : Type w} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] [semimodule R M] [semimodule R M₂] (q q' : submodule R M₂) (f : M →ₗ[R] M₂) :

@[simp]
theorem submodule.comap_infi {R : Type u} {M : Type v} {M₂ : Type w} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] [semimodule R M] [semimodule R M₂] {ι : Sort u_1} (f : M →ₗ[R] M₂) (p : ι → submodule R M₂) :
submodule.comap f (⨅ (i : ι), p i) = ⨅ (i : ι), submodule.comap f (p i)

@[simp]
theorem submodule.comap_zero {R : Type u} {M : Type v} {M₂ : Type w} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] [semimodule R M] [semimodule R M₂] (q : submodule R M₂) :

theorem submodule.map_comap_le {R : Type u} {M : Type v} {M₂ : Type w} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] [semimodule R M] [semimodule R M₂] (f : M →ₗ[R] M₂) (q : submodule R M₂) :

theorem submodule.le_comap_map {R : Type u} {M : Type v} {M₂ : Type w} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] [semimodule R M] [semimodule R M₂] (f : M →ₗ[R] M₂) (p : submodule R M) :

theorem submodule.map_inf_eq_map_inf_comap {R : Type u} {M : Type v} {M₂ : Type w} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] [semimodule R M] [semimodule R M₂] {f : M →ₗ[R] M₂} {p : submodule R M} {p' : submodule R M₂} :

theorem submodule.map_comap_subtype {R : Type u} {M : Type v} [semiring R] [add_comm_monoid M] [semimodule R M] (p p' : submodule R M) :

theorem submodule.eq_zero_of_bot_submodule {R : Type u} {M : Type v} [semiring R] [add_comm_monoid M] [semimodule R M] (b : ) :
b = 0

def submodule.span (R : Type u) {M : Type v} [semiring R] [add_comm_monoid M] [semimodule R M] (s : set M) :

The span of a set s ⊆ M is the smallest submodule of M that contains s.

Equations
theorem submodule.mem_span {R : Type u} {M : Type v} [semiring R] [add_comm_monoid M] [semimodule R M] {x : M} {s : set M} :
x submodule.span R s ∀ (p : submodule R M), s px p

theorem submodule.subset_span {R : Type u} {M : Type v} [semiring R] [add_comm_monoid M] [semimodule R M] {s : set M} :

theorem submodule.span_le {R : Type u} {M : Type v} [semiring R] [add_comm_monoid M] [semimodule R M] {s : set M} {p : submodule R M} :

theorem submodule.span_mono {R : Type u} {M : Type v} [semiring R] [add_comm_monoid M] [semimodule R M] {s t : set M} (h : s t) :

theorem submodule.span_eq_of_le {R : Type u} {M : Type v} [semiring R] [add_comm_monoid M] [semimodule R M] (p : submodule R M) {s : set M} (h₁ : s p) (h₂ : p submodule.span R s) :

@[simp]
theorem submodule.span_eq {R : Type u} {M : Type v} [semiring R] [add_comm_monoid M] [semimodule R M] (p : submodule R M) :

theorem submodule.map_span {R : Type u} {M : Type v} {M₂ : Type w} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] [semimodule R M] [semimodule R M₂] (f : M →ₗ[R] M₂) (s : set M) :

theorem submodule.span_induction {R : Type u} {M : Type v} [semiring R] [add_comm_monoid M] [semimodule R M] {x : M} {s : set M} {p : M → Prop} (h : x submodule.span R s) (Hs : ∀ (x : M), x sp x) (H0 : p 0) (H1 : ∀ (x y : M), p xp yp (x + y)) (H2 : ∀ (a : R) (x : M), p xp (a x)) :
p x

An induction principle for span membership. If p holds for 0 and all elements of s, and is preserved under addition and scalar multiplication, then p holds for all elements of the span of s.

def submodule.gi (R : Type u) (M : Type v) [semiring R] [add_comm_monoid M] [semimodule R M] :

span forms a Galois insertion with the coercion from submodule to set.

Equations
@[simp]
theorem submodule.span_empty {R : Type u} {M : Type v} [semiring R] [add_comm_monoid M] [semimodule R M] :

@[simp]
theorem submodule.span_univ {R : Type u} {M : Type v} [semiring R] [add_comm_monoid M] [semimodule R M] :

theorem submodule.span_union {R : Type u} {M : Type v} [semiring R] [add_comm_monoid M] [semimodule R M] (s t : set M) :

theorem submodule.span_Union {R : Type u} {M : Type v} [semiring R] [add_comm_monoid M] [semimodule R M] {ι : Sort u_1} (s : ι → set M) :
submodule.span R (⋃ (i : ι), s i) = ⨆ (i : ι), submodule.span R (s i)

@[simp]
theorem submodule.coe_supr_of_directed {R : Type u} {M : Type v} [semiring R] [add_comm_monoid M] [semimodule R M] {ι : Sort u_1} [hι : nonempty ι] (S : ι → submodule R M) (H : directed has_le.le S) :
(supr S) = ⋃ (i : ι), (S i)

theorem submodule.mem_sup_left {R : Type u} {M : Type v} [semiring R] [add_comm_monoid M] [semimodule R M] {S T : submodule R M} {x : M} (a : x S) :
x S T

theorem submodule.mem_sup_right {R : Type u} {M : Type v} [semiring R] [add_comm_monoid M] [semimodule R M] {S T : submodule R M} {x : M} (a : x T) :
x S T

theorem submodule.mem_supr_of_mem {R : Type u} {M : Type v} [semiring R] [add_comm_monoid M] [semimodule R M] {ι : Sort u_1} {b : M} {p : ι → submodule R M} (i : ι) (h : b p i) :
b ⨆ (i : ι), p i

theorem submodule.mem_Sup_of_mem {R : Type u} {M : Type v} [semiring R] [add_comm_monoid M] [semimodule R M] {S : set (submodule R M)} {s : submodule R M} (hs : s S) {x : M} (a : x s) :
x Sup S

@[simp]
theorem submodule.mem_supr_of_directed {R : Type u} {M : Type v} [semiring R] [add_comm_monoid M] [semimodule R M] {ι : Sort u_1} [nonempty ι] (S : ι → submodule R M) (H : directed has_le.le S) {x : M} :
x supr S ∃ (i : ι), x S i

theorem submodule.mem_Sup_of_directed {R : Type u} {M : Type v} [semiring R] [add_comm_monoid M] [semimodule R M] {s : set (submodule R M)} {z : M} (hs : s.nonempty) (hdir : directed_on has_le.le s) :
z Sup s ∃ (y : submodule R M) (H : y s), z y

theorem submodule.mem_sup {R : Type u} {M : Type v} [semiring R] [add_comm_monoid M] [semimodule R M] {p p' : submodule R M} {x : M} :
x p p' ∃ (y : M) (H : y p) (z : M) (H : z p'), y + z = x

theorem submodule.mem_sup' {R : Type u} {M : Type v} [semiring R] [add_comm_monoid M] [semimodule R M] {p p' : submodule R M} {x : M} :
x p p' ∃ (y : p) (z : p'), y + z = x

theorem submodule.mem_span_singleton_self {R : Type u} {M : Type v} [semiring R] [add_comm_monoid M] [semimodule R M] (x : M) :

theorem submodule.nontrivial_span_singleton {R : Type u} {M : Type v} [semiring R] [add_comm_monoid M] [semimodule R M] {x : M} (h : x 0) :

theorem submodule.mem_span_singleton {R : Type u} {M : Type v} [semiring R] [add_comm_monoid M] [semimodule R M] {x y : M} :
x submodule.span R {y} ∃ (a : R), a y = x

theorem submodule.span_singleton_eq_range {R : Type u} {M : Type v} [semiring R] [add_comm_monoid M] [semimodule R M] (y : M) :
(submodule.span R {y}) = set.range (λ (_x : R), _x y)

theorem submodule.disjoint_span_singleton {K : Type u_1} {E : Type u_2} [division_ring K] [add_comm_group E] [module K E] {s : submodule K E} {x : E} :
disjoint s (submodule.span K {x}) x sx = 0

theorem submodule.mem_span_insert {R : Type u} {M : Type v} [semiring R] [add_comm_monoid M] [semimodule R M] {x : M} {s : set M} {y : M} :
x submodule.span R (insert y s) ∃ (a : R) (z : M) (H : z submodule.span R s), x = a y + z

theorem submodule.span_insert_eq_span {R : Type u} {M : Type v} [semiring R] [add_comm_monoid M] [semimodule R M] {x : M} {s : set M} (h : x submodule.span R s) :

theorem submodule.span_span {R : Type u} {M : Type v} [semiring R] [add_comm_monoid M] [semimodule R M] {s : set M} :

theorem submodule.span_eq_bot {R : Type u} {M : Type v} [semiring R] [add_comm_monoid M] [semimodule R M] {s : set M} :
submodule.span R s = ∀ (x : M), x sx = 0

@[simp]
theorem submodule.span_singleton_eq_bot {R : Type u} {M : Type v} [semiring R] [add_comm_monoid M] [semimodule R M] {x : M} :

@[simp]
theorem submodule.span_zero {R : Type u} {M : Type v} [semiring R] [add_comm_monoid M] [semimodule R M] :

@[simp]
theorem submodule.span_image {R : Type u} {M : Type v} {M₂ : Type w} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] [semimodule R M] [semimodule R M₂] {s : set M} (f : M →ₗ[R] M₂) :

theorem submodule.linear_eq_on {R : Type u} {M : Type v} {M₂ : Type w} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] [semimodule R M] [semimodule R M₂] (s : set M) {f g : M →ₗ[R] M₂} (H : ∀ (x : M), x sf x = g x) {x : M} (h : x submodule.span R s) :
f x = g x

theorem submodule.linear_map.ext_on {R : Type u} {M : Type v} {M₂ : Type w} {ι : Type x} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] [semimodule R M] [semimodule R M₂] {v : ι → M} {f g : M →ₗ[R] M₂} (hv : submodule.span R (set.range v) = ) (h : ∀ (i : ι), f (v i) = g (v i)) :
f = g

theorem submodule.supr_eq_span {R : Type u} {M : Type v} [semiring R] [add_comm_monoid M] [semimodule R M] {ι : Sort w} (p : ι → submodule R M) :
(⨆ (i : ι), p i) = submodule.span R (⋃ (i : ι), (p i))

theorem submodule.span_singleton_le_iff_mem {R : Type u} {M : Type v} [semiring R] [add_comm_monoid M] [semimodule R M] (m : M) (p : submodule R M) :

theorem submodule.lt_add_iff_not_mem {R : Type u} {M : Type v} [semiring R] [add_comm_monoid M] [semimodule R M] {I : submodule R M} {a : M} :
I < I + submodule.span R {a} a I

theorem submodule.mem_supr {R : Type u} {M : Type v} [semiring R] [add_comm_monoid M] [semimodule R M] {ι : Sort w} (p : ι → submodule R M) {m : M} :
(m ⨆ (i : ι), p i) ∀ (N : submodule R M), (∀ (i : ι), p i N)m N

def submodule.prod {R : Type u} {M : Type v} {M₂ : Type w} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] [semimodule R M] [semimodule R M₂] (p : submodule R M) (q : submodule R M₂) :
submodule R (M × M₂)

The product of two submodules is a submodule.

Equations
@[simp]
theorem submodule.prod_coe {R : Type u} {M : Type v} {M₂ : Type w} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] [semimodule R M] [semimodule R M₂] (p : submodule R M) (q : submodule R M₂) :
(p.prod q) = p.prod q

@[simp]
theorem submodule.mem_prod {R : Type u} {M : Type v} {M₂ : Type w} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] [semimodule R M] [semimodule R M₂] {p : submodule R M} {q : submodule R M₂} {x : M × M₂} :
x p.prod q x.fst p x.snd q

theorem submodule.span_prod_le {R : Type u} {M : Type v} {M₂ : Type w} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] [semimodule R M] [semimodule R M₂] (s : set M) (t : set M₂) :

@[simp]
theorem submodule.prod_top {R : Type u} {M : Type v} {M₂ : Type w} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] [semimodule R M] [semimodule R M₂] :

@[simp]
theorem submodule.prod_bot {R : Type u} {M : Type v} {M₂ : Type w} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] [semimodule R M] [semimodule R M₂] :

theorem submodule.prod_mono {R : Type u} {M : Type v} {M₂ : Type w} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] [semimodule R M] [semimodule R M₂] {p p' : submodule R M} {q q' : submodule R M₂} (a : p p') (a_1 : q q') :
p.prod q p'.prod q'

@[simp]
theorem submodule.prod_inf_prod {R : Type u} {M : Type v} {M₂ : Type w} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] [semimodule R M] [semimodule R M₂] (p p' : submodule R M) (q q' : submodule R M₂) :
p.prod q p'.prod q' = (p p').prod (q q')

@[simp]
theorem submodule.prod_sup_prod {R : Type u} {M : Type v} {M₂ : Type w} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] [semimodule R M] [semimodule R M₂] (p p' : submodule R M) (q q' : submodule R M₂) :
p.prod q p'.prod q' = (p p').prod (q q')

theorem submodule.mem_span_insert' {R : Type u} {M : Type v} [ring R] [add_comm_group M] [semimodule R M] {x y : M} {s : set M} :
x submodule.span R (insert y s) ∃ (a : R), x + a y submodule.span R s

def submodule.quotient_rel {R : Type u} {M : Type v} [ring R] [add_comm_group M] [semimodule R M] (p : submodule R M) :

The equivalence relation associated to a submodule p, defined by x ≈ y iff y - x ∈ p.

Equations
def submodule.quotient {R : Type u} {M : Type v} [ring R] [add_comm_group M] [semimodule R M] (p : submodule R M) :
Type v

The quotient of a module M by a submodule p ⊆ M.

Equations
def submodule.quotient.mk {R : Type u} {M : Type v} [ring R] [add_comm_group M] [semimodule R M] {p : submodule R M} (a : M) :

Map associating to an element of M the corresponding element of M/p, when p is a submodule of M.

Equations
@[simp]
theorem submodule.quotient.mk_eq_mk {R : Type u} {M : Type v} [ring R] [add_comm_group M] [semimodule R M] {p : submodule R M} (x : M) :

@[simp]
theorem submodule.quotient.mk'_eq_mk {R : Type u} {M : Type v} [ring R] [add_comm_group M] [semimodule R M] {p : submodule R M} (x : M) :

@[simp]
theorem submodule.quotient.quot_mk_eq_mk {R : Type u} {M : Type v} [ring R] [add_comm_group M] [semimodule R M] {p : submodule R M} (x : M) :

theorem submodule.quotient.eq {R : Type u} {M : Type v} [ring R] [add_comm_group M] [semimodule R M] (p : submodule R M) {x y : M} :

@[instance]
def submodule.quotient.has_zero {R : Type u} {M : Type v} [ring R] [add_comm_group M] [semimodule R M] (p : submodule R M) :

Equations
@[instance]
def submodule.quotient.inhabited {R : Type u} {M : Type v} [ring R] [add_comm_group M] [semimodule R M] (p : submodule R M) :

Equations
@[simp]
theorem submodule.quotient.mk_zero {R : Type u} {M : Type v} [ring R] [add_comm_group M] [semimodule R M] (p : submodule R M) :

@[simp]
theorem submodule.quotient.mk_eq_zero {R : Type u} {M : Type v} [ring R] [add_comm_group M] [semimodule R M] (p : submodule R M) {x : M} :

@[instance]
def submodule.quotient.has_add {R : Type u} {M : Type v} [ring R] [add_comm_group M] [semimodule R M] (p : submodule R M) :

Equations
@[simp]
theorem submodule.quotient.mk_add {R : Type u} {M : Type v} [ring R] [add_comm_group M] [semimodule R M] (p : submodule R M) {x y : M} :

@[instance]
def submodule.quotient.has_neg {R : Type u} {M : Type v} [ring R] [add_comm_group M] [semimodule R M] (p : submodule R M) :

Equations
@[simp]
theorem submodule.quotient.mk_neg {R : Type u} {M : Type v} [ring R] [add_comm_group M] [semimodule R M] (p : submodule R M) {x : M} :

@[instance]
def submodule.quotient.has_scalar {R : Type u} {M : Type v} [ring R] [add_comm_group M] [semimodule R M] (p : submodule R M) :

Equations
@[simp]
theorem submodule.quotient.mk_smul {R : Type u} {M : Type v} [ring R] [add_comm_group M] [semimodule R M] (p : submodule R M) {r : R} {x : M} :

theorem submodule.quotient.nontrivial_of_lt_top {R : Type u} {M : Type v} [ring R] [add_comm_group M] [semimodule R M] (p : submodule R M) (h : p < ) :

theorem submodule.quot_hom_ext {R : Type u} {M : Type v} {M₂ : Type w} [ring R] [add_comm_group M] [add_comm_group M₂] [semimodule R M] [semimodule R M₂] (p : submodule R M) ⦃f g : p.quotient →ₗ[R] M₂⦄ (h : ∀ (x : M), f (submodule.quotient.mk x) = g (submodule.quotient.mk x)) :
f = g

theorem submodule.comap_smul {K : Type u'} {V : Type v'} {V₂ : Type w'} [field K] [add_comm_group V] [vector_space K V] [add_comm_group V₂] [vector_space K V₂] (f : V →ₗ[K] V₂) (p : submodule K V₂) (a : K) (h : a 0) :

theorem submodule.map_smul {K : Type u'} {V : Type v'} {V₂ : Type w'} [field K] [add_comm_group V] [vector_space K V] [add_comm_group V₂] [vector_space K V₂] (f : V →ₗ[K] V₂) (p : submodule K V) (a : K) (h : a 0) :

theorem submodule.comap_smul' {K : Type u'} {V : Type v'} {V₂ : Type w'} [field K] [add_comm_group V] [vector_space K V] [add_comm_group V₂] [vector_space K V₂] (f : V →ₗ[K] V₂) (p : submodule K V₂) (a : K) :
submodule.comap (a f) p = ⨅ (h : a 0), submodule.comap f p

theorem submodule.map_smul' {K : Type u'} {V : Type v'} {V₂ : Type w'} [field K] [add_comm_group V] [vector_space K V] [add_comm_group V₂] [vector_space K V₂] (f : V →ₗ[K] V₂) (p : submodule K V) (a : K) :
submodule.map (a f) p = ⨆ (h : a 0), submodule.map f p

Properties of linear maps

@[simp]
theorem linear_map.finsupp_sum {R : Type u} {M : Type v} {M₂ : Type w} {ι : Type x} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] [semimodule R M] [semimodule R M₂] {γ : Type u_1} [has_zero γ] (f : M →ₗ[R] M₂) {t : ι →₀ γ} {g : ι → γ → M} :
f (t.sum g) = t.sum (λ (i : ι) (d : γ), f (g i d))

theorem linear_map.map_cod_restrict {R : Type u} {M : Type v} {M₂ : Type w} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] [semimodule R M] [semimodule R M₂] (p : submodule R M) (f : M₂ →ₗ[R] M) (h : ∀ (c : M₂), f c p) (p' : submodule R M₂) :

theorem linear_map.comap_cod_restrict {R : Type u} {M : Type v} {M₂ : Type w} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] [semimodule R M] [semimodule R M₂] (p : submodule R M) (f : M₂ →ₗ[R] M) (hf : ∀ (c : M₂), f c p) (p' : submodule R p) :

def linear_map.range {R : Type u} {M : Type v} {M₂ : Type w} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] [semimodule R M] [semimodule R M₂] (f : M →ₗ[R] M₂) :
submodule R M₂

The range of a linear map f : M → M₂ is a submodule of M₂.

Equations
theorem linear_map.range_coe {R : Type u} {M : Type v} {M₂ : Type w} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] [semimodule R M] [semimodule R M₂] (f : M →ₗ[R] M₂) :

@[simp]
theorem linear_map.mem_range {R : Type u} {M : Type v} {M₂ : Type w} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] [semimodule R M] [semimodule R M₂] {f : M →ₗ[R] M₂} {x : M₂} :
x f.range ∃ (y : M), f y = x

theorem linear_map.mem_range_self {R : Type u} {M : Type v} {M₂ : Type w} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] [semimodule R M] [semimodule R M₂] (f : M →ₗ[R] M₂) (x : M) :

@[simp]
theorem linear_map.range_id {R : Type u} {M : Type v} [semiring R] [add_comm_monoid M] [semimodule R M] :

theorem linear_map.range_comp {R : Type u} {M : Type v} {M₂ : Type w} {M₃ : Type y} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] [add_comm_monoid M₃] [semimodule R M] [semimodule R M₂] [semimodule R M₃] (f : M →ₗ[R] M₂) (g : M₂ →ₗ[R] M₃) :

theorem linear_map.range_comp_le_range {R : Type u} {M : Type v} {M₂ : Type w} {M₃ : Type y} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] [add_comm_monoid M₃] [semimodule R M] [semimodule R M₂] [semimodule R M₃] (f : M →ₗ[R] M₂) (g : M₂ →ₗ[R] M₃) :

theorem linear_map.range_eq_top {R : Type u} {M : Type v} {M₂ : Type w} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] [semimodule R M] [semimodule R M₂] {f : M →ₗ[R] M₂} :

theorem linear_map.range_le_iff_comap {R : Type u} {M : Type v} {M₂ : Type w} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] [semimodule R M] [semimodule R M₂] {f : M →ₗ[R] M₂} {p : submodule R M₂} :

theorem linear_map.map_le_range {R : Type u} {M : Type v} {M₂ : Type w} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] [semimodule R M] [semimodule R M₂] {f : M →ₗ[R] M₂} {p : submodule R M} :

theorem linear_map.range_coprod {R : Type u} {M : Type v} {M₂ : Type w} {M₃ : Type y} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] [add_comm_monoid M₃] [semimodule R M] [semimodule R M₂] [semimodule R M₃] (f : M →ₗ[R] M₃) (g : M₂ →ₗ[R] M₃) :

theorem linear_map.sup_range_inl_inr {R : Type u} {M : Type v} {M₂ : Type w} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] [semimodule R M] [semimodule R M₂] :

def linear_map.range_restrict {R : Type u} {M : Type v} {M₂ : Type w} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] [semimodule R M] [semimodule R M₂] (f : M →ₗ[R] M₂) :

Restrict the codomain of a linear map f to f.range.

Equations
def linear_map.to_span_singleton (R : Type u) (M : Type v) [semiring R] [add_comm_monoid M] [semimodule R M] (x : M) :

Given an element x of a module M over R, the natural map from R to scalar multiples of x.

Equations
theorem linear_map.span_singleton_eq_range (R : Type u) (M : Type v) [semiring R] [add_comm_monoid M] [semimodule R M] (x : M) :

The range of to_span_singleton x is the span of x.

theorem linear_map.to_span_singleton_one (R : Type u) (M : Type v) [semiring R] [add_comm_monoid M] [semimodule R M] (x : M) :

def linear_map.ker {R : Type u} {M : Type v} {M₂ : Type w} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] [semimodule R M] [semimodule R M₂] (f : M →ₗ[R] M₂) :

The kernel of a linear map f : M → M₂ is defined to be comap f ⊥. This is equivalent to the set of x : M such that f x = 0. The kernel is a submodule of M.

Equations
@[simp]
theorem linear_map.mem_ker {R : Type u} {M : Type v} {M₂ : Type w} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] [semimodule R M] [semimodule R M₂] {f : M →ₗ[R] M₂} {y : M} :
y f.ker f y = 0

@[simp]
theorem linear_map.ker_id {R : Type u} {M : Type v} [semiring R] [add_comm_monoid M] [semimodule R M] :

@[simp]
theorem linear_map.map_coe_ker {R : Type u} {M : Type v} {M₂ : Type w} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] [semimodule R M] [semimodule R M₂] (f : M →ₗ[R] M₂) (x : (f.ker)) :
f x = 0

theorem linear_map.comp_ker_subtype {R : Type u} {M : Type v} {M₂ : Type w} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] [semimodule R M] [semimodule R M₂] (f : M →ₗ[R] M₂) :

theorem linear_map.ker_comp {R : Type u} {M : Type v} {M₂ : Type w} {M₃ : Type y} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] [add_comm_monoid M₃] [semimodule R M] [semimodule R M₂] [semimodule R M₃] (f : M →ₗ[R] M₂) (g : M₂ →ₗ[R] M₃) :

theorem linear_map.ker_le_ker_comp {R : Type u} {M : Type v} {M₂ : Type w} {M₃ : Type y} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] [add_comm_monoid M₃] [semimodule R M] [semimodule R M₂] [semimodule R M₃] (f : M →ₗ[R] M₂) (g : M₂ →ₗ[R] M₃) :
f.ker (g.comp f).ker

theorem linear_map.disjoint_ker {R : Type u} {M : Type v} {M₂ : Type w} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] [semimodule R M] [semimodule R M₂] {f : M →ₗ[R] M₂} {p : submodule R M} :
disjoint p f.ker ∀ (x : M), x pf x = 0x = 0

theorem linear_map.disjoint_inl_inr {R : Type u} {M : Type v} {M₂ : Type w} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] [semimodule R M] [semimodule R M₂] :

theorem linear_map.ker_eq_bot' {R : Type u} {M : Type v} {M₂ : Type w} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] [semimodule R M] [semimodule R M₂] {f : M →ₗ[R] M₂} :
f.ker = ∀ (m : M), f m = 0m = 0

theorem linear_map.le_ker_iff_map {R : Type u} {M : Type v} {M₂ : Type w} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] [semimodule R M] [semimodule R M₂] {f : M →ₗ[R] M₂} {p : submodule R M} :

theorem linear_map.ker_cod_restrict {R : Type u} {M : Type v} {M₂ : Type w} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] [semimodule R M] [semimodule R M₂] (p : submodule R M) (f : M₂ →ₗ[R] M) (hf : ∀ (c : M₂), f c p) :

theorem linear_map.range_cod_restrict {R : Type u} {M : Type v} {M₂ : Type w} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] [semimodule R M] [semimodule R M₂] (p : submodule R M) (f : M₂ →ₗ[R] M) (hf : ∀ (c : M₂), f c p) :

theorem linear_map.ker_restrict {R : Type u} {M : Type v} [semiring R] [add_comm_monoid M] [semimodule R M] {p : submodule R M} {f : M →ₗ[R] M} (hf : ∀ (x : M), x pf x p) :

theorem linear_map.map_comap_eq {R : Type u} {M : Type v} {M₂ : Type w} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] [semimodule R M] [semimodule R M₂] (f : M →ₗ[R] M₂) (q : submodule R M₂) :

theorem linear_map.map_comap_eq_self {R : Type u} {M : Type v} {M₂ : Type w} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] [semimodule R M] [semimodule R M₂] {f : M →ₗ[R] M₂} {q : submodule R M₂} (h : q f.range) :

@[simp]
theorem linear_map.ker_zero {R : Type u} {M : Type v} {M₂ : Type w} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] [semimodule R M] [semimodule R M₂] :

@[simp]
theorem linear_map.range_zero {R : Type u} {M : Type v} {M₂ : Type w} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] [semimodule R M] [semimodule R M₂] :

theorem linear_map.ker_eq_top {R : Type u} {M : Type v} {M₂ : Type w} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] [semimodule R M] [semimodule R M₂] {f : M →ₗ[R] M₂} :
f.ker = f = 0

theorem linear_map.range_le_bot_iff {R : Type u} {M : Type v} {M₂ : Type w} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] [semimodule R M] [semimodule R M₂] (f : M →ₗ[R] M₂) :
f.range f = 0

theorem linear_map.range_le_ker_iff {R : Type u} {M : Type v} {M₂ : Type w} {M₃ : Type y} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] [add_comm_monoid M₃] [semimodule R M] [semimodule R M₂] [semimodule R M₃] {f : M →ₗ[R] M₂} {g : M₂ →ₗ[R] M₃} :
f.range g.ker g.comp f = 0

theorem linear_map.comap_le_comap_iff {R : Type u} {M : Type v} {M₂ : Type w} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] [semimodule R M] [semimodule R M₂] {f : M →ₗ[R] M₂} (hf : f.range = ) {p p' : submodule R M₂} :

theorem linear_map.comap_injective {R : Type u} {M : Type v} {M₂ : Type w} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] [semimodule R M] [semimodule R M₂] {f : M →ₗ[R] M₂} (hf : f.range = ) :

theorem linear_map.map_coprod_prod {R : Type u} {M : Type v} {M₂ : Type w} {M₃ : Type y} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] [add_comm_monoid M₃] [semimodule R M] [semimodule R M₂] [semimodule R M₃] (f : M →ₗ[R] M₃) (g : M₂ →ₗ[R] M₃) (p : submodule R M) (q : submodule R M₂) :

theorem linear_map.comap_prod_prod {R : Type u} {M : Type v} {M₂ : Type w} {M₃ : Type y} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] [add_comm_monoid M₃] [semimodule R M] [semimodule R M₂] [semimodule R M₃] (f : M →ₗ[R] M₂) (g : M →ₗ[R] M₃) (p : submodule R M₂) (q : submodule R M₃) :

theorem linear_map.prod_eq_inf_comap {R : Type u} {M : Type v} {M₂ : Type w} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] [semimodule R M] [semimodule R M₂] (p : submodule R M) (q : submodule R M₂) :

theorem linear_map.prod_eq_sup_map {R : Type u} {M : Type v} {M₂ : Type w} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] [semimodule R M] [semimodule R M₂] (p : submodule R M) (q : submodule R M₂) :

theorem linear_map.span_inl_union_inr {R : Type u} {M : Type v} {M₂ : Type w} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] [semimodule R M] [semimodule R M₂] {s : set M} {t : set M₂} :

@[simp]
theorem linear_map.ker_prod {R : Type u} {M : Type v} {M₂ : Type w} {M₃ : Type y} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] [add_comm_monoid M₃] [semimodule R M] [semimodule R M₂] [semimodule R M₃] (f : M →ₗ[R] M₂) (g : M →ₗ[R] M₃) :
(f.prod g).ker = f.ker g.ker

theorem linear_map.range_prod_le {R : Type u} {M : Type v} {M₂ : Type w} {M₃ : Type y} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] [add_comm_monoid M₃] [semimodule R M] [semimodule R M₂] [semimodule R M₃] (f : M →ₗ[R] M₂) (g : M →ₗ[R] M₃) :

theorem linear_map.ker_eq_bot_of_injective {R : Type u} {M : Type v} {M₂ : Type w} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] [semimodule R M] [semimodule R M₂] {f : M →ₗ[R] M₂} (hf : function.injective f) :

theorem linear_map.comap_map_eq {R : Type u} {M : Type v} {M₂ : Type w} [semiring R] [add_comm_group M] [add_comm_group M₂] [semimodule R M] [semimodule R M₂] (f : M →ₗ[R] M₂) (p : submodule R M) :

theorem linear_map.comap_map_eq_self {R : Type u} {M : Type v} {M₂ : Type w} [semiring R] [add_comm_group M] [add_comm_group M₂] [semimodule R M] [semimodule R M₂] {f : M →ₗ[R] M₂} {p : submodule R M} (h : f.ker p) :

theorem linear_map.map_le_map_iff {R : Type u} {M : Type v} {M₂ : Type w} [semiring R] [add_comm_group M] [add_comm_group M₂] [semimodule R M] [semimodule R M₂] (f : M →ₗ[R] M₂) {p p' : submodule R M} :

theorem linear_map.map_le_map_iff' {R : Type u} {M : Type v} {M₂ : Type w} [semiring R] [add_comm_group M] [add_comm_group M₂] [semimodule R M] [semimodule R M₂] {f : M →ₗ[R] M₂} (hf : f.ker = ) {p p' : submodule R M} :

theorem linear_map.map_injective {R : Type u} {M : Type v} {M₂ : Type w} [semiring R] [add_comm_group M] [add_comm_group M₂] [semimodule R M] [semimodule R M₂] {f : M →ₗ[R] M₂} (hf : f.ker = ) :

theorem linear_map.map_eq_top_iff {R : Type u} {M : Type v} {M₂ : Type w} [semiring R] [add_comm_group M] [add_comm_group M₂] [semimodule R M] [semimodule R M₂] {f : M →ₗ[R] M₂} (hf : f.range = ) {p : submodule R M} :

theorem linear_map.sub_mem_ker_iff {R : Type u} {M : Type v} {M₂ : Type w} [ring R] [add_comm_group M] [add_comm_group M₂] [semimodule R M] [semimodule R M₂] {f : M →ₗ[R] M₂} {x y : M} :
x - y f.ker f x = f y

theorem linear_map.disjoint_ker' {R : Type u} {M : Type v} {M₂ : Type w} [ring R] [add_comm_group M] [add_comm_group M₂] [semimodule R M] [semimodule R M₂] {f : M →ₗ[R] M₂} {p : submodule R M} :
disjoint p f.ker ∀ (x y : M), x py pf x = f yx = y

theorem linear_map.inj_of_disjoint_ker {R : Type u} {M : Type v} {M₂ : Type w} [ring R] [add_comm_group M] [add_comm_group M₂] [semimodule R M] [semimodule R M₂] {f : M →ₗ[R] M₂} {p : submodule R M} {s : set M} (h : s p) (hd : disjoint p f.ker) (x y : M) (H : x s) (H_1 : y s) (a : f x = f y) :
x = y

theorem linear_map.ker_eq_bot {R : Type u} {M : Type v} {M₂ : Type w} [ring R] [add_comm_group M] [add_comm_group M₂] [semimodule R M] [semimodule R M₂] {f : M →ₗ[R] M₂} :

theorem linear_map.range_prod_eq {R : Type u} {M : Type v} {M₂ : Type w} {M₃ : Type y} [ring R] [add_comm_group M] [add_comm_group M₂] [add_comm_group M₃] [semimodule R M] [semimodule R M₂] [semimodule R M₃] {f : M →ₗ[R] M₂} {g : M →ₗ[R] M₃} (h : f.ker g.ker = ) :

If the union of the kernels ker f and ker g spans the domain, then the range of prod f g is equal to the product of range f and range g.

theorem linear_map.ker_smul {K : Type u'} {V : Type v'} {V₂ : Type w'} [field K] [add_comm_group V] [vector_space K V] [add_comm_group V₂] [vector_space K V₂] (f : V →ₗ[K] V₂) (a : K) (h : a 0) :
(a f).ker = f.ker

theorem linear_map.ker_smul' {K : Type u'} {V : Type v'} {V₂ : Type w'} [field K] [add_comm_group V] [vector_space K V] [add_comm_group V₂] [vector_space K V₂] (f : V →ₗ[K] V₂) (a : K) :
(a f).ker = ⨅ (h : a 0), f.ker

theorem linear_map.range_smul {K : Type u'} {V : Type v'} {V₂ : Type w'} [field K] [add_comm_group V] [vector_space K V] [add_comm_group V₂] [vector_space K V₂] (f : V →ₗ[K] V₂) (a : K) (h : a 0) :
(a f).range = f.range

theorem linear_map.range_smul' {K : Type u'} {V : Type v'} {V₂ : Type w'} [field K] [add_comm_group V] [vector_space K V] [add_comm_group V₂] [vector_space K V₂] (f : V →ₗ[K] V₂) (a : K) :
(a f).range = ⨆ (h : a 0), f.range

theorem submodule.sup_eq_range {R : Type u} {M : Type v} [semiring R] [add_comm_monoid M] [semimodule R M] (p q : submodule R M) :

theorem is_linear_map.is_linear_map_add {R : Type u} {M : Type v} [semiring R] [add_comm_monoid M] [semimodule R M] :
is_linear_map R (λ (x : M × M), x.fst + x.snd)

theorem is_linear_map.is_linear_map_sub {R : Type u_1} {M : Type u_2} [semiring R] [add_comm_group M] [semimodule R M] :
is_linear_map R (λ (x : M × M), x.fst - x.snd)

@[simp]
theorem submodule.map_top {R : Type u} {M : Type v} {M₂ : Type w} {T : semiring R} [add_comm_monoid M] [add_comm_monoid M₂] [semimodule R M] [semimodule R M₂] (f : M →ₗ[R] M₂) :

@[simp]
theorem submodule.comap_bot {R : Type u} {M : Type v} {M₂ : Type w} {T : semiring R} [add_comm_monoid M] [add_comm_monoid M₂] [semimodule R M] [semimodule R M₂] (f : M →ₗ[R] M₂) :

@[simp]
theorem submodule.ker_subtype {R : Type u} {M : Type v} {T : semiring R} [add_comm_monoid M] [semimodule R M] (p : submodule R M) :

@[simp]
theorem submodule.range_subtype {R : Type u} {M : Type v} {T : semiring R} [add_comm_monoid M] [semimodule R M] (p : submodule R M) :

theorem submodule.map_subtype_le {R : Type u} {M : Type v} {T : semiring R} [add_comm_monoid M] [semimodule R M] (p : submodule R M) (p' : submodule R p) :

@[simp]
theorem submodule.map_subtype_top {R : Type u} {M : Type v} {T : semiring R} [add_comm_monoid M] [semimodule R M] (p : submodule R M) :

Under the canonical linear map from a submodule p to the ambient space M, the image of the maximal submodule of p is just p.

@[simp]
theorem submodule.comap_subtype_eq_top {R : Type u} {M : Type v} {T : semiring R} [add_comm_monoid M] [semimodule R M] {p p' : submodule R M} :

@[simp]
theorem submodule.comap_subtype_self {R : Type u} {M : Type v} {T : semiring R} [add_comm_monoid M] [semimodule R M] (p : submodule R M) :

@[simp]
theorem submodule.ker_of_le {R : Type u} {M : Type v} {T : semiring R} [add_comm_monoid M] [semimodule R M] (p p' : submodule R M) (h : p p') :

theorem submodule.range_of_le {R : Type u} {M : Type v} {T : semiring R} [add_comm_monoid M] [semimodule R M] (p q : submodule R M) (h : p q) :

@[simp]
theorem submodule.map_inl {R : Type u} {M : Type v} {M₂ : Type w} {T : semiring R} [add_comm_monoid M] [add_comm_monoid M₂] [semimodule R M] [semimodule R M₂] (p : submodule R M) :

@[simp]
theorem submodule.map_inr {R : Type u} {M : Type v} {M₂ : Type w} {T : semiring R} [add_comm_monoid M] [add_comm_monoid M₂] [semimodule R M] [semimodule R M₂] (q : submodule R M₂) :

@[simp]
theorem submodule.comap_fst {R : Type u} {M : Type v} {M₂ : Type w} {T : semiring R} [add_comm_monoid M] [add_comm_monoid M₂] [semimodule R M] [semimodule R M₂] (p : submodule R M) :

@[simp]
theorem submodule.comap_snd {R : Type u} {M : Type v} {M₂ : Type w} {T : semiring R} [add_comm_monoid M] [add_comm_monoid M₂] [semimodule R M] [semimodule R M₂] (q : submodule R M₂) :

@[simp]
theorem submodule.prod_comap_inl {R : Type u} {M : Type v} {M₂ : Type w} {T : semiring R} [add_comm_monoid M] [add_comm_monoid M₂] [semimodule R M] [semimodule R M₂] (p : submodule R M) (q : submodule R M₂) :

@[simp]
theorem submodule.prod_comap_inr {R : Type u} {M : Type v} {M₂ : Type w} {T : semiring R} [add_comm_monoid M] [add_comm_monoid M₂] [semimodule R M] [semimodule R M₂] (p : submodule R M) (q : submodule R M₂) :

@[simp]
theorem submodule.prod_map_fst {R : Type u} {M : Type v} {M₂ : Type w} {T : semiring R} [add_comm_monoid M] [add_comm_monoid M₂] [semimodule R M] [semimodule R M₂] (p : submodule R M) (q : submodule R M₂) :
submodule.map (linear_map.fst R M M₂) (p.prod q) = p

@[simp]
theorem submodule.prod_map_snd {R : Type u} {M : Type v} {M₂ : Type w} {T : semiring R} [add_comm_monoid M] [add_comm_monoid M₂] [semimodule R M] [semimodule R M₂] (p : submodule R M) (q : submodule R M₂) :
submodule.map (linear_map.snd R M M₂) (p.prod q) = q

@[simp]
theorem submodule.ker_inl {R : Type u} {M : Type v} {M₂ : Type w} {T : semiring R} [add_comm_monoid M] [add_comm_monoid M₂] [semimodule R M] [semimodule R M₂] :
(linear_map.inl R M M₂).ker =

@[simp]
theorem submodule.ker_inr {R : Type u} {M : Type v} {M₂ : Type w} {T : semiring R} [add_comm_monoid M] [add_comm_monoid M₂] [semimodule R M] [semimodule R M₂] :
(linear_map.inr R M M₂).ker =

@[simp]
theorem submodule.range_fst {R : Type u} {M : Type v} {M₂ : Type w} {T : semiring R} [add_comm_monoid M] [add_comm_monoid M₂] [semimodule R M] [semimodule R M₂] :

@[simp]
theorem submodule.range_snd {R : Type u} {M : Type v} {M₂ : Type w} {T : semiring R} [add_comm_monoid M] [add_comm_monoid M₂] [semimodule R M] [semimodule R M₂] :

theorem submodule.disjoint_iff_comap_eq_bot {R : Type u} {M : Type v} {T : ring R} [add_comm_group M] [semimodule R M] {p q : submodule R M} :

def submodule.map_subtype.rel_iso {R : Type u} {M : Type v} {T : ring R} [add_comm_group M] [semimodule R M] (p : submodule R M) :
submodule R p ≃o {p' // p' p}

If N ⊆ M then submodules of N are the same as submodules of M contained in N

Equations
def submodule.map_subtype.order_embedding {R : Type u} {M : Type v} {T : ring R} [add_comm_group M] [semimodule R M] (p : submodule R M) :

If p ⊆ M is a submodule, the ordering of submodules of p is embedded in the ordering of submodules of M.

Equations
@[simp]
theorem submodule.map_subtype_embedding_eq {R : Type u} {M : Type v} {T : ring R} [add_comm_group M] [semimodule R M] (p : submodule R M) (p' : submodule R p) :

def submodule.mkq {R : Type u} {M : Type v} {T : ring R} [add_comm_group M] [semimodule R M] (p : submodule R M) :

The map from a module M to the quotient of M by a submodule p as a linear map.

Equations
@[simp]
theorem submodule.mkq_apply {R : Type u} {M : Type v} {T : ring R} [add_comm_group M] [semimodule R M] (p : submodule R M) (x : M) :

def submodule.liftq {R : Type u} {M : Type v} {M₂ : Type w} {T : ring R} [add_comm_group M] [add_comm_group M₂] [semimodule R M] [semimodule R M₂] (p : submodule R M) (f : M →ₗ[R] M₂) (h : p f.ker) :

The map from the quotient of M by a submodule p to M₂ induced by a linear map f : M → M₂ vanishing on p, as a linear map.

Equations
@[simp]
theorem submodule.liftq_apply {R : Type u} {M : Type v} {M₂ : Type w} {T : ring R} [add_comm_group M] [add_comm_group M₂] [semimodule R M] [semimodule R M₂] (p : submodule R M) (f : M →ₗ[R] M₂) {h : p f.ker} (x : M) :

@[simp]
theorem submodule.liftq_mkq {R : Type u} {M : Type v} {M₂ : Type w} {T : ring R} [add_comm_group M] [add_comm_group M₂] [semimodule R M] [semimodule R M₂] (p : submodule R M) (f : M →ₗ[R] M₂) (h : p f.ker) :
(p.liftq f h).comp p.mkq = f

@[simp]
theorem submodule.range_mkq {R : Type u} {M : Type v} {T : ring R} [add_comm_group M] [semimodule R M] (p : submodule R M) :

@[simp]
theorem submodule.ker_mkq {R : Type u} {M : Type v} {T : ring R} [add_comm_group M] [semimodule R M] (p : submodule R M) :
p.mkq.ker = p

theorem submodule.le_comap_mkq {R : Type u} {M : Type v} {T : ring R} [add_comm_group M] [semimodule R M] (p : submodule R M) (p' : submodule R p.quotient) :

@[simp]
theorem submodule.mkq_map_self {R : Type u} {M : Type v} {T : ring R} [add_comm_group M] [semimodule R M] (p : submodule R M) :

@[simp]
theorem submodule.comap_map_mkq {R : Type u} {M : Type v} {T : ring R} [add_comm_group M] [semimodule R M] (p p' : submodule R M) :

@[simp]
theorem submodule.map_mkq_eq_top {R : Type u} {M : Type v} {T : ring R} [add_comm_group M] [semimodule R M] (p p' : submodule R M) :

def submodule.mapq {R : Type u} {M : Type v} {M₂ : Type w} {T : ring R} [add_comm_group M] [add_comm_group M₂] [semimodule R M] [semimodule R M₂] (p : submodule R M) (q : submodule R M₂) (f : M →ₗ[R] M₂) (h : p submodule.comap f q) :

The map from the quotient of M by submodule p to the quotient of M₂ by submodule q along f : M → M₂ is linear.

Equations
@[simp]
theorem submodule.mapq_apply {R : Type u} {M : Type v} {M₂ : Type w} {T : ring R} [add_comm_group M] [add_comm_group M₂] [semimodule R M] [semimodule R M₂] (p : submodule R M) (q : submodule R M₂) (f : M →ₗ[R] M₂) {h : p submodule.comap f q} (x : M) :

theorem submodule.mapq_mkq {R : Type u} {M : Type v} {M₂ : Type w} {T : ring R} [add_comm_group M] [add_comm_group M₂] [semimodule R M] [semimodule R M₂] (p : submodule R M) (q : submodule R M₂) (f : M →ₗ[R] M₂) {h : p submodule.comap f q} :
(p.mapq q f h).comp p.mkq = q.mkq.comp f

theorem submodule.comap_liftq {R : Type u} {M : Type v} {M₂ : Type w} {T : ring R} [add_comm_group M] [add_comm_group M₂] [semimodule R M] [semimodule R M₂] (p : submodule R M) (q : submodule R M₂) (f : M →ₗ[R] M₂) (h : p f.ker) :

theorem submodule.map_liftq {R : Type u} {M : Type v} {M₂ : Type w} {T : ring R} [add_comm_group M] [add_comm_group M₂] [semimodule R M] [semimodule R M₂] (p : submodule R M) (f : M →ₗ[R] M₂) (h : p f.ker) (q : submodule R p.quotient) :

theorem submodule.ker_liftq {R : Type u} {M : Type v} {M₂ : Type w} {T : ring R} [add_comm_group M] [add_comm_group M₂] [semimodule R M] [semimodule R M₂] (p : submodule R M) (f : M →ₗ[R] M₂) (h : p f.ker) :

theorem submodule.range_liftq {R : Type u} {M : Type v} {M₂ : Type w} {T : ring R} [add_comm_group M] [add_comm_group M₂] [semimodule R M] [semimodule R M₂] (p : submodule R M) (f : M →ₗ[R] M₂) (h : p f.ker) :
(p.liftq f h).range = f.range

theorem submodule.ker_liftq_eq_bot {R : Type u} {M : Type v} {M₂ : Type w} {T : ring R} [add_comm_group M] [add_comm_group M₂] [semimodule R M] [semimodule R M₂] (p : submodule R M) (f : M →ₗ[R] M₂) (h : p f.ker) (h' : f.ker p) :
(p.liftq f h).ker =

def submodule.comap_mkq.rel_iso {R : Type u} {M : Type v} {T : ring R} [add_comm_group M] [semimodule R M] (p : submodule R M) :
submodule R p.quotient ≃o {p' // p p'}

The correspondence theorem for modules: there is an order isomorphism between submodules of the quotient of M by p, and submodules of M larger than p.

Equations
def submodule.comap_mkq.order_embedding {R : Type u} {M : Type v} {T : ring R} [add_comm_group M] [semimodule R M] (p : submodule R M) :

The ordering on submodules of the quotient of M by p embeds into the ordering on submodules of M.

Equations
@[simp]
theorem submodule.comap_mkq_embedding_eq {R : Type u} {M : Type v} {T : ring R} [add_comm_group M] [semimodule R M] (p : submodule R M) (p' : submodule R p.quotient) :

theorem linear_map.range_mkq_comp {R : Type u} {M : Type v} {M₂ : Type w} [ring R] [add_comm_group M] [add_comm_group M₂] [module R M] [module R M₂] (f : M →ₗ[R] M₂) :
f.range.mkq.comp f = 0

theorem linear_map.ker_le_range_iff {R : Type u} {M : Type v} {M₂ : Type w} {M₃ : Type y} [ring R] [add_comm_group M] [add_comm_group M₂] [add_comm_group M₃] [module R M] [module R M₂] [module R M₃] {f : M →ₗ[R] M₂} {g : M₂ →ₗ[R] M₃} :

theorem linear_map.ker_eq_bot_of_cancel {R : Type u} {M : Type v} {M₂ : Type w} [ring R] [add_comm_group M] [add_comm_group M₂] [module R M] [module R M₂] {f : M →ₗ[R] M₂} (h : ∀ (u v : (f.ker) →ₗ[R] M), f.comp u = f.comp vu = v) :

A monomorphism is injective.

theorem linear_map.range_eq_top_of_cancel {R : Type u} {M : Type v} {M₂ : Type w} [ring R] [add_comm_group M] [add_comm_group M₂] [module R M] [module R M₂] {f : M →ₗ[R] M₂} (h : ∀ (u v : M₂ →ₗ[R] f.range.quotient), u.comp f = v.comp fu = v) :

An epimorphism is surjective.

@[simp]
theorem linear_map.range_range_restrict {R : Type u} {M : Type v} {M₂ : Type w} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] [semimodule R M] [semimodule R M₂] (f : M →ₗ[R] M₂) :

Linear equivalences

@[nolint]
structure linear_equiv (R : Type u) (M : Type v) (M₂ : Type w) [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] [semimodule R M] [semimodule R M₂] :
Type (max v w)

A linear equivalence is an invertible linear map.

@[nolint]
def linear_equiv.to_equiv {R : Type u} {M : Type v} {M₂ : Type w} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] [semimodule R M] [semimodule R M₂] (s : M ≃ₗ[R] M₂) :
M M₂

@[nolint]
def linear_equiv.to_linear_map {R : Type u} {M : Type v} {M₂ : Type w} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] [semimodule R M] [semimodule R M₂] (s : M ≃ₗ[R] M₂) :
M →ₗ[R] M₂

@[instance]
def linear_equiv.has_coe {R : Type u} {M : Type v} {M₂ : Type w} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] [semimodule R M] [semimodule R M₂] :
has_coe (M ≃ₗ[R] M₂) (M →ₗ[R] M₂)

Equations
@[instance]
def linear_equiv.has_coe_to_fun {R : Type u} {M : Type v} {M₂ : Type w} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] [semimodule R M] [semimodule R M₂] :

Equations
@[simp]
theorem linear_equiv.mk_apply {R : Type u} {M : Type v} {M₂ : Type w} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] [semimodule R M] [semimodule R M₂] {to_fun : M → M₂} {inv_fun : M₂ → M} {map_add : ∀ (x y : M), to_fun (x + y) = to_fun x + to_fun y} {map_smul : ∀ (c : R) (x : M), to_fun (c x) = c to_fun x} {left_inv : function.left_inverse inv_fun to_fun} {right_inv : function.right_inverse inv_fun to_fun} {a : M} :
{to_fun := to_fun, map_add' := map_add, map_smul' := map_smul, inv_fun := inv_fun, left_inv := left_inv, right_inv := right_inv} a = to_fun a

theorem linear_equiv.to_equiv_injective {R : Type u} {M : Type v} {M₂ : Type w} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] [semimodule R M] [semimodule R M₂] :

@[simp]
theorem linear_equiv.coe_coe {R : Type u} {M : Type v} {M₂ : Type w} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] {semimodule_M : semimodule R M} {semimodule_M₂ : semimodule R M₂} (e : M ≃ₗ[R] M₂) :

@[simp]
theorem linear_equiv.coe_to_equiv {R : Type u} {M : Type v} {M₂ : Type w} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] {semimodule_M : semimodule R M} {semimodule_M₂ : semimodule R M₂} (e : M ≃ₗ[R] M₂) :

@[simp]
theorem linear_equiv.to_fun_apply {R : Type u} {M : Type v} {M₂ : Type w} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] {semimodule_M : semimodule R M} {semimodule_M₂ : semimodule R M₂} (e : M ≃ₗ[R] M₂) {m : M} :
e.to_fun m = e m

@[ext]
theorem linear_equiv.ext {R : Type u} {M : Type v} {M₂ : Type w} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] {semimodule_M : semimodule R M} {semimodule_M₂ : semimodule R M₂} {e e' : M ≃ₗ[R] M₂} (h : ∀ (x : M), e x = e' x) :
e = e'

theorem linear_equiv.eq_of_linear_map_eq {R : Type u} {M : Type v} {M₂ : Type w} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] [semimodule R M] [semimodule R M₂] {f f' : M ≃ₗ[R] M₂} (h : f = f') :
f = f'

def linear_equiv.refl (R : Type u) (M : Type v) [semiring R] [add_comm_monoid M] [semimodule R M] :

The identity map is a linear equivalence.

Equations
@[simp]
theorem linear_equiv.refl_apply {R : Type u} {M : Type v} [semiring R] [add_comm_monoid M] [semimodule R M] (x : M) :

def linear_equiv.symm {R : Type u} {M : Type v} {M₂ : Type w} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] {semimodule_M : semimodule R M} {semimodule_M₂ : semimodule R M₂} (e : M ≃ₗ[R] M₂) :
M₂ ≃ₗ[R] M

Linear equivalences are symmetric.

Equations
@[simp]
theorem linear_equiv.inv_fun_apply {R : Type u} {M : Type v} {M₂ : Type w} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] {semimodule_M : semimodule R M} {semimodule_M₂ : semimodule R M₂} (e : M ≃ₗ[R] M₂) {m : M₂} :
e.inv_fun m = (e.symm) m

def linear_equiv.trans {R : Type u} {M : Type v} {M₂ : Type w} {M₃ : Type y} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] [add_comm_monoid M₃] {semimodule_M : semimodule R M} {semimodule_M₂ : semimodule R M₂} {semimodule_M₃ : semimodule R M₃} (e₁ : M ≃ₗ[R] M₂) (e₂ : M₂ ≃ₗ[R] M₃) :
M ≃ₗ[R] M₃

Linear equivalences are transitive.

Equations
def linear_equiv.to_add_equiv {R : Type u} {M : Type v} {M₂ : Type w} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] {semimodule_M : semimodule R M} {semimodule_M₂ : semimodule R M₂} (e : M ≃ₗ[R] M₂) :
M ≃+ M₂

A linear equivalence is an additive equivalence.

Equations
@[simp]
theorem linear_equiv.coe_to_add_equiv {R : Type u} {M : Type v} {M₂ : Type w} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] {semimodule_M : semimodule R M} {semimodule_M₂ : semimodule R M₂} (e : M ≃ₗ[R] M₂) :

@[simp]
theorem linear_equiv.trans_apply {R : Type u} {M : Type v} {M₂ : Type w} {M₃ : Type y} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] [add_comm_monoid M₃] {semimodule_M : semimodule R M} {semimodule_M₂ : semimodule R M₂} {semimodule_M₃ : semimodule R M₃} (e₁ : M ≃ₗ[R] M₂) (e₂ : M₂ ≃ₗ[R] M₃) (c : M) :
(e₁.trans e₂) c = e₂ (e₁ c)

@[simp]
theorem linear_equiv.apply_symm_apply {R : Type u} {M : Type v} {M₂ : Type w} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] {semimodule_M : semimodule R M} {semimodule_M₂ : semimodule R M₂} (e : M ≃ₗ[R] M₂) (c : M₂) :
e ((e.symm) c) = c

@[simp]
theorem linear_equiv.symm_apply_apply {R : Type u} {M : Type v} {M₂ : Type w} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] {semimodule_M : semimodule R M} {semimodule_M₂ : semimodule R M₂} (e : M ≃ₗ[R] M₂) (b : M) :
(e.symm) (e b) = b

@[simp]
theorem linear_equiv.symm_trans_apply {R : Type u} {M : Type v} {M₂ : Type w} {M₃ : Type y} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] [add_comm_monoid M₃] {semimodule_M : semimodule R M} {semimodule_M₂ : semimodule R M₂} {semimodule_M₃ : semimodule R M₃} (e₁ : M ≃ₗ[R] M₂) (e₂ : M₂ ≃ₗ[R] M₃) (c : M₃) :
((e₁.trans e₂).symm) c = (e₁.symm) ((e₂.symm) c)

@[simp]
theorem linear_equiv.trans_refl {R : Type u} {M : Type v} {M₂ : Type w} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] {semimodule_M : semimodule R M} {semimodule_M₂ : semimodule R M₂} (e : M ≃ₗ[R] M₂) :

@[simp]
theorem linear_equiv.refl_trans {R : Type u} {M : Type v} {M₂ : Type w} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] {semimodule_M : semimodule R M} {semimodule_M₂ : semimodule R M₂} (e : M ≃ₗ[R] M₂) :

theorem linear_equiv.symm_apply_eq {R : Type u} {M : Type v} {M₂ : Type w} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] {semimodule_M : semimodule R M} {semimodule_M₂ : semimodule R M₂} (e : M ≃ₗ[R] M₂) {x : M₂} {y : M} :
(e.symm) x = y x = e y

theorem linear_equiv.eq_symm_apply {R : Type u} {M : Type v} {M₂ : Type w} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] {semimodule_M : semimodule R M} {semimodule_M₂ : semimodule R M₂} (e : M ≃ₗ[R] M₂) {x : M₂} {y : M} :
y = (e.symm) x e y = x

@[simp]
theorem linear_equiv.trans_symm {R : Type u} {M : Type v} {M₂ : Type w} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] [semimodule R M] [semimodule R M₂] (f : M ≃ₗ[R] M₂) :

@[simp]
theorem linear_equiv.symm_trans {R : Type u} {M : Type v} {M₂ : Type w} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] [semimodule R M] [semimodule R M₂] (f : M ≃ₗ[R] M₂) :

@[simp]

@[simp]
theorem linear_equiv.comp_coe {R : Type u} {M : Type v} {M₂ : Type w} {M₃ : Type y} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] [add_comm_monoid M₃] [semimodule R M] [semimodule R M₂] [semimodule R M₃] (f : M ≃ₗ[R] M₂) (f' : M₂ ≃ₗ[R] M₃) :
f'.comp f = (f.trans f')

@[simp]
theorem linear_equiv.map_add {R : Type u} {M : Type v} {M₂ : Type w} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] {semimodule_M : semimodule R M} {semimodule_M₂ : semimodule R M₂} (e : M ≃ₗ[R] M₂) (a b : M) :
e (a + b) = e a + e b

@[simp]
theorem linear_equiv.map_zero {R : Type u} {M : Type v} {M₂ : Type w} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] {semimodule_M : semimodule R M} {semimodule_M₂ : semimodule R M₂} (e : M ≃ₗ[R] M₂) :
e 0 = 0

@[simp]
theorem linear_equiv.map_smul {R : Type u} {M : Type v} {M₂ : Type w} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] {semimodule_M : semimodule R M} {semimodule_M₂ : semimodule R M₂} (e : M ≃ₗ[R] M₂) (c : R) (x : M) :
e (c x) = c e x

@[simp]
theorem linear_equiv.map_sum {R : Type u} {M : Type v} {M₂ : Type w} {ι : Type x} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] {semimodule_M : semimodule R M} {semimodule_M₂ : semimodule R M₂} (e : M ≃ₗ[R] M₂) {s : finset ι} (u : ι → M) :
e (∑ (i : ι) in s, u i) = ∑ (i : ι) in s, e (u i)

@[simp]
theorem linear_equiv.map_eq_zero_iff {R : Type u} {M : Type v} {M₂ : Type w} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] {semimodule_M : semimodule R M} {semimodule_M₂ : semimodule R M₂} (e : M ≃ₗ[R] M₂) {x : M} :
e x = 0 x = 0

theorem linear_equiv.map_ne_zero_iff {R : Type u} {M : Type v} {M₂ : Type w} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] {semimodule_M : semimodule R M} {semimodule_M₂ : semimodule R M₂} (e : M ≃ₗ[R] M₂) {x : M} :
e x 0 x 0

@[simp]
theorem linear_equiv.symm_symm {R : Type u} {M : Type v} {M₂ : Type w} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] {semimodule_M : semimodule R M} {semimodule_M₂ : semimodule R M₂} (e : M ≃ₗ[R] M₂) :
e.symm.symm = e

theorem linear_equiv.bijective {R : Type u} {M : Type v} {M₂ : Type w} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] {semimodule_M : semimodule R M} {semimodule_M₂ : semimodule R M₂} (e : M ≃ₗ[R] M₂) :

theorem linear_equiv.injective {R : Type u} {M : Type v} {M₂ : Type w} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] {semimodule_M : semimodule R M} {semimodule_M₂ : semimodule R M₂} (e : M ≃ₗ[R] M₂) :

theorem linear_equiv.surjective {R : Type u} {M : Type v} {M₂ : Type w} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] {semimodule_M : semimodule R M} {semimodule_M₂ : semimodule R M₂} (e : M ≃ₗ[R] M₂) :

theorem linear_equiv.image_eq_preimage {R : Type u} {M : Type v} {M₂ : Type w} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] {semimodule_M : semimodule R M} {semimodule_M₂ : semimodule R M₂} (e : M ≃ₗ[R] M₂) (s : set M) :
e '' s = (e.symm) ⁻¹' s

theorem linear_equiv.map_eq_comap {R : Type u} {M : Type v} {M₂ : Type w} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] {semimodule_M : semimodule R M} {semimodule_M₂ : semimodule R M₂} (e : M ≃ₗ[R] M₂) {p : submodule R M} :

def linear_equiv.of_submodule {R : Type u} {M : Type v} {M₂ : Type w} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] {semimodule_M : semimodule R M} {semimodule_M₂ : semimodule R M₂} (e : M ≃ₗ[R] M₂) (p : submodule R M) :

A linear equivalence of two modules restricts to a linear equivalence from any submodule of the domain onto the image of the submodule.

Equations
@[simp]
theorem linear_equiv.of_submodule_apply {R : Type u} {M : Type v} {M₂ : Type w} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] {semimodule_M : semimodule R M} {semimodule_M₂ : semimodule R M₂} (e : M ≃ₗ[R] M₂) (p : submodule R M) (x : p) :

@[simp]
theorem linear_equiv.of_submodule_symm_apply {R : Type u} {M : Type v} {M₂ : Type w} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] {semimodule_M : semimodule R M} {semimodule_M₂ : semimodule R M₂} (e : M ≃ₗ[R] M₂) (p : submodule R M) (x : (submodule.map e p)) :

def linear_equiv.prod {R : Type u} {M : Type v} {M₂ : Type w} {M₃ : Type y} {M₄ : Type z} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] [add_comm_monoid M₃] [add_comm_monoid M₄] {semimodule_M : semimodule R M} {semimodule_M₂ : semimodule R M₂} {semimodule_M₃ : semimodule R M₃} {semimodule_M₄ : semimodule R M₄} (e₁ : M ≃ₗ[R] M₂) (e₂ : M₃ ≃ₗ[R] M₄) :
(M × M₃) ≃ₗ[R] M₂ × M₄

Product of linear equivalences; the maps come from equiv.prod_congr.

Equations
theorem linear_equiv.prod_symm {R : Type u} {M : Type v} {M₂ : Type w} {M₃ : Type y} {M₄ : Type z} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] [add_comm_monoid M₃] [add_comm_monoid M₄] {semimodule_M : semimodule R M} {semimodule_M₂ : semimodule R M₂} {semimodule_M₃ : semimodule R M₃} {semimodule_M₄ : semimodule R M₄} (e₁ : M ≃ₗ[R] M₂) (e₂ : M₃ ≃ₗ[R] M₄) :
(e₁.prod e₂).symm = e₁.symm.prod e₂.symm

@[simp]
theorem linear_equiv.prod_apply {R : Type u} {M : Type v} {M₂ : Type w} {M₃ : Type y} {M₄ : Type z} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] [add_comm_monoid M₃] [add_comm_monoid M₄] {semimodule_M : semimodule R M} {semimodule_M₂ : semimodule R M₂} {semimodule_M₃ : semimodule R M₃} {semimodule_M₄ : semimodule R M₄} (e₁ : M ≃ₗ[R] M₂) (e₂ : M₃ ≃ₗ[R] M₄) (p : M × M₃) :
(e₁.prod e₂) p = (e₁ p.fst, e₂ p.snd)

@[simp]
theorem linear_equiv.coe_prod {R : Type u} {M : Type v} {M₂ : Type w} {M₃ : Type y} {M₄ : Type z} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] [add_comm_monoid M₃] [add_comm_monoid M₄] {semimodule_M : semimodule R M} {semimodule_M₂ : semimodule R M₂} {semimodule_M₃ : semimodule R M₃} {semimodule_M₄ : semimodule R M₄} (e₁ : M ≃ₗ[R] M₂) (e₂ : M₃ ≃ₗ[R] M₄) :
(e₁.prod e₂) = e₁.prod_map e₂

def linear_equiv.uncurry (R : Type u) (V : Type v') (V₂ : Type w') [semiring R] :
(V → V₂ → R) ≃ₗ[R] V × V₂ → R

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

Equations
@[simp]
theorem linear_equiv.coe_uncurry (R : Type u) (V : Type v') (V₂ : Type w') [semiring R] :

@[simp]
theorem linear_equiv.coe_uncurry_symm (R : Type u) (V : Type v') (V₂ : Type w') [semiring R] :

def linear_equiv.of_eq {R : Type u} {M : Type v} [semiring R] [add_comm_monoid M] {semimodule_M : semimodule R M} (p q : submodule R M) (h : p = q) :

Linear equivalence between two equal submodules.

Equations
@[simp]
theorem linear_equiv.coe_of_eq_apply {R : Type u} {M : Type v} [semiring R] [add_comm_monoid M] {semimodule_M : semimodule R M} {p q : submodule R M} (h : p = q) (x : p) :

@[simp]
theorem linear_equiv.of_eq_symm {R : Type u} {M : Type v} [semiring R] [add_comm_monoid M] {semimodule_M : semimodule R M} {p q : submodule R M} (h : p = q) :

def linear_equiv.of_submodules {R : Type u} {M : Type v} {M₂ : Type w} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] {semimodule_M : semimodule R M} {semimodule_M₂ : semimodule R M₂} (e : M ≃ₗ[R] M₂) (p : submodule R M) (q : submodule R M₂) (h : submodule.map e p = q) :

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

Equations
@[simp]
theorem linear_equiv.of_submodules_apply {R : Type u} {M : Type v} {M₂ : Type w} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] {semimodule_M : semimodule R M} {semimodule_M₂ : semimodule R M₂} (e : M ≃ₗ[R] M₂) {p : submodule R M} {q : submodule R M₂} (h : submodule.map e p = q) (x : p) :
((e.of_submodules p q h) x) = e x

@[simp]
theorem linear_equiv.of_submodules_symm_apply {R : Type u} {M : Type v} {M₂ : Type w} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] {semimodule_M : semimodule R M} {semimodule_M₂ : semimodule R M₂} (e : M ≃ₗ[R] M₂) {p : submodule R M} {q : submodule R M₂} (h : submodule.map e p = q) (x : q) :
(((e.of_submodules p q h).symm) x) = (e.symm) x

def linear_equiv.of_top {R : Type u} {M : Type v} [semiring R] [add_comm_monoid M] {semimodule_M : semimodule R M} (p : submodule R M) (h : p = ) :

The top submodule of M is linearly equivalent to M.

Equations
@[simp]
theorem linear_equiv.of_top_apply {R : Type u} {M : Type v} [semiring R] [add_comm_monoid M] {semimodule_M : semimodule R M} (p : submodule R M) {h : p = } (x : p) :

@[simp]
theorem linear_equiv.coe_of_top_symm_apply {R : Type u} {M : Type v} [semiring R] [add_comm_monoid M] {semimodule_M : semimodule R M} (p : submodule R M) {h : p = } (x : M) :

theorem linear_equiv.of_top_symm_apply {R : Type u} {M : Type v} [semiring R] [add_comm_monoid M] {semimodule_M : semimodule R M} (p : submodule R M) {h : p = } (x : M) :
((linear_equiv.of_top p h).symm) x = x, _⟩

def linear_equiv.of_linear {R : Type u} {M : Type v} {M₂ : Type w} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] {semimodule_M : semimodule R M} {semimodule_M₂ : semimodule R M₂} (f : M →ₗ[R] M₂) (g : M₂ →ₗ[R] M) (h₁ : f.comp g = linear_map.id) (h₂ : g.comp f = linear_map.id) :
M ≃ₗ[R] M₂

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

Equations
@[simp]
theorem linear_equiv.of_linear_apply {R : Type u} {M : Type v} {M₂ : Type w} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] {semimodule_M : semimodule R M} {semimodule_M₂ : semimodule R M₂} (f : M →ₗ[R] M₂) (g : M₂ →ₗ[R] M) {h₁ : f.comp g = linear_map.id} {h₂ : g.comp f = linear_map.id} (x : M) :
(linear_equiv.of_linear f g h₁ h₂) x = f x

@[simp]
theorem linear_equiv.of_linear_symm_apply {R : Type u} {M : Type v} {M₂ : Type w} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] {semimodule_M : semimodule R M} {semimodule_M₂ : semimodule R M₂} (f : M →ₗ[R] M₂) (g : M₂ →ₗ[R] M) {h₁ : f.comp g = linear_map.id} {h₂ : g.comp f = linear_map.id} (x : M₂) :
((linear_equiv.of_linear f g h₁ h₂).symm) x = g x

@[simp]
theorem linear_equiv.range {R : Type u} {M : Type v} {M₂ : Type w} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] {semimodule_M : semimodule R M} {semimodule_M₂ : semimodule R M₂} (e : M ≃ₗ[R] M₂) :

theorem linear_equiv.eq_bot_of_equiv {R : Type u} {M : Type v} {M₂ : Type w} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] {semimodule_M : semimodule R M} (p : submodule R M) [semimodule R M₂] (e : p ≃ₗ[R] ) :
p =

@[simp]
theorem linear_equiv.ker {R : Type u} {M : Type v} {M₂ : Type w} [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] {semimodule_M : semimodule R M} {semimodule_M₂ : semimodule R M₂} (e : M ≃ₗ[R] M₂) :

@[simp]
theorem linear_equiv.map_neg {R : Type u} {M : Type v} {M₂ : Type w} [semiring R] [add_comm_group M] [add_comm_group M₂] {semimodule_M : semimodule R M} {semimodule_M₂ : semimodule R M₂} (e : M ≃ₗ[R] M₂) (a : M) :
e (-a) = -e a

@[simp]
theorem linear_equiv.map_sub {R : Type u} {M : Type v} {M₂ : Type w} [semiring R] [add_comm_group M] [add_comm_group M₂] {semimodule_M : semimodule R M} {semimodule_M₂ : semimodule R M₂} (e : M ≃ₗ[R] M₂) (a b : M) :
e (a - b) = e a - e b

def linear_equiv.skew_prod {R : Type u} {M : Type v} {M₂ : Type w} {M₃ : Type y} {M₄ : Type z} [semiring R] [add_comm_group M] [add_comm_group M₂] [add_comm_group M₃] [add_comm_group M₄] {semimodule_M : semimodule R M} {semimodule_M₂ : semimodule R M₂} {semimodule_M₃ : semimodule R M₃} {semimodule_M₄ : semimodule R M₄} (e₁ : M ≃ₗ[R] M₂) (e₂ : M₃ ≃ₗ[R] M₄) (f : M →ₗ[R] M₄) :
(M × M₃) ≃ₗ[R] M₂ × M₄

Equivalence given by a block lower diagonal matrix. e₁ and e₂ are diagonal square blocks, and f is a rectangular block below the diagonal.

Equations
@[simp]
theorem linear_equiv.skew_prod_apply {R : Type u} {M : Type v} {M₂ : Type w} {M₃ : Type y} {M₄ : Type z} [semiring R] [add_comm_group M] [add_comm_group M₂] [add_comm_group M₃] [add_comm_group M₄] {semimodule_M : semimodule R M} {semimodule_M₂ : semimodule R M₂} {semimodule_M₃ : semimodule R M₃} {semimodule_M₄ : semimodule R M₄} (e₁ : M ≃ₗ[R] M₂) (e₂ : M₃ ≃ₗ[R] M₄) (f : M →ₗ[R] M₄) (x : M × M₃) :
(e₁.skew_prod e₂ f) x = (e₁ x.fst, e₂ x.snd + f x.fst)

@[simp]
theorem linear_equiv.skew_prod_symm_apply {R : Type u} {M : Type v} {M₂ : Type w} {M₃ : Type y} {M₄ : Type z} [semiring R] [add_comm_group M] [add_comm_group M₂] [add_comm_group M₃] [add_comm_group M₄] {semimodule_M : semimodule R M} {semimodule_M₂ : semimodule R M₂} {semimodule_M₃ : semimodule R M₃} {semimodule_M₄ : semimodule R M₄} (e₁ : M ≃ₗ[R] M₂) (e₂ : M₃ ≃ₗ[R] M₄) (f : M →ₗ[R] M₄) (x : M₂ × M₄) :
((e₁.skew_prod e₂ f).symm) x = ((e₁.symm) x.fst, (e₂.symm) (x.snd - f ((e₁.symm) x.fst)))

def linear_equiv.of_injective {R : Type u} {M : Type v} {M₂ : Type w} [ring R] [add_comm_group M] [add_comm_group M₂] {semimodule_M : semimodule R M} {semimodule_M₂ : semimodule R M₂} (f : M →ₗ[R] M₂) (h : f.ker = ) :

An injective linear map f : M →ₗ[R] M₂ defines a linear equivalence between M and f.range.

Equations
@[simp]
theorem linear_equiv.of_injective_apply {R : Type u} {M : Type v} {M₂ : Type w} [ring R] [add_comm_group M] [add_comm_group M₂] {semimodule_M : semimodule R M} {semimodule_M₂ : semimodule R M₂} (f : M →ₗ[R] M₂) {h : f.ker = } (x : M) :

def linear_equiv.of_bijective {R : Type u} {M : Type v} {M₂ : Type w} [ring R] [add_comm_group M] [add_comm_group M₂] {semimodule_M : semimodule R M} {semimodule_M₂ : semimodule R M₂} (f : M →ₗ[R] M₂) (hf₁ : f.ker = ) (hf₂ : f.range = ) :
M ≃ₗ[R] M₂

A bijective linear map is a linear equivalence. Here, bijectivity is described by saying that the kernel of f is {0} and the range is the universal set.

Equations
@[simp]
theorem linear_equiv.of_bijective_apply {R : Type u} {M : Type v} {M₂ : Type w} [ring R] [add_comm_group M] [add_comm_group M₂] {semimodule_M : semimodule R M} {semimodule_M₂ : semimodule R M₂} (f : M →ₗ[R] M₂) {hf₁ : f.ker = } {hf₂ : f.range = } (x : M) :
(linear_equiv.of_bijective f hf₁ hf₂) x = f x

def linear_equiv.smul_of_unit {R : Type u} {M : Type v} [comm_ring R] [add_comm_group M] [semimodule R M] (a : units R) :

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

Equations
def linear_equiv.arrow_congr {R : Type u_1} {M₁ : Type u_2} {M₂ : Type u_3} {M₂₁ : Type u_4} {M₂₂ : Type u_5} [comm_ring R] [add_comm_group M₁] [add_comm_group M₂] [add_comm_group M₂₁] [add_comm_group 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
@[simp]
theorem linear_equiv.arrow_congr_apply {R : Type u_1} {M₁ : Type u_2} {M₂ : Type u_3} {M₂₁ : Type u_4} {M₂₂ : Type u_5} [comm_ring R] [add_comm_group M₁] [add_comm_group M₂] [add_comm_group M₂₁] [add_comm_group 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₁.arrow_congr e₂) f) x = e₂ (f ((e₁.symm) x))

@[simp]
theorem linear_equiv.arrow_congr_symm_apply {R : Type u_1} {M₁ : Type u_2} {M₂ : Type u_3} {M₂₁ : Type u_4} {M₂₂ : Type u_5} [comm_ring R] [add_comm_group M₁] [add_comm_group M₂] [add_comm_group M₂₁] [add_comm_group 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₁.arrow_congr e₂).symm) f) x = (e₂.symm) (f (e₁ x))

theorem linear_equiv.arrow_congr_comp {R : Type u} {M : Type v} {M₂ : Type w} {M₃ : Type y} [comm_ring R] [add_comm_group M] [add_comm_group M₂] [add_comm_group M₃] [semimodule R M] [semimodule R M₂] [semimodule R M₃] {N : Type u_1} {N₂ : Type u_2} {N₃ : Type u_3} [add_comm_group N] [add_comm_group N₂] [add_comm_group 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₁.arrow_congr e₃) (g.comp f) = ((e₂.arrow_congr e₃) g).comp ((e₁.arrow_congr e₂) f)

theorem linear_equiv.arrow_congr_trans {R : Type u} [comm_ring R] {M₁ : Type u_1} {M₂ : Type u_2} {M₃ : Type u_3} {N₁ : Type u_4} {N₂ : Type u_5} {N₃ : Type u_6} [add_comm_group M₁] [module R M₁] [add_comm_group M₂] [module R M₂] [add_comm_group M₃] [module R M₃] [add_comm_group N₁] [module R N₁] [add_comm_group N₂] [module R N₂] [add_comm_group N₃] [module R N₃] (e₁ : M₁ ≃ₗ[R] M₂) (e₂ : N₁ ≃ₗ[R] N₂) (e₃ : M₂ ≃ₗ[R] M₃) (e₄ : N₂ ≃ₗ[R] N₃) :
(e₁.arrow_congr e₂).trans (e₃.arrow_congr e₄) = (e₁.trans e₃).arrow_congr (e₂.trans e₄)

def linear_equiv.congr_right {R : Type u} {M : Type v} {M₂ : Type w} {M₃ : Type y} [comm_ring R] [add_comm_group M] [add_comm_group M₂] [add_comm_group M₃] [semimodule R M] [semimodule R M₂] [semimodule 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
def linear_equiv.conj {R : Type u} {M : Type v} {M₂ : Type w} [comm_ring R] [add_comm_group M] [add_comm_group M₂] [semimodule R M] [semimodule 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
theorem linear_equiv.conj_apply {R : Type u} {M : Type v} {M₂ : Type w} [comm_ring R] [add_comm_group M] [add_comm_group M₂] [semimodule R M] [semimodule R M₂] (e : M ≃ₗ[R] M₂) (f : module.End R M) :
(e.conj) f = (e.comp f).comp (e.symm)

theorem linear_equiv.symm_conj_apply {R : Type u} {M : Type v} {M₂ : Type w} [comm_ring R] [add_comm_group M] [add_comm_group M₂] [semimodule R M] [semimodule R M₂] (e : M ≃ₗ[R] M₂) (f : module.End R M₂) :
(e.symm.conj) f = ((e.symm).comp f).comp e

theorem linear_equiv.conj_comp {R : Type u} {M : Type v} {M₂ : Type w} [comm_ring R] [add_comm_group M] [add_comm_group M₂] [semimodule R M] [semimodule R M₂] (e : M ≃ₗ[R] M₂) (f g : module.End R M) :

theorem linear_equiv.conj_trans {R : Type u} {M : Type v} {M₂ : Type w} {M₃ : Type y} [comm_ring R] [add_comm_group M] [add_comm_group M₂] [add_comm_group M₃] [semimodule R M] [semimodule R M₂] [semimodule R M₃] (e₁ : M ≃ₗ[R] M₂) (e₂ : M₂ ≃ₗ[R] M₃) :
e₁.conj.trans e₂.conj = (e₁.trans e₂).conj

@[simp]
theorem linear_equiv.conj_id {R : Type u} {M : Type v} {M₂ : Type w} [comm_ring R] [add_comm_group M] [add_comm_group M₂] [semimodule R M] [semimodule R M₂] (e : M ≃ₗ[R] M₂) :

def linear_equiv.smul_of_ne_zero (K : Type u') (M : Type v) [field K] [add_comm_group M] [module K M] (a : K) (ha : a 0) :

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

Equations
theorem linear_equiv.ker_to_span_singleton (K : Type u') (M : Type v) [field K] [add_comm_group M] [module K M] {x : M} (h : x 0) :

def linear_equiv.to_span_nonzero_singleton (K : Type u') (M : Type v) [field K] [add_comm_group M] [module K M] (x : M) (h : x 0) :

Given a nonzero element x of a vector space M over a field K, the natural map from K to the span of x, with invertibility check to consider it as an isomorphism.

Equations
theorem linear_equiv.to_span_nonzero_singleton_one (K : Type u') (M : Type v) [field K] [add_comm_group M] [module K M] (x : M) (h : x 0) :

def linear_equiv.coord (K : Type u') (M : Type v) [field K] [add_comm_group M] [module K M] (x : M) (h : x 0) :

Given a nonzero element x of a vector space M over a field K, the natural map from the span of x to K.

theorem linear_equiv.coord_self (K : Type u') (M : Type v) [field K] [add_comm_group M] [module K M] (x : M) (h : x 0) :
(linear_equiv.coord K M x h) x, _⟩ = 1

def submodule.comap_subtype_equiv_of_le {R : Type u} {M : Type v} [semiring R] [add_comm_monoid M] [semimodule R M] {p q : submodule R M} (hpq : p q) :

If s ≤ t, then we can view s as a submodule of t by taking the comap of t.subtype.

Equations
def submodule.quot_equiv_of_eq_bot {R : Type u} {M : Type v} [ring R] [add_comm_group M] [module R M] (p : submodule R M) (hp : p = ) :

If p = ⊥, then M / p ≃ₗ[R] M.

Equations
@[simp]
theorem submodule.quot_equiv_of_eq_bot_apply_mk {R : Type u} {M : Type v} [ring R] [add_comm_group M] [module R M] (p : submodule R M) (hp : p = ) (x : M) :

@[simp]
theorem submodule.quot_equiv_of_eq_bot_symm_apply {R : Type u} {M : Type v} [ring R] [add_comm_group M] [module R M] (p : submodule R M) (hp : p = ) (x : M) :

@[simp]
theorem submodule.coe_quot_equiv_of_eq_bot_symm {R : Type u} {M : Type v} [ring R] [add_comm_group M] [module R M] (p : submodule R M) (hp : p = ) :

def submodule.quot_equiv_of_eq {R : Type u} {M : Type v} [ring R] [add_comm_group M] [module R M] (p q : submodule R M) (h : p = q) :

Quotienting by equal submodules gives linearly equivalent quotients.

Equations
@[simp]
theorem submodule.mem_map_equiv {R : Type u} {M : Type v} {M₂ : Type w} [comm_ring R] [add_comm_group M] [add_comm_group M₂] [module R M] [module R M₂] (p : submodule R M) {e : M ≃ₗ[R] M₂} {x : M₂} :

theorem submodule.comap_le_comap_smul {R : Type u} {M : Type v} {M₂ : Type w} [comm_ring R] [add_comm_group M] [add_comm_group M₂] [module R M] [module R M₂] (q : submodule R M₂) (f : M →ₗ[R] M₂) (c : R) :

theorem submodule.inf_comap_le_comap_add {R : Type u} {M : Type v} {M₂ : Type w} [comm_ring R] [add_comm_group M] [add_comm_group M₂] [module R M] [module R M₂] (q : submodule R M₂) (f₁ f₂ : M →ₗ[R] M₂) :

def submodule.compatible_maps {R : Type u} {M : Type v} {M₂ : Type w} [comm_ring R] [add_comm_group M] [add_comm_group M₂] [module R M] [module R M₂] (p : submodule R M) (q : submodule R M₂) :
submodule R (M →ₗ[R] M₂)

Given modules M, M₂ over a commutative ring, together with submodules p ⊆ M, q ⊆ M₂, the set of maps $\{f ∈ Hom(M, M₂) | f(p) ⊆ q \}$ is a submodule of Hom(M, M₂).

Equations
def submodule.mapq_linear {R : Type u} {M : Type v} {M₂ : Type w} [comm_ring R] [add_comm_group M] [add_comm_group M₂] [module R M] [module R M₂] (p : submodule R M) (q : submodule R M₂) :

Given modules M, M₂ over a commutative ring, together with submodules p ⊆ M, q ⊆ M₂, the natural map $\{f ∈ Hom(M, M₂) | f(p) ⊆ q \} \to Hom(M/p, M₂/q)$ is linear.

Equations
def equiv.to_linear_equiv {R : Type u} {M : Type v} {M₂ : Type w} [semiring R] [add_comm_monoid M] [semimodule R M] [add_comm_monoid M₂] [semimodule R M₂] (e : M M₂) (h : is_linear_map R e) :
M ≃ₗ[R] M₂

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

Equations
def linear_map.quot_ker_equiv_range {R : Type u} {M : Type v} {M₂ : Type w} [ring R] [add_comm_group M] [add_comm_group M₂] [module R M] [module R M₂] (f : M →ₗ[R] M₂) :

The first isomorphism law for modules. The quotient of M by the kernel of f is linearly equivalent to the range of f.

Equations
@[simp]
theorem linear_map.quot_ker_equiv_range_apply_mk {R : Type u} {M : Type v} {M₂ : Type w} [ring R] [add_comm_group M] [add_comm_group M₂] [module R M] [module R M₂] (f : M →ₗ[R] M₂) (x : M) :

@[simp]
theorem linear_map.quot_ker_equiv_range_symm_apply_image {R : Type u} {M : Type v} {M₂ : Type w} [ring R] [add_comm_group M] [add_comm_group M₂] [module R M] [module R M₂] (f : M →ₗ[R] M₂) (x : M) (h : f x f.range) :

def linear_map.quotient_inf_to_sup_quotient {R : Type u} {M : Type v} [ring R] [add_comm_group M] [module R M] (p p' : submodule R M) :

Canonical linear map from the quotient p/(p ∩ p') to (p+p')/p', mapping x + (p ∩ p') to x + p', where p and p' are submodules of an ambient module.

Equations

Second Isomorphism Law : the canonical map from p/(p ∩ p') to (p+p')/p' as a linear isomorphism.

Equations
theorem linear_map.is_linear_map_prod_iso {R : Type u_1} {M : Type u_2} {M₂ : Type u_3} {M₃ : Type u_4} [comm_semiring R] [add_comm_monoid M] [add_comm_monoid M₂] [add_comm_group M₃] [semimodule R M] [semimodule R M₂] [semimodule R M₃] :
is_linear_map R (λ (p : (M →ₗ[R] M₂) × (M →ₗ[R] M₃)), p.fst.prod p.snd)

def linear_map.pi {R : Type u} {M₂ : Type w} {ι : Type x} [semiring R] [add_comm_monoid M₂] [semimodule R M₂] {φ : ι → Type i} [Π (i : ι), add_comm_monoid (φ i)] [Π (i : ι), semimodule R (φ i)] (f : Π (i : ι), M₂ →ₗ[R] φ i) :
M₂ →ₗ[R] Π (i : ι), φ i

pi construction for linear functions. From a family of linear functions it produces a linear function into a family of modules.

Equations
@[simp]
theorem linear_map.pi_apply {R : Type u} {M₂ : Type w} {ι : Type x} [semiring R] [add_comm_monoid M₂] [semimodule R M₂] {φ : ι → Type i} [Π (i : ι), add_comm_monoid (φ i)] [Π (i : ι), semimodule R (φ i)] (f : Π (i : ι), M₂ →ₗ[R] φ i) (c : M₂) (i : ι) :
(linear_map.pi f) c i = (f i) c

theorem linear_map.ker_pi {R : Type u} {M₂ : Type w} {ι : Type x} [semiring R] [add_comm_monoid M₂] [semimodule R M₂] {φ : ι → Type i} [Π (i : ι), add_comm_monoid (φ i)] [Π (i : ι), semimodule R (φ i)] (f : Π (i : ι), M₂ →ₗ[R] φ i) :
(linear_map.pi f).ker = ⨅ (i : ι), (f i).ker

theorem linear_map.pi_eq_zero {R : Type u} {M₂ : Type w} {ι : Type x} [semiring R] [add_comm_monoid M₂] [semimodule R M₂] {φ : ι → Type i} [Π (i : ι), add_comm_monoid (φ i)] [Π (i : ι), semimodule R (φ i)] (f : Π (i : ι), M₂ →ₗ[R] φ i) :
linear_map.pi f = 0 ∀ (i : ι), f i = 0

theorem linear_map.pi_zero {R : Type u} {M₂ : Type w} {ι : Type x} [semiring R] [add_comm_monoid M₂] [semimodule R M₂] {φ : ι → Type i} [Π (i : ι), add_comm_monoid (φ i)] [Π (i : ι), semimodule R (φ i)] :
linear_map.pi (λ (i : ι), 0) = 0

theorem linear_map.pi_comp {R : Type u} {M₂ : Type w} {M₃ : Type y} {ι : Type x} [semiring R] [add_comm_monoid M₂] [semimodule R M₂] [add_comm_monoid M₃] [semimodule R M₃] {φ : ι → Type i} [Π (i : ι), add_comm_monoid (φ i)] [Π (i : ι), semimodule R (φ i)] (f : Π (i : ι), M₂ →ₗ[R] φ i) (g : M₃ →ₗ[R] M₂) :
(linear_map.pi f).comp g = linear_map.pi (λ (i : ι), (f i).comp g)

def linear_map.proj {R : Type u} {ι : Type x} [semiring R] {φ : ι → Type i} [Π (i : ι), add_comm_monoid (φ i)] [Π (i : ι), semimodule R (φ i)] (i : ι) :
(Π (i : ι), φ i) →ₗ[R] φ i

The projections from a family of modules are linear maps.

Equations
@[simp]
theorem linear_map.proj_apply {R : Type u} {ι : Type x} [semiring R] {φ : ι → Type i} [Π (i : ι), add_comm_monoid (φ i)] [Π (i : ι), semimodule R (φ i)] (i : ι) (b : Π (i : ι), φ i) :

theorem linear_map.proj_pi {R : Type u} {M₂ : Type w} {ι : Type x} [semiring R] [add_comm_monoid M₂] [semimodule R M₂] {φ : ι → Type i} [Π (i : ι), add_comm_monoid (φ i)] [Π (i : ι), semimodule R (φ i)] (f : Π (i : ι), M₂ →ₗ[R] φ i) (i : ι) :

theorem linear_map.infi_ker_proj {R : Type u} {ι : Type x} [semiring R] {φ : ι → Type i} [Π (i : ι), add_comm_monoid (φ i)] [Π (i : ι), semimodule R (φ i)] :
(⨅ (i : ι), (linear_map.proj i).ker) =

def linear_map.infi_ker_proj_equiv (R : Type u) {ι : Type x} [semiring R] (φ : ι → Type i) [Π (i : ι), add_comm_monoid (φ i)] [Π (i : ι), semimodule R (φ i)] {I J : set ι} [decidable_pred (λ (i : ι), i I)] (hd : disjoint I J) (hu : set.univ I J) :
(⨅ (i : ι) (H : i J), (linear_map.proj i).ker) ≃ₗ[R] Π (i : I), φ i

If I and J are disjoint index sets, the product of the kernels of the Jth projections of φ is linearly equivalent to the product over I.

Equations
def linear_map.diag {R : Type u} {ι : Type x} [semiring R] {φ : ι → Type i} [Π (i : ι), add_comm_monoid (φ i)] [Π (i : ι), semimodule R (φ i)] [decidable_eq ι] (i j : ι) :
φ i →ₗ[R] φ j

diag i j is the identity map if i = j. Otherwise it is the constant 0 map.

Equations
theorem linear_map.update_apply {R : Type u} {M₂ : Type w} {ι : Type x} [semiring R] [add_comm_monoid M₂] [semimodule R M₂] {φ : ι → Type i} [Π (i : ι), add_comm_monoid (φ i)] [Π (i : ι), semimodule R (φ i)] [decidable_eq ι] (f : Π (i : ι), M₂ →ₗ[R] φ i) (c : M₂) (i j : ι) (b : M₂ →ₗ[R] φ i) :
(function.update f i b j) c = function.update (λ (i : ι), (f i) c) i (b c) j

def linear_map.std_basis (R : Type u) {ι : Type x} [semiring R] (φ : ι → Type i) [Π (i : ι), add_comm_monoid (φ i)] [Π (i : ι), semimodule R (φ i)] [decidable_eq ι] (i : ι) :
φ i →ₗ[R] Π (i : ι), φ i

The standard basis of the product of φ.

Equations
theorem linear_map.std_basis_apply (R : Type u) {ι : Type x} [semiring R] (φ : ι → Type i) [Π (i : ι), add_comm_monoid (φ i)] [Π (i : ι), semimodule R (φ i)] [decidable_eq ι] (i : ι) (b : φ i) :

@[simp]
theorem linear_map.std_basis_same (R : Type u) {ι : Type x} [semiring R] (φ : ι → Type i) [Π (i : ι), add_comm_monoid (φ i)] [Π (i : ι), semimodule R (φ i)] [decidable_eq ι] (i : ι) (b : φ i) :
(linear_map.std_basis R φ i) b i = b

theorem linear_map.std_basis_ne (R : Type u) {ι : Type x} [semiring R] (φ : ι → Type i) [Π (i : ι), add_comm_monoid (φ i)] [Π (i : ι), semimodule R (φ i)] [decidable_eq ι] (i j : ι) (h : j i) (b : φ i) :
(linear_map.std_basis R φ i) b j = 0

theorem linear_map.ker_std_basis (R : Type u) {ι : Type x} [semiring R] (φ : ι → Type i) [Π (i : ι), add_comm_monoid (φ i)] [Π (i : ι), semimodule R (φ i)] [decidable_eq ι] (i : ι) :

theorem linear_map.proj_comp_std_basis (R : Type u) {ι : Type x} [semiring R] (φ : ι → Type i) [Π (i : ι), add_comm_monoid (φ i)] [Π (i : ι), semimodule R (φ i)] [decidable_eq ι] (i j : ι) :

theorem linear_map.proj_std_basis_same (R : Type u) {ι : Type x} [semiring R] (φ : ι → Type i) [Π (i : ι), add_comm_monoid (φ i)] [Π (i : ι), semimodule R (φ i)] [decidable_eq ι] (i : ι) :

theorem linear_map.proj_std_basis_ne (R : Type u) {ι : Type x} [semiring R] (φ : ι → Type i) [Π (i : ι), add_comm_monoid (φ i)] [Π (i : ι), semimodule R (φ i)] [decidable_eq ι] (i j : ι) (h : i j) :

theorem linear_map.supr_range_std_basis_le_infi_ker_proj (R : Type u) {ι : Type x} [semiring R] (φ : ι → Type i) [Π (i : ι), add_comm_monoid (φ i)] [Π (i : ι), semimodule R (φ i)] [decidable_eq ι] (I J : set ι) (h : disjoint I J) :
(⨆ (i : ι) (H : i I), (linear_map.std_basis R φ i).range) ⨅ (i : ι) (H : i J), (linear_map.proj i).ker

theorem linear_map.infi_ker_proj_le_supr_range_std_basis (R : Type u) {ι : Type x} [semiring R] (φ : ι → Type i) [Π (i : ι), add_comm_monoid (φ i)] [Π (i : ι), semimodule R (φ i)] [decidable_eq ι] {I : finset ι} {J : set ι} (hu : set.univ I J) :
(⨅ (i : ι) (H : i J), (linear_map.proj i).ker) ⨆ (i : ι) (H : i I), (linear_map.std_basis R φ i).range

theorem linear_map.supr_range_std_basis_eq_infi_ker_proj (R : Type u) {ι : Type x} [semiring R] (φ : ι → Type i) [Π (i : ι), add_comm_monoid (φ i)] [Π (i : ι), semimodule R (φ i)] [decidable_eq ι] {I J : set ι} (hd : disjoint I J) (hu : set.univ I J) (hI : I.finite) :
(⨆ (i : ι) (H : i I), (linear_map.std_basis R φ i).range) = ⨅ (i : ι) (H : i J), (linear_map.proj i).ker

theorem linear_map.supr_range_std_basis (R : Type u) {ι : Type x} [semiring R] (φ : ι → Type i) [Π (i : ι), add_comm_monoid (φ i)] [Π (i : ι), semimodule R (φ i)] [decidable_eq ι] [fintype ι] :
(⨆ (i : ι), (linear_map.std_basis R φ i).range) =

theorem linear_map.disjoint_std_basis_std_basis (R : Type u) {ι : Type x} [semiring R] (φ : ι → Type i) [Π (i : ι), add_comm_monoid (φ i)] [Π (i : ι), semimodule R (φ i)] [decidable_eq ι] (I J : set ι) (h : disjoint I J) :
disjoint (⨆ (i : ι) (H : i I), (linear_map.std_basis R φ i).range) (⨆ (i : ι) (H : i J), (linear_map.std_basis R φ i).range)

theorem linear_map.std_basis_eq_single (R : Type u) {ι : Type x} [semiring R] [decidable_eq ι] {a : R} :
(λ (i : ι), (linear_map.std_basis R (λ (_x : ι), R) i) a) = λ (i : ι), (finsupp.single i a)

def linear_map.fun_left (R : Type u) (M : Type v) [semiring R] [add_comm_monoid M] [semimodule R M] {m : Type u_1} {n : Type u_2} (f : m → n) :
(n → M) →ₗ[R] m → M

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

Equations
@[simp]
theorem linear_map.fun_left_apply (R : Type u) (M : Type v) [semiring R] [add_comm_monoid M] [semimodule R M] {m : Type u_1} {n : Type u_2} (f : m → n) (g : n → M) (i : m) :
(linear_map.fun_left R M f) g i = g (f i)

@[simp]
theorem linear_map.fun_left_id (R : Type u) (M : Type v) [semiring R] [add_comm_monoid M] [semimodule R M] {n : Type u_2} (g : n → M) :

theorem linear_map.fun_left_comp (R : Type u) (M : Type v) [semiring R] [add_comm_monoid M] [semimodule R M] {m : Type u_1} {n : Type u_2} {p : Type u_3} (f₁ : n → p) (f₂ : m → n) :

def linear_map.fun_congr_left (R : Type u) (M : Type v) [semiring R] [add_comm_monoid M] [semimodule R M] {m : Type u_1} {n : Type u_2} (e : m n) :
(n → M) ≃ₗ[R] m → M

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

Equations
@[simp]
theorem linear_map.fun_congr_left_apply (R : Type u) (M : Type v) [semiring R] [add_comm_monoid M] [semimodule R M] {m : Type u_1} {n : Type u_2} (e : m n) (x : n → M) :

@[simp]
theorem linear_map.fun_congr_left_id (R : Type u) (M : Type v) [semiring R] [add_comm_monoid M] [semimodule R M] {n : Type u_2} :

@[simp]
theorem linear_map.fun_congr_left_comp (R : Type u) (M : Type v) [semiring R] [add_comm_monoid M] [semimodule R M] {m : Type u_1} {n : Type u_2} {p : Type u_3} (e₁ : m n) (e₂ : n p) :

@[simp]
theorem linear_map.fun_congr_left_symm (R : Type u) (M : Type v) [semiring R] [add_comm_monoid M] [semimodule R M] {m : Type u_1} {n : Type u_2} (e : m n) :

@[instance]
def linear_map.automorphism_group (R : Type u) (M : Type v) [semiring R] [add_comm_monoid M] [semimodule R M] :

Equations
def linear_map.general_linear_group (R : Type u) (M : Type v) [semiring R] [add_comm_monoid M] [semimodule R M] :
Type v

The group of invertible linear maps from M to itself

Equations

An invertible linear map f determines an equivalence from M to itself.

Equations

An equivalence from M to itself determines an invertible linear map.

Equations

The general linear group on R and M is multiplicatively equivalent to the type of linear equivalences between M and itself.

Equations