mathlib documentation

linear_algebra.dimension

Dimension of modules and vector spaces #

Main definitions #

Main statements #

For modules over rings satisfying the rank condition

For modules over rings satisfying the strong rank condition

For modules over rings with invariant basis number (including all commutative rings and all noetherian rings)

For vector spaces (i.e. modules over a field), we have

Implementation notes #

There is a naming discrepancy: most of the theorem names refer to dim, even though the definition is of module.rank. This reflects that module.rank was originally called dim, and only defined for vector spaces.

Many theorems in this file are not universe-generic when they relate dimensions in different universes. They should be as general as they can be without inserting lifts. The types V, V', ... all live in different universes, and V₁, V₂, ... all live in the same universe.

@[protected]
noncomputable def module.rank (K : Type u) (V : Type v) [semiring K] [add_comm_monoid V] [module K V] :

The rank of a module, defined as a term of type cardinal.

We define this as the supremum of the cardinalities of linearly independent subsets.

For a free module over any ring satisfying the strong rank condition (e.g. left-noetherian rings, commutative rings, and in particular division rings and fields), this is the same as the dimension of the space (i.e. the cardinality of any basis).

In particular this agrees with the usual notion of the dimension of a vector space.

The definition is marked as protected to avoid conflicts with _root_.rank, the rank of a linear map.

Equations
theorem linear_map.lift_dim_le_of_injective {R : Type u} [ring R] {M : Type v} [add_comm_group M] [module R M] {M' : Type v'} [add_comm_group M'] [module R M'] (f : M →ₗ[R] M') (i : function.injective f) :
theorem linear_map.dim_le_of_injective {R : Type u} [ring R] {M : Type v} [add_comm_group M] [module R M] {M₁ : Type v} [add_comm_group M₁] [module R M₁] (f : M →ₗ[R] M₁) (i : function.injective f) :
theorem dim_le {R : Type u} [ring R] {M : Type v} [add_comm_group M] [module R M] {n : } (H : ∀ (s : finset M), linear_independent R (λ (i : s), i)s.card n) :
theorem lift_dim_range_le {R : Type u} [ring R] {M : Type v} [add_comm_group M] [module R M] {M' : Type v'} [add_comm_group M'] [module R M'] (f : M →ₗ[R] M') :
theorem dim_range_le {R : Type u} [ring R] {M : Type v} [add_comm_group M] [module R M] {M₁ : Type v} [add_comm_group M₁] [module R M₁] (f : M →ₗ[R] M₁) :
theorem lift_dim_map_le {R : Type u} [ring R] {M : Type v} [add_comm_group M] [module R M] {M' : Type v'} [add_comm_group M'] [module R M'] (f : M →ₗ[R] M') (p : submodule R M) :
theorem dim_map_le {R : Type u} [ring R] {M : Type v} [add_comm_group M] [module R M] {M₁ : Type v} [add_comm_group M₁] [module R M₁] (f : M →ₗ[R] M₁) (p : submodule R M) :
theorem dim_le_of_submodule {R : Type u} [ring R] {M : Type v} [add_comm_group M] [module R M] (s t : submodule R M) (h : s t) :
theorem linear_equiv.lift_dim_eq {R : Type u} [ring R] {M : Type v} [add_comm_group M] [module R M] {M' : Type v'} [add_comm_group M'] [module R M'] (f : M ≃ₗ[R] M') :

Two linearly equivalent vector spaces have the same dimension, a version with different universes.

theorem linear_equiv.dim_eq {R : Type u} [ring R] {M : Type v} [add_comm_group M] [module R M] {M₁ : Type v} [add_comm_group M₁] [module R M₁] (f : M ≃ₗ[R] M₁) :

Two linearly equivalent vector spaces have the same dimension.

theorem dim_eq_of_injective {R : Type u} [ring R] {M : Type v} [add_comm_group M] [module R M] {M₁ : Type v} [add_comm_group M₁] [module R M₁] (f : M →ₗ[R] M₁) (h : function.injective f) :
theorem linear_equiv.dim_map_eq {R : Type u} [ring R] {M : Type v} [add_comm_group M] [module R M] {M₁ : Type v} [add_comm_group M₁] [module R M₁] (f : M ≃ₗ[R] M₁) (p : submodule R M) :

Pushforwards of submodules along a linear_equiv have the same dimension.

@[simp]
theorem dim_top (R : Type u) [ring R] (M : Type v) [add_comm_group M] [module R M] :
theorem dim_range_of_surjective {R : Type u} [ring R] {M : Type v} [add_comm_group M] [module R M] {M' : Type v'} [add_comm_group M'] [module R M'] (f : M →ₗ[R] M') (h : function.surjective f) :
theorem dim_submodule_le {R : Type u} [ring R] {M : Type v} [add_comm_group M] [module R M] (s : submodule R M) :
theorem linear_map.dim_le_of_surjective {R : Type u} [ring R] {M : Type v} [add_comm_group M] [module R M] {M₁ : Type v} [add_comm_group M₁] [module R M₁] (f : M →ₗ[R] M₁) (h : function.surjective f) :
theorem dim_quotient_le {R : Type u} [ring R] {M : Type v} [add_comm_group M] [module R M] (p : submodule R M) :
theorem cardinal_lift_le_dim_of_linear_independent {R : Type u} [ring R] {M : Type v} [add_comm_group M] [module R M] [nontrivial R] {ι : Type w} {v : ι → M} (hv : linear_independent R v) :
theorem cardinal_lift_le_dim_of_linear_independent' {R : Type u} [ring R] {M : Type v} [add_comm_group M] [module R M] [nontrivial R] {ι : Type w} {v : ι → M} (hv : linear_independent R v) :
theorem cardinal_le_dim_of_linear_independent {R : Type u} [ring R] {M : Type v} [add_comm_group M] [module R M] [nontrivial R] {ι : Type v} {v : ι → M} (hv : linear_independent R v) :
theorem cardinal_le_dim_of_linear_independent' {R : Type u} [ring R] {M : Type v} [add_comm_group M] [module R M] [nontrivial R] {s : set M} (hs : linear_independent R (λ (x : s), x)) :
@[simp]
theorem dim_punit (R : Type u) [ring R] [nontrivial R] :
@[simp]
theorem dim_bot (R : Type u) [ring R] (M : Type v) [add_comm_group M] [module R M] [nontrivial R] :
noncomputable def basis_fintype_of_finite_spans {R : Type u} [ring R] {M : Type v} [add_comm_group M] [module R M] [nontrivial R] (w : set M) [fintype w] (s : submodule.span R w = ) {ι : Type w} (b : basis ι R M) :

Over any nontrivial ring, the existence of a finite spanning set implies that any basis is finite.

Equations
theorem union_support_maximal_linear_independent_eq_range_basis {R : Type u} [ring R] {M : Type v} [add_comm_group M] [module R M] [nontrivial R] {ι : Type w} (b : basis ι R M) {κ : Type w'} (v : κ → M) (i : linear_independent R v) (m : i.maximal) :
(⋃ (k : κ), (((b.repr) (v k)).support)) = set.univ

Over any ring R, if b is a basis for a module M, and s is a maximal linearly independent set, then the union of the supports of x ∈ s (when written out in the basis b) is all of b.

theorem infinite_basis_le_maximal_linear_independent' {R : Type u} [ring R] {M : Type v} [add_comm_group M] [module R M] [nontrivial R] {ι : Type w} (b : basis ι R M) [infinite ι] {κ : Type w'} (v : κ → M) (i : linear_independent R v) (m : i.maximal) :
(# ι).lift (# κ).lift

Over any ring R, if b is an infinite basis for a module M, and s is a maximal linearly independent set, then the cardinality of b is bounded by the cardinality of s.

theorem infinite_basis_le_maximal_linear_independent {R : Type u} [ring R] {M : Type v} [add_comm_group M] [module R M] [nontrivial R] {ι : Type w} (b : basis ι R M) [infinite ι] {κ : Type w} (v : κ → M) (i : linear_independent R v) (m : i.maximal) :
# ι # κ

Over any ring R, if b is an infinite basis for a module M, and s is a maximal linearly independent set, then the cardinality of b is bounded by the cardinality of s.

theorem complete_lattice.independent.subtype_ne_bot_le_rank {ι : Type w} {R : Type u} [ring R] {M : Type v} [add_comm_group M] [module R M] [nontrivial R] [no_zero_smul_divisors R M] {V : ι → submodule R M} (hV : complete_lattice.independent V) :
(# {i // V i }).lift (module.rank R M).lift
theorem dim_zero_iff_forall_zero {R : Type u} {M : Type v} [ring R] [nontrivial R] [add_comm_group M] [module R M] [no_zero_smul_divisors R M] :
module.rank R M = 0 ∀ (x : M), x = 0
theorem dim_zero_iff {R : Type u} {M : Type v} [ring R] [nontrivial R] [add_comm_group M] [module R M] [no_zero_smul_divisors R M] :
theorem dim_pos_iff_exists_ne_zero {R : Type u} {M : Type v} [ring R] [nontrivial R] [add_comm_group M] [module R M] [no_zero_smul_divisors R M] :
0 < module.rank R M ∃ (x : M), x 0
theorem dim_pos_iff_nontrivial {R : Type u} {M : Type v} [ring R] [nontrivial R] [add_comm_group M] [module R M] [no_zero_smul_divisors R M] :
theorem dim_pos {R : Type u} {M : Type v} [ring R] [nontrivial R] [add_comm_group M] [module R M] [no_zero_smul_divisors R M] [h : nontrivial M] :
theorem mk_eq_mk_of_basis {ι : Type w} {ι' : Type w'} {R : Type u} [ring R] [invariant_basis_number R] {M : Type v} [add_comm_group M] [module R M] (v : basis ι R M) (v' : basis ι' R M) :
(# ι).lift = (# ι').lift

The dimension theorem: if v and v' are two bases, their index types have the same cardinalities.

noncomputable def basis.index_equiv {ι : Type w} {ι' : Type w'} {R : Type u} [ring R] [invariant_basis_number R] {M : Type v} [add_comm_group M] [module R M] (v : basis ι R M) (v' : basis ι' R M) :
ι ι'

Given two basis indexed by ι and ι' of an R-module, where R satisfies the invariant basis number property, an equiv ι ≃ ι'.

Equations
theorem mk_eq_mk_of_basis' {ι : Type w} {R : Type u} [ring R] [invariant_basis_number R] {M : Type v} [add_comm_group M] [module R M] {ι' : Type w} (v : basis ι R M) (v' : basis ι' R M) :
# ι = # ι'
theorem basis.le_span'' {R : Type u} [ring R] [rank_condition R] {M : Type v} [add_comm_group M] [module R M] {ι : Type u_1} [fintype ι] (b : basis ι R M) {w : set M} [fintype w] (s : submodule.span R w = ) :

An auxiliary lemma for basis.le_span.

If R satisfies the rank condition, then for any finite basis b : basis ι R M, and any finite spanning set w : set M, the cardinality of ι is bounded by the cardinality of w.

theorem basis_le_span' {R : Type u} [ring R] [rank_condition R] {M : Type v} [add_comm_group M] [module R M] {ι : Type u_1} (b : basis ι R M) {w : set M} [fintype w] (s : submodule.span R w = ) :

Another auxiliary lemma for basis.le_span, which does not require assuming the basis is finite, but still assumes we have a finite spanning set.

theorem basis.le_span {ι : Type w} {R : Type u} [ring R] [rank_condition R] {M : Type v} [add_comm_group M] [module R M] {J : set M} (v : basis ι R M) (hJ : submodule.span R J = ) :

If R satisfies the rank condition, then the cardinality of any basis is bounded by the cardinality of any spanning set.

theorem linear_independent_le_span_aux' {R : Type u} [ring R] [strong_rank_condition R] {M : Type v} [add_comm_group M] [module R M] {ι : Type u_1} [fintype ι] (v : ι → M) (i : linear_independent R v) (w : set M) [fintype w] (s : set.range v (submodule.span R w)) :
noncomputable def linear_independent_fintype_of_le_span_fintype {R : Type u} [ring R] [strong_rank_condition R] {M : Type v} [add_comm_group M] [module R M] {ι : Type u_1} (v : ι → M) (i : linear_independent R v) (w : set M) [fintype w] (s : set.range v (submodule.span R w)) :

If R satisfies the strong rank condition, then any linearly independent family v : ι → M contained in the span of some finite w : set M, is itself finite.

Equations
theorem linear_independent_le_span' {R : Type u} [ring R] [strong_rank_condition R] {M : Type v} [add_comm_group M] [module R M] {ι : Type u_1} (v : ι → M) (i : linear_independent R v) (w : set M) [fintype w] (s : set.range v (submodule.span R w)) :

If R satisfies the strong rank condition, then for any linearly independent family v : ι → M contained in the span of some finite w : set M, the cardinality of ι is bounded by the cardinality of w.

theorem linear_independent_le_span {R : Type u} [ring R] [strong_rank_condition R] {M : Type v} [add_comm_group M] [module R M] {ι : Type u_1} (v : ι → M) (i : linear_independent R v) (w : set M) [fintype w] (s : submodule.span R w = ) :

If R satisfies the strong rank condition, then for any linearly independent family v : ι → M and any finite spanning set w : set M, the cardinality of ι is bounded by the cardinality of w.

noncomputable def fintype_of_is_noetherian_linear_independent {ι : Type w} {R : Type u} [ring R] [strong_rank_condition R] {M : Type v} [add_comm_group M] [module R M] [is_noetherian R M] {v : ι → M} (hi : linear_independent R v) :

A linearly-independent family of vectors in a module over a ring satisfying the strong rank condition must be finite if the module is Noetherian.

Equations
theorem finite_of_is_noetherian_linear_independent {R : Type u} [ring R] [strong_rank_condition R] {M : Type v} [add_comm_group M] [module R M] [is_noetherian R M] {s : set M} (hi : linear_independent R coe) :

A linearly-independent subset of a module over a ring satisfying the strong rank condition must be finite if the module is Noetherian.

theorem linear_independent_le_infinite_basis {R : Type u} [ring R] [strong_rank_condition R] {M : Type v} [add_comm_group M] [module R M] {ι : Type u_1} (b : basis ι R M) [infinite ι] {κ : Type u_1} (v : κ → M) (i : linear_independent R v) :
# κ # ι

An auxiliary lemma for linear_independent_le_basis: we handle the case where the basis b is infinite.

theorem linear_independent_le_basis {R : Type u} [ring R] [strong_rank_condition R] {M : Type v} [add_comm_group M] [module R M] {ι : Type u_1} (b : basis ι R M) {κ : Type u_1} (v : κ → M) (i : linear_independent R v) :
# κ # ι

Over any ring R satisfying the strong rank condition, if b is a basis for a module M, and s is a linearly independent set, then the cardinality of s is bounded by the cardinality of b.

theorem basis.card_le_card_of_linear_independent_aux {R : Type u_1} [ring R] [strong_rank_condition R] (n : ) {m : } (v : fin mfin n → R) :

In an n-dimensional space, the rank is at most m.

theorem maximal_linear_independent_eq_infinite_basis {R : Type u} [ring R] [strong_rank_condition R] {M : Type v} [add_comm_group M] [module R M] {ι : Type u_1} (b : basis ι R M) [infinite ι] {κ : Type u_1} (v : κ → M) (i : linear_independent R v) (m : i.maximal) :
# κ = # ι

Over any ring R satisfying the strong rank condition, if b is an infinite basis for a module M, then every maximal linearly independent set has the same cardinality as b.

This proof (along with some of the lemmas above) comes from Les familles libres maximales d'un module ont-elles le meme cardinal?

theorem basis.mk_eq_dim'' {R : Type u} [ring R] [strong_rank_condition R] {M : Type v} [add_comm_group M] [module R M] {ι : Type v} (v : basis ι R M) :
# ι = module.rank R M
theorem basis.mk_range_eq_dim {ι : Type w} {R : Type u} [ring R] [strong_rank_condition R] {M : Type v} [add_comm_group M] [module R M] (v : basis ι R M) :
theorem dim_eq_card_basis {R : Type u} [ring R] [strong_rank_condition R] {M : Type v} [add_comm_group M] [module R M] {ι : Type w} [fintype ι] (h : basis ι R M) :

If a vector space has a finite basis, then its dimension (seen as a cardinal) is equal to the cardinality of the basis.

theorem basis.card_le_card_of_linear_independent {R : Type u} [ring R] [strong_rank_condition R] {M : Type v} [add_comm_group M] [module R M] {ι : Type u_1} [fintype ι] (b : basis ι R M) {ι' : Type u_2} [fintype ι'] {v : ι' → M} (hv : linear_independent R v) :
theorem basis.card_le_card_of_submodule {ι : Type w} {ι' : Type w'} {R : Type u} [ring R] [strong_rank_condition R] {M : Type v} [add_comm_group M] [module R M] (N : submodule R M) [fintype ι] (b : basis ι R M) [fintype ι'] (b' : basis ι' R N) :
theorem basis.card_le_card_of_le {ι : Type w} {ι' : Type w'} {R : Type u} [ring R] [strong_rank_condition R] {M : Type v} [add_comm_group M] [module R M] {N O : submodule R M} (hNO : N O) [fintype ι] (b : basis ι R O) [fintype ι'] (b' : basis ι' R N) :
theorem basis.mk_eq_dim {ι : Type w} {R : Type u} [ring R] [strong_rank_condition R] {M : Type v} [add_comm_group M] [module R M] (v : basis ι R M) :
(# ι).lift = (module.rank R M).lift
theorem basis.mk_eq_dim' {ι : Type w} {R : Type u} [ring R] [strong_rank_condition R] {M : Type v} [add_comm_group M] [module R M] (v : basis ι R M) :
(# ι).lift = (module.rank R M).lift
theorem basis.nonempty_fintype_index_of_dim_lt_omega {R : Type u} [ring R] [strong_rank_condition R] {M : Type v} [add_comm_group M] [module R M] {ι : Type u_1} (b : basis ι R M) (h : module.rank R M < ω) :

If a module has a finite dimension, all bases are indexed by a finite type.

noncomputable def basis.fintype_index_of_dim_lt_omega {R : Type u} [ring R] [strong_rank_condition R] {M : Type v} [add_comm_group M] [module R M] {ι : Type u_1} (b : basis ι R M) (h : module.rank R M < ω) :

If a module has a finite dimension, all bases are indexed by a finite type.

Equations
theorem basis.finite_index_of_dim_lt_omega {R : Type u} [ring R] [strong_rank_condition R] {M : Type v} [add_comm_group M] [module R M] {ι : Type u_1} {s : set ι} (b : basis s R M) (h : module.rank R M < ω) :

If a module has a finite dimension, all bases are indexed by a finite set.

theorem dim_span {ι : Type w} {R : Type u} [ring R] [strong_rank_condition R] {M : Type v} [add_comm_group M] [module R M] {v : ι → M} (hv : linear_independent R v) :
theorem dim_span_set {R : Type u} [ring R] [strong_rank_condition R] {M : Type v} [add_comm_group M] [module R M] {s : set M} (hs : linear_independent R (λ (x : s), x)) :
def submodule.induction_on_rank {ι : Type w} {R : Type u} [ring R] [strong_rank_condition R] {M : Type v} [add_comm_group M] [module R M] [is_domain R] [fintype ι] (b : basis ι R M) (P : submodule R MSort u_1) (ih : Π (N : submodule R M), (Π (N' : submodule R M), N' NΠ (x : M), x N(∀ (c : R) (y : M), y N'c x + y = 0c = 0)P N')P N) (N : submodule R M) :
P N

If N is a submodule in a free, finitely generated module, do induction on adjoining a linear independent element to a submodule.

Equations
theorem ideal.rank_eq {R : Type u_1} {S : Type u_2} [comm_ring R] [strong_rank_condition R] [ring S] [is_domain S] [algebra R S] {n : Type u_3} {m : Type u_4} [fintype n] [fintype m] (b : basis n R S) {I : ideal S} (hI : I ) (c : basis m R I) :

If S a finite-dimensional ring extension of R which is free as an R-module, then the rank of an ideal I of S over R is the same as the rank of S.

@[simp]
theorem dim_self (R : Type u) [ring R] [strong_rank_condition R] :

If a vector space has a finite dimension, the index set of basis.of_vector_space is finite.

theorem nonempty_linear_equiv_of_lift_dim_eq {K : Type u} {V : Type v} {V' : Type v'} [division_ring K] [add_comm_group V] [module K V] [add_comm_group V'] [module K V'] (cond : (module.rank K V).lift = (module.rank K V').lift) :

Two vector spaces are isomorphic if they have the same dimension.

theorem nonempty_linear_equiv_of_dim_eq {K : Type u} {V V₁ : Type v} [division_ring K] [add_comm_group V] [module K V] [add_comm_group V₁] [module K V₁] (cond : module.rank K V = module.rank K V₁) :
nonempty (V ≃ₗ[K] V₁)

Two vector spaces are isomorphic if they have the same dimension.

noncomputable def linear_equiv.of_lift_dim_eq {K : Type u} (V : Type v) (V' : Type v') [division_ring K] [add_comm_group V] [module K V] [add_comm_group V'] [module K V'] (cond : (module.rank K V).lift = (module.rank K V').lift) :
V ≃ₗ[K] V'

Two vector spaces are isomorphic if they have the same dimension.

Equations
noncomputable def linear_equiv.of_dim_eq {K : Type u} (V V₁ : Type v) [division_ring K] [add_comm_group V] [module K V] [add_comm_group V₁] [module K V₁] (cond : module.rank K V = module.rank K V₁) :
V ≃ₗ[K] V₁

Two vector spaces are isomorphic if they have the same dimension.

Equations
theorem linear_equiv.nonempty_equiv_iff_lift_dim_eq {K : Type u} {V : Type v} {V' : Type v'} [division_ring K] [add_comm_group V] [module K V] [add_comm_group V'] [module K V'] :

Two vector spaces are isomorphic if and only if they have the same dimension.

theorem linear_equiv.nonempty_equiv_iff_dim_eq {K : Type u} {V V₁ : Type v} [division_ring K] [add_comm_group V] [module K V] [add_comm_group V₁] [module K V₁] :

Two vector spaces are isomorphic if and only if they have the same dimension.

theorem dim_span_le {K : Type u} {V : Type v} [division_ring K] [add_comm_group V] [module K V] (s : set V) :
theorem dim_span_of_finset {K : Type u} {V : Type v} [division_ring K] [add_comm_group V] [module K V] (s : finset V) :
theorem dim_prod {K : Type u} {V V₁ : Type v} [division_ring K] [add_comm_group V] [module K V] [add_comm_group V₁] [module K V₁] :
module.rank K (V × V₁) = module.rank K V + module.rank K V₁
theorem dim_pi {K : Type u} {η : Type u₁'} {φ : η → Type u_1} [division_ring K] [fintype η] [Π (i : η), add_comm_group (φ i)] [Π (i : η), module K (φ i)] :
module.rank K (Π (i : η), φ i) = cardinal.sum (λ (i : η), module.rank K (φ i))
theorem dim_fun {K : Type u} [division_ring K] {V η : Type u} [fintype η] [add_comm_group V] [module K V] :
module.rank K (η → V) = ((fintype.card η)) * module.rank K V
theorem dim_fun_eq_lift_mul {K : Type u} {V : Type v} {η : Type u₁'} [division_ring K] [add_comm_group V] [module K V] [fintype η] :
module.rank K (η → V) = ((fintype.card η)) * (module.rank K V).lift
theorem dim_fun' {K : Type u} {η : Type u₁'} [division_ring K] [fintype η] :
module.rank K (η → K) = (fintype.card η)
theorem dim_fin_fun {K : Type u} [division_ring K] (n : ) :
module.rank K (fin n → K) = n
theorem dim_quotient_add_dim {K : Type u} {V : Type v} [field K] [add_comm_group V] [module K V] (p : submodule K V) :
theorem dim_range_add_dim_ker {K : Type u} {V V₁ : Type v} [field K] [add_comm_group V] [module K V] [add_comm_group V₁] [module K V₁] (f : V →ₗ[K] V₁) :

rank-nullity theorem

theorem dim_eq_of_surjective {K : Type u} {V V₁ : Type v} [field K] [add_comm_group V] [module K V] [add_comm_group V₁] [module K V₁] (f : V →ₗ[K] V₁) (h : function.surjective f) :
theorem dim_add_dim_split {K : Type u} {V V₁ V₂ V₃ : Type v} [field K] [add_comm_group V] [module K V] [add_comm_group V₁] [module K V₁] [add_comm_group V₂] [module K V₂] [add_comm_group V₃] [module K V₃] (db : V₂ →ₗ[K] V) (eb : V₃ →ₗ[K] V) (cd : V₁ →ₗ[K] V₂) (ce : V₁ →ₗ[K] V₃) (hde : db.range eb.range) (hgd : cd.ker = ) (eq : db.comp cd = eb.comp ce) (eq₂ : ∀ (d : V₂) (e : V₃), db d = eb e(∃ (c : V₁), cd c = d ce c = e)) :

This is mostly an auxiliary lemma for dim_sup_add_dim_inf_eq.

theorem dim_sup_add_dim_inf_eq {K : Type u} {V : Type v} [field K] [add_comm_group V] [module K V] (s t : submodule K V) :
theorem dim_add_le_dim_add_dim {K : Type u} {V : Type v} [field K] [add_comm_group V] [module K V] (s t : submodule K V) :
theorem exists_mem_ne_zero_of_dim_pos {K : Type u} {V : Type v} [field K] [add_comm_group V] [module K V] {s : submodule K V} (h : 0 < module.rank K s) :
∃ (b : V), b s b 0
noncomputable def rank {K : Type u} {V : Type v} {V' : Type v'} [field K] [add_comm_group V] [module K V] [add_comm_group V'] [module K V'] (f : V →ₗ[K] V') :

rank f is the rank of a linear_map f, defined as the dimension of f.range.

Equations
theorem rank_le_domain {K : Type u} {V V₁ : Type v} [field K] [add_comm_group V] [module K V] [add_comm_group V₁] [module K V₁] (f : V →ₗ[K] V₁) :
theorem rank_le_range {K : Type u} {V V₁ : Type v} [field K] [add_comm_group V] [module K V] [add_comm_group V₁] [module K V₁] (f : V →ₗ[K] V₁) :
theorem rank_add_le {K : Type u} {V : Type v} {V' : Type v'} [field K] [add_comm_group V] [module K V] [add_comm_group V'] [module K V'] (f g : V →ₗ[K] V') :
rank (f + g) rank f + rank g
@[simp]
theorem rank_zero {K : Type u} {V : Type v} {V' : Type v'} [field K] [add_comm_group V] [module K V] [add_comm_group V'] [module K V'] :
rank 0 = 0
theorem rank_finset_sum_le {K : Type u} {V : Type v} {V' : Type v'} [field K] [add_comm_group V] [module K V] [add_comm_group V'] [module K V'] {η : Type u_1} (s : finset η) (f : η → (V →ₗ[K] V')) :
rank (∑ (d : η) in s, f d) ∑ (d : η) in s, rank (f d)
theorem rank_comp_le1 {K : Type u} {V : Type v} {V' : Type v'} {V'' : Type v''} [field K] [add_comm_group V] [module K V] [add_comm_group V'] [module K V'] [add_comm_group V''] [module K V''] (g : V →ₗ[K] V') (f : V' →ₗ[K] V'') :
rank (f.comp g) rank f
theorem rank_comp_le2 {K : Type u} {V : Type v} {V' V'₁ : Type v'} [field K] [add_comm_group V] [module K V] [add_comm_group V'] [module K V'] [add_comm_group V'₁] [module K V'₁] (g : V →ₗ[K] V') (f : V' →ₗ[K] V'₁) :
rank (f.comp g) rank g
noncomputable def basis.of_dim_eq_zero {K : Type u} {V : Type v} [field K] [add_comm_group V] [module K V] {ι : Type u_1} [is_empty ι] (hV : module.rank K V = 0) :
basis ι K V

The ι indexed basis on V, where ι is an empty type and V is zero-dimensional.

See also finite_dimensional.fin_basis.

Equations
@[simp]
theorem basis.of_dim_eq_zero_apply {K : Type u} {V : Type v} [field K] [add_comm_group V] [module K V] {ι : Type u_1} [is_empty ι] (hV : module.rank K V = 0) (i : ι) :
theorem le_dim_iff_exists_linear_independent {K : Type u} {V : Type v} [field K] [add_comm_group V] [module K V] {c : cardinal} :
c module.rank K V ∃ (s : set V), # s = c linear_independent K coe
theorem le_dim_iff_exists_linear_independent_finset {K : Type u} {V : Type v} [field K] [add_comm_group V] [module K V] {n : } :
theorem le_rank_iff_exists_linear_independent {K : Type u} {V : Type v} {V' : Type v'} [field K] [add_comm_group V] [module K V] [add_comm_group V'] [module K V'] {c : cardinal} {f : V →ₗ[K] V'} :
c rank f ∃ (s : set V), (# s).lift = c.lift linear_independent K (λ (x : s), f x)
theorem le_rank_iff_exists_linear_independent_finset {K : Type u} {V : Type v} {V' : Type v'} [field K] [add_comm_group V] [module K V] [add_comm_group V'] [module K V'] {n : } {f : V →ₗ[K] V'} :
n rank f ∃ (s : finset V), s.card = n linear_independent K (λ (x : s), f x)
theorem dim_le_one_iff {K : Type u} {V : Type v} [field K] [add_comm_group V] [module K V] :
module.rank K V 1 ∃ (v₀ : V), ∀ (v : V), ∃ (r : K), r v₀ = v

A vector space has dimension at most 1 if and only if there is a single vector of which all vectors are multiples.

theorem dim_submodule_le_one_iff {K : Type u} {V : Type v} [field K] [add_comm_group V] [module K V] (s : submodule K V) :
module.rank K s 1 ∃ (v₀ : V) (H : v₀ s), s submodule.span K {v₀}

A submodule has dimension at most 1 if and only if there is a single vector in the submodule such that the submodule is contained in its span.

theorem dim_submodule_le_one_iff' {K : Type u} {V : Type v} [field K] [add_comm_group V] [module K V] (s : submodule K V) :
module.rank K s 1 ∃ (v₀ : V), s submodule.span K {v₀}

A submodule has dimension at most 1 if and only if there is a single vector, not necessarily in the submodule, such that the submodule is contained in its span.