Documentation

Mathlib.LinearAlgebra.LinearIndependent

Linear independence #

This file defines linear independence in a module or vector space.

It is inspired by Isabelle/HOL's linear algebra, and hence indirectly by HOL Light.

We define LinearIndependent R v as ker (Finsupp.total ι M R v) = ⊥. Here Finsupp.total is the linear map sending a function f : ι →₀ R with finite support to the linear combination of vectors from v with these coefficients. Then we prove that several other statements are equivalent to this one, including injectivity of Finsupp.total ι M R v and some versions with explicitly written linear combinations.

Main definitions #

All definitions are given for families of vectors, i.e. v : ι → M where M is the module or vector space and ι : Type* is an arbitrary indexing type.

Main statements #

We prove several specialized tests for linear independence of families of vectors and of sets of vectors.

In many cases we additionally provide dot-style operations (e.g., LinearIndependent.union) to make the linear independence tests usable as hv.insert ha etc.

We also prove that, when working over a division ring, any family of vectors includes a linear independent subfamily spanning the same subspace.

Implementation notes #

We use families instead of sets because it allows us to say that two identical vectors are linearly dependent.

If you want to use sets, use the family (fun x ↦ x : s → M) given a set s : Set M. The lemmas LinearIndependent.to_subtype_range and LinearIndependent.of_subtype_range connect those two worlds.

Tags #

linearly dependent, linear dependence, linearly independent, linear independence

def LinearIndependent {ι : Type u'} (R : Type u_2) {M : Type u_4} (v : ιM) [Semiring R] [AddCommMonoid M] [Module R M] :

LinearIndependent R v states the family of vectors v is linearly independent over R.

Equations
Instances For

    Delaborator for LinearIndependent that suggests pretty printing with type hints in case the family of vectors is over a Set.

    Type hints look like LinearIndependent fun (v : ↑s) => ↑v or LinearIndependent (ι := ↑s) f, depending on whether the family is a lambda expression or not.

    Equations
    • One or more equations did not get rendered due to their size.
    Instances For
      theorem linearIndependent_iff {ι : Type u'} {R : Type u_2} {M : Type u_4} {v : ιM} [Semiring R] [AddCommMonoid M] [Module R M] :
      LinearIndependent R v ∀ (l : ι →₀ R), (Finsupp.total ι M R v) l = 0l = 0
      theorem linearIndependent_iff' {ι : Type u'} {R : Type u_2} {M : Type u_4} {v : ιM} [Semiring R] [AddCommMonoid M] [Module R M] :
      LinearIndependent R v ∀ (s : Finset ι) (g : ιR), (s.sum fun (i : ι) => g i v i) = 0is, g i = 0
      theorem linearIndependent_iff'' {ι : Type u'} {R : Type u_2} {M : Type u_4} {v : ιM} [Semiring R] [AddCommMonoid M] [Module R M] :
      LinearIndependent R v ∀ (s : Finset ι) (g : ιR), (is, g i = 0)(s.sum fun (i : ι) => g i v i) = 0∀ (i : ι), g i = 0
      theorem not_linearIndependent_iff {ι : Type u'} {R : Type u_2} {M : Type u_4} {v : ιM} [Semiring R] [AddCommMonoid M] [Module R M] :
      ¬LinearIndependent R v ∃ (s : Finset ι) (g : ιR), (s.sum fun (i : ι) => g i v i) = 0 is, g i 0
      theorem Fintype.linearIndependent_iff {ι : Type u'} {R : Type u_2} {M : Type u_4} {v : ιM} [Semiring R] [AddCommMonoid M] [Module R M] [Fintype ι] :
      LinearIndependent R v ∀ (g : ιR), (Finset.univ.sum fun (i : ι) => g i v i) = 0∀ (i : ι), g i = 0
      theorem Fintype.linearIndependent_iff' {ι : Type u'} {R : Type u_2} {M : Type u_4} {v : ιM} [Semiring R] [AddCommMonoid M] [Module R M] [Fintype ι] [DecidableEq ι] :
      LinearIndependent R v LinearMap.ker ((LinearMap.lsum R (fun (x : ι) => R) ) fun (i : ι) => LinearMap.id.smulRight (v i)) =

      A finite family of vectors v i is linear independent iff the linear map that sends c : ι → R to ∑ i, c i • v i has the trivial kernel.

      theorem Fintype.not_linearIndependent_iff {ι : Type u'} {R : Type u_2} {M : Type u_4} {v : ιM} [Semiring R] [AddCommMonoid M] [Module R M] [Fintype ι] :
      ¬LinearIndependent R v ∃ (g : ιR), (Finset.univ.sum fun (i : ι) => g i v i) = 0 ∃ (i : ι), g i 0
      theorem linearIndependent_empty_type {ι : Type u'} {R : Type u_2} {M : Type u_4} {v : ιM} [Semiring R] [AddCommMonoid M] [Module R M] [IsEmpty ι] :
      theorem LinearIndependent.ne_zero {ι : Type u'} {R : Type u_2} {M : Type u_4} {v : ιM} [Semiring R] [AddCommMonoid M] [Module R M] [Nontrivial R] (i : ι) (hv : LinearIndependent R v) :
      v i 0
      theorem LinearIndependent.eq_zero_of_pair {R : Type u_2} {M : Type u_4} [Semiring R] [AddCommMonoid M] [Module R M] {x : M} {y : M} (h : LinearIndependent R ![x, y]) {s : R} {t : R} (h' : s x + t y = 0) :
      s = 0 t = 0
      theorem LinearIndependent.pair_iff {R : Type u_2} {M : Type u_4} [Semiring R] [AddCommMonoid M] [Module R M] {x : M} {y : M} :
      LinearIndependent R ![x, y] ∀ (s t : R), s x + t y = 0s = 0 t = 0
      theorem LinearIndependent.comp {ι : Type u'} {ι' : Type u_1} {R : Type u_2} {M : Type u_4} {v : ιM} [Semiring R] [AddCommMonoid M] [Module R M] (h : LinearIndependent R v) (f : ι'ι) (hf : Function.Injective f) :

      A subfamily of a linearly independent family (i.e., a composition with an injective map) is a linearly independent family.

      theorem linearIndependent_iff_finset_linearIndependent {ι : Type u'} {R : Type u_2} {M : Type u_4} {v : ιM} [Semiring R] [AddCommMonoid M] [Module R M] :
      LinearIndependent R v ∀ (s : Finset ι), LinearIndependent R (v Subtype.val)

      A family is linearly independent if and only if all of its finite subfamily is linearly independent.

      theorem LinearIndependent.coe_range {ι : Type u'} {R : Type u_2} {M : Type u_4} {v : ιM} [Semiring R] [AddCommMonoid M] [Module R M] (i : LinearIndependent R v) :
      LinearIndependent (ι := { x : M // x Set.range v }) R Subtype.val
      theorem LinearIndependent.map {ι : Type u'} {R : Type u_2} {M : Type u_4} {M' : Type u_5} {v : ιM} [Semiring R] [AddCommMonoid M] [AddCommMonoid M'] [Module R M] [Module R M'] (hv : LinearIndependent R v) {f : M →ₗ[R] M'} (hf_inj : Disjoint (Submodule.span R (Set.range v)) (LinearMap.ker f)) :

      If v is a linearly independent family of vectors and the kernel of a linear map f is disjoint with the submodule spanned by the vectors of v, then f ∘ v is a linearly independent family of vectors. See also LinearIndependent.map' for a special case assuming ker f = ⊥.

      theorem Submodule.range_ker_disjoint {ι : Type u'} {R : Type u_2} {M : Type u_4} {M' : Type u_5} {v : ιM} [Semiring R] [AddCommMonoid M] [AddCommMonoid M'] [Module R M] [Module R M'] {f : M →ₗ[R] M'} (hv : LinearIndependent R (f v)) :

      If v is an injective family of vectors such that f ∘ v is linearly independent, then v spans a submodule disjoint from the kernel of f

      theorem LinearIndependent.map' {ι : Type u'} {R : Type u_2} {M : Type u_4} {M' : Type u_5} {v : ιM} [Semiring R] [AddCommMonoid M] [AddCommMonoid M'] [Module R M] [Module R M'] (hv : LinearIndependent R v) (f : M →ₗ[R] M') (hf_inj : LinearMap.ker f = ) :

      An injective linear map sends linearly independent families of vectors to linearly independent families of vectors. See also LinearIndependent.map for a more general statement.

      theorem LinearIndependent.map_of_injective_injective {ι : Type u'} {R : Type u_2} {M : Type u_4} {v : ιM} [Semiring R] [AddCommMonoid M] [Module R M] {R' : Type u_8} {M' : Type u_9} [Semiring R'] [AddCommMonoid M'] [Module R' M'] (hv : LinearIndependent R v) (i : R'R) (j : M →+ M') (hi : ∀ (r : R'), i r = 0r = 0) (hj : ∀ (m : M), j m = 0m = 0) (hc : ∀ (r : R') (m : M), j (i r m) = r j m) :

      If M / R and M' / R' are modules, i : R' → R is a map, j : M →+ M' is a monoid map, such that they send non-zero elements to non-zero elements, and compatible with the scalar multiplications on M and M', then j sends linearly independent families of vectors to linearly independent families of vectors. As a special case, taking R = R' it is LinearIndependent.map'.

      theorem LinearIndependent.map_of_surjective_injective {ι : Type u'} {R : Type u_2} {M : Type u_4} {v : ιM} [Semiring R] [AddCommMonoid M] [Module R M] {R' : Type u_8} {M' : Type u_9} [Semiring R'] [AddCommMonoid M'] [Module R' M'] (hv : LinearIndependent R v) (i : ZeroHom R R') (j : M →+ M') (hi : Function.Surjective i) (hj : ∀ (m : M), j m = 0m = 0) (hc : ∀ (r : R) (m : M), j (r m) = i r j m) :

      If M / R and M' / R' are modules, i : R → R' is a surjective map which maps zero to zero, j : M →+ M' is a monoid map which sends non-zero elements to non-zero elements, such that the scalar multiplications on M and M' are compatible, then j sends linearly independent families of vectors to linearly independent families of vectors. As a special case, taking R = R' it is LinearIndependent.map'.

      theorem LinearIndependent.of_comp {ι : Type u'} {R : Type u_2} {M : Type u_4} {M' : Type u_5} {v : ιM} [Semiring R] [AddCommMonoid M] [AddCommMonoid M'] [Module R M] [Module R M'] (f : M →ₗ[R] M') (hfv : LinearIndependent R (f v)) :

      If the image of a family of vectors under a linear map is linearly independent, then so is the original family.

      theorem LinearMap.linearIndependent_iff {ι : Type u'} {R : Type u_2} {M : Type u_4} {M' : Type u_5} {v : ιM} [Semiring R] [AddCommMonoid M] [AddCommMonoid M'] [Module R M] [Module R M'] (f : M →ₗ[R] M') (hf_inj : LinearMap.ker f = ) :

      If f is an injective linear map, then the family f ∘ v is linearly independent if and only if the family v is linearly independent.

      theorem linearIndependent_of_subsingleton {ι : Type u'} {R : Type u_2} {M : Type u_4} {v : ιM} [Semiring R] [AddCommMonoid M] [Module R M] [Subsingleton R] :
      theorem linearIndependent_equiv {ι : Type u'} {ι' : Type u_1} {R : Type u_2} {M : Type u_4} [Semiring R] [AddCommMonoid M] [Module R M] (e : ι ι') {f : ι'M} :
      theorem linearIndependent_equiv' {ι : Type u'} {ι' : Type u_1} {R : Type u_2} {M : Type u_4} [Semiring R] [AddCommMonoid M] [Module R M] (e : ι ι') {f : ι'M} {g : ιM} (h : f e = g) :
      theorem linearIndependent_subtype_range {R : Type u_2} {M : Type u_4} [Semiring R] [AddCommMonoid M] [Module R M] {ι : Type u_8} {f : ιM} (hf : Function.Injective f) :
      LinearIndependent (ι := { x : M // x Set.range f }) R Subtype.val LinearIndependent R f
      theorem LinearIndependent.of_subtype_range {R : Type u_2} {M : Type u_4} [Semiring R] [AddCommMonoid M] [Module R M] {ι : Type u_8} {f : ιM} (hf : Function.Injective f) :
      LinearIndependent (ι := { x : M // x Set.range f }) R Subtype.valLinearIndependent R f

      Alias of the forward direction of linearIndependent_subtype_range.

      theorem linearIndependent_image {R : Type u_2} {M : Type u_4} [Semiring R] [AddCommMonoid M] [Module R M] {ι : Type u_8} {s : Set ι} {f : ιM} (hf : Set.InjOn f s) :
      (LinearIndependent R fun (x : s) => f x) LinearIndependent R fun (x : (f '' s)) => x
      theorem linearIndependent_span {ι : Type u'} {R : Type u_2} {M : Type u_4} {v : ιM} [Semiring R] [AddCommMonoid M] [Module R M] (hs : LinearIndependent R v) :
      LinearIndependent R fun (i : ι) => v i,
      theorem LinearIndependent.fin_cons' {R : Type u_2} {M : Type u_4} [Semiring R] [AddCommMonoid M] [Module R M] {m : } (x : M) (v : Fin mM) (hli : LinearIndependent R v) (x_ortho : ∀ (c : R) (y : (Submodule.span R (Set.range v))), c x + y = 0c = 0) :

      See LinearIndependent.fin_cons for a family of elements in a vector space.

      theorem LinearIndependent.restrict_scalars {ι : Type u'} {R : Type u_2} {K : Type u_3} {M : Type u_4} {v : ιM} [Semiring R] [AddCommMonoid M] [Module R M] [Semiring K] [SMulWithZero R K] [Module K M] [IsScalarTower R K M] (hinj : Function.Injective fun (r : R) => r 1) (li : LinearIndependent K v) :

      A set of linearly independent vectors in a module M over a semiring K is also linearly independent over a subring R of K. The implementation uses minimal assumptions about the relationship between R, K and M. The version where K is an R-algebra is LinearIndependent.restrict_scalars_algebras.

      theorem linearIndependent_finset_map_embedding_subtype {R : Type u_2} {M : Type u_4} [Semiring R] [AddCommMonoid M] [Module R M] (s : Set M) (li : LinearIndependent (ι := { x : M // x s }) R Subtype.val) (t : Finset s) :
      LinearIndependent R Subtype.val

      Every finite subset of a linearly independent set is linearly independent.

      theorem linearIndependent_bounded_of_finset_linearIndependent_bounded {R : Type u_2} {M : Type u_4} [Semiring R] [AddCommMonoid M] [Module R M] {n : } (H : ∀ (s : Finset M), (LinearIndependent R fun (i : { x : M // x s }) => i)s.card n) (s : Set M) :
      LinearIndependent (ι := { x : M // x s }) R Subtype.valCardinal.mk s n

      If every finite set of linearly independent vectors has cardinality at most n, then the same is true for arbitrary sets of linearly independent vectors.

      The following lemmas use the subtype defined by a set in M as the index set ι.

      theorem linearIndependent_comp_subtype {ι : Type u'} {R : Type u_2} {M : Type u_4} {v : ιM} [Semiring R] [AddCommMonoid M] [Module R M] {s : Set ι} :
      LinearIndependent (ι := s) R (v Subtype.val) lFinsupp.supported R R s, (Finsupp.total ι M R v) l = 0l = 0
      theorem linearDependent_comp_subtype' {ι : Type u'} {R : Type u_2} {M : Type u_4} {v : ιM} [Semiring R] [AddCommMonoid M] [Module R M] {s : Set ι} :
      ¬LinearIndependent (ι := s) R (v Subtype.val) fFinsupp.supported R R s, (Finsupp.total ι M R v) f = 0 f 0
      theorem linearDependent_comp_subtype {ι : Type u'} {R : Type u_2} {M : Type u_4} {v : ιM} [Semiring R] [AddCommMonoid M] [Module R M] {s : Set ι} :
      ¬LinearIndependent (ι := s) R (v Subtype.val) fFinsupp.supported R R s, (f.support.sum fun (i : ι) => f i v i) = 0 f 0

      A version of linearDependent_comp_subtype' with Finsupp.total unfolded.

      theorem linearIndependent_subtype {R : Type u_2} {M : Type u_4} [Semiring R] [AddCommMonoid M] [Module R M] {s : Set M} :
      (LinearIndependent R fun (x : s) => x) lFinsupp.supported R R s, (Finsupp.total M M R id) l = 0l = 0
      theorem linearIndependent_comp_subtype_disjoint {ι : Type u'} {R : Type u_2} {M : Type u_4} {v : ιM} [Semiring R] [AddCommMonoid M] [Module R M] {s : Set ι} :
      LinearIndependent (ι := s) R (v Subtype.val) Disjoint (Finsupp.supported R R s) (LinearMap.ker (Finsupp.total ι M R v))
      theorem linearIndependent_subtype_disjoint {R : Type u_2} {M : Type u_4} [Semiring R] [AddCommMonoid M] [Module R M] {s : Set M} :
      (LinearIndependent R fun (x : s) => x) Disjoint (Finsupp.supported R R s) (LinearMap.ker (Finsupp.total M M R id))
      theorem linearIndependent_iff_totalOn {R : Type u_2} {M : Type u_4} [Semiring R] [AddCommMonoid M] [Module R M] {s : Set M} :
      (LinearIndependent R fun (x : s) => x) LinearMap.ker (Finsupp.totalOn M M R id s) =
      theorem LinearIndependent.restrict_of_comp_subtype {ι : Type u'} {R : Type u_2} {M : Type u_4} {v : ιM} [Semiring R] [AddCommMonoid M] [Module R M] {s : Set ι} (hs : LinearIndependent (ι := s) R (v Subtype.val)) :
      LinearIndependent (ι := s) R (s.restrict v)
      theorem linearIndependent_empty (R : Type u_2) (M : Type u_4) [Semiring R] [AddCommMonoid M] [Module R M] :
      LinearIndependent R fun (x : ) => x
      theorem LinearIndependent.mono {R : Type u_2} {M : Type u_4} [Semiring R] [AddCommMonoid M] [Module R M] {t : Set M} {s : Set M} (h : t s) :
      (LinearIndependent R fun (x : s) => x)LinearIndependent R fun (x : t) => x
      theorem linearIndependent_of_finite {R : Type u_2} {M : Type u_4} [Semiring R] [AddCommMonoid M] [Module R M] (s : Set M) (H : ts, t.FiniteLinearIndependent R fun (x : t) => x) :
      LinearIndependent R fun (x : s) => x
      theorem linearIndependent_iUnion_of_directed {R : Type u_2} {M : Type u_4} [Semiring R] [AddCommMonoid M] [Module R M] {η : Type u_8} {s : ηSet M} (hs : Directed (fun (x x_1 : Set M) => x x_1) s) (h : ∀ (i : η), LinearIndependent R fun (x : (s i)) => x) :
      LinearIndependent R fun (x : (⋃ (i : η), s i)) => x
      theorem linearIndependent_sUnion_of_directed {R : Type u_2} {M : Type u_4} [Semiring R] [AddCommMonoid M] [Module R M] {s : Set (Set M)} (hs : DirectedOn (fun (x x_1 : Set M) => x x_1) s) (h : as, LinearIndependent (ι := { x : M // x a }) R Subtype.val) :
      LinearIndependent R fun (x : (⋃₀ s)) => x
      theorem linearIndependent_biUnion_of_directed {R : Type u_2} {M : Type u_4} [Semiring R] [AddCommMonoid M] [Module R M] {η : Type u_8} {s : Set η} {t : ηSet M} (hs : DirectedOn (t ⁻¹'o fun (x x_1 : Set M) => x x_1) s) (h : as, LinearIndependent R fun (x : (t a)) => x) :
      LinearIndependent R fun (x : (as, t a)) => x

      Properties which require Ring R #

      theorem linearIndependent_iff_injective_total {ι : Type u'} {R : Type u_2} {M : Type u_4} {v : ιM} [Ring R] [AddCommGroup M] [Module R M] :
      theorem LinearIndependent.injective_total {ι : Type u'} {R : Type u_2} {M : Type u_4} {v : ιM} [Ring R] [AddCommGroup M] [Module R M] :

      Alias of the forward direction of linearIndependent_iff_injective_total.

      theorem LinearIndependent.injective {ι : Type u'} {R : Type u_2} {M : Type u_4} {v : ιM} [Ring R] [AddCommGroup M] [Module R M] [Nontrivial R] (hv : LinearIndependent R v) :
      theorem LinearIndependent.to_subtype_range {R : Type u_2} {M : Type u_4} [Ring R] [AddCommGroup M] [Module R M] {ι : Type u_8} {f : ιM} (hf : LinearIndependent R f) :
      LinearIndependent (ι := { x : M // x Set.range f }) R Subtype.val
      theorem LinearIndependent.to_subtype_range' {R : Type u_2} {M : Type u_4} [Ring R] [AddCommGroup M] [Module R M] {ι : Type u_8} {f : ιM} (hf : LinearIndependent R f) {t : Set M} (ht : Set.range f = t) :
      LinearIndependent (ι := { x : M // x t }) R Subtype.val
      theorem LinearIndependent.image_of_comp {R : Type u_2} {M : Type u_4} [Ring R] [AddCommGroup M] [Module R M] {ι : Type u_8} {ι' : Type u_9} (s : Set ι) (f : ιι') (g : ι'M) (hs : LinearIndependent R fun (x : s) => g (f x)) :
      LinearIndependent R fun (x : (f '' s)) => g x
      theorem LinearIndependent.image {R : Type u_2} {M : Type u_4} [Ring R] [AddCommGroup M] [Module R M] {ι : Type u_8} {s : Set ι} {f : ιM} (hs : LinearIndependent R fun (x : s) => f x) :
      LinearIndependent R fun (x : (f '' s)) => x
      theorem LinearIndependent.group_smul {ι : Type u'} {R : Type u_2} {M : Type u_4} [Ring R] [AddCommGroup M] [Module R M] {G : Type u_8} [hG : Group G] [DistribMulAction G R] [DistribMulAction G M] [IsScalarTower G R M] [SMulCommClass G R M] {v : ιM} (hv : LinearIndependent R v) (w : ιG) :
      theorem LinearIndependent.units_smul {ι : Type u'} {R : Type u_2} {M : Type u_4} [Ring R] [AddCommGroup M] [Module R M] {v : ιM} (hv : LinearIndependent R v) (w : ιRˣ) :
      theorem LinearIndependent.eq_of_pair {R : Type u_2} {M : Type u_4} [Ring R] [AddCommGroup M] [Module R M] {x : M} {y : M} (h : LinearIndependent R ![x, y]) {s : R} {t : R} {s' : R} {t' : R} (h' : s x + t y = s' x + t' y) :
      s = s' t = t'
      theorem LinearIndependent.eq_zero_of_pair' {R : Type u_2} {M : Type u_4} [Ring R] [AddCommGroup M] [Module R M] {x : M} {y : M} (h : LinearIndependent R ![x, y]) {s : R} {t : R} (h' : s x = t y) :
      s = 0 t = 0
      theorem LinearIndependent.linear_combination_pair_of_det_ne_zero {R : Type u_8} {M : Type u_9} [CommRing R] [NoZeroDivisors R] [AddCommGroup M] [Module R M] {x : M} {y : M} (h : LinearIndependent R ![x, y]) {a : R} {b : R} {c : R} {d : R} (h' : a * d - b * c 0) :
      LinearIndependent R ![a x + b y, c x + d y]

      If two vectors x and y are linearly independent, so are their linear combinations a x + b y and c x + d y provided the determinant a * d - b * c is nonzero.

      def LinearIndependent.Maximal {ι : Type w} {R : Type u} [Semiring R] {M : Type v} [AddCommMonoid M] [Module R M] {v : ιM} (_i : LinearIndependent R v) :

      A linearly independent family is maximal if there is no strictly larger linearly independent family.

      Equations
      Instances For
        theorem LinearIndependent.maximal_iff {ι : Type w} {R : Type u} [Ring R] [Nontrivial R] {M : Type v} [AddCommGroup M] [Module R M] {v : ιM} (i : LinearIndependent R v) :
        i.Maximal ∀ (κ : Type v) (w : κM), LinearIndependent R w∀ (j : ικ), w j = vFunction.Surjective j

        An alternative characterization of a maximal linearly independent family, quantifying over types (in the same universe as M) into which the indexing family injects.

        theorem LinearIndependent.eq_of_smul_apply_eq_smul_apply {ι : Type u'} {R : Type u_2} [Ring R] {M : Type u_8} [AddCommGroup M] [Module R M] {v : ιM} (li : LinearIndependent R v) (c : R) (d : R) (i : ι) (j : ι) (hc : c 0) (h : c v i = d v j) :
        i = j

        Linear independent families are injective, even if you multiply either side.

        The following lemmas use the subtype defined by a set in M as the index set ι.

        theorem LinearIndependent.disjoint_span_image {ι : Type u'} {R : Type u_2} {M : Type u_4} {v : ιM} [Ring R] [AddCommGroup M] [Module R M] (hv : LinearIndependent R v) {s : Set ι} {t : Set ι} (hs : Disjoint s t) :
        theorem LinearIndependent.not_mem_span_image {ι : Type u'} {R : Type u_2} {M : Type u_4} {v : ιM} [Ring R] [AddCommGroup M] [Module R M] [Nontrivial R] (hv : LinearIndependent R v) {s : Set ι} {x : ι} (h : xs) :
        v xSubmodule.span R (v '' s)
        theorem LinearIndependent.total_ne_of_not_mem_support {ι : Type u'} {R : Type u_2} {M : Type u_4} {v : ιM} [Ring R] [AddCommGroup M] [Module R M] [Nontrivial R] (hv : LinearIndependent R v) {x : ι} (f : ι →₀ R) (h : xf.support) :
        (Finsupp.total ι M R v) f v x
        theorem linearIndependent_sum {ι : Type u'} {ι' : Type u_1} {R : Type u_2} {M : Type u_4} [Ring R] [AddCommGroup M] [Module R M] {v : ι ι'M} :
        theorem LinearIndependent.sum_type {ι : Type u'} {ι' : Type u_1} {R : Type u_2} {M : Type u_4} {v : ιM} [Ring R] [AddCommGroup M] [Module R M] {v' : ι'M} (hv : LinearIndependent R v) (hv' : LinearIndependent R v') (h : Disjoint (Submodule.span R (Set.range v)) (Submodule.span R (Set.range v'))) :
        theorem LinearIndependent.union {R : Type u_2} {M : Type u_4} [Ring R] [AddCommGroup M] [Module R M] {s : Set M} {t : Set M} (hs : LinearIndependent R fun (x : s) => x) (ht : LinearIndependent R fun (x : t) => x) (hst : Disjoint (Submodule.span R s) (Submodule.span R t)) :
        LinearIndependent R fun (x : (s t)) => x
        theorem linearIndependent_iUnion_finite_subtype {R : Type u_2} {M : Type u_4} [Ring R] [AddCommGroup M] [Module R M] {ι : Type u_8} {f : ιSet M} (hl : ∀ (i : ι), LinearIndependent R fun (x : (f i)) => x) (hd : ∀ (i : ι) (t : Set ι), t.FiniteitDisjoint (Submodule.span R (f i)) (it, Submodule.span R (f i))) :
        LinearIndependent R fun (x : (⋃ (i : ι), f i)) => x
        theorem linearIndependent_iUnion_finite {R : Type u_2} {M : Type u_4} [Ring R] [AddCommGroup M] [Module R M] {η : Type u_8} {ιs : ηType u_9} {f : (j : η) → ιs jM} (hindep : ∀ (j : η), LinearIndependent R (f j)) (hd : ∀ (i : η) (t : Set η), t.FiniteitDisjoint (Submodule.span R (Set.range (f i))) (it, Submodule.span R (Set.range (f i)))) :
        LinearIndependent R fun (ji : (j : η) × ιs j) => f ji.fst ji.snd
        @[simp]
        theorem LinearIndependent.totalEquiv_symm_apply {ι : Type u'} {R : Type u_2} {M : Type u_4} {v : ιM} [Ring R] [AddCommGroup M] [Module R M] (hv : LinearIndependent R v) :
        ∀ (a : (Submodule.span R (Set.range v))), hv.totalEquiv.symm a = ((LinearEquiv.ofInjective (LinearMap.codRestrict (Submodule.span R (Set.range v)) (Finsupp.total ι M R v) ) ).toEquiv.trans (LinearEquiv.ofTop (LinearMap.range (LinearMap.codRestrict (Submodule.span R (Set.range v)) (Finsupp.total ι M R v) )) ).toEquiv).invFun a
        def LinearIndependent.totalEquiv {ι : Type u'} {R : Type u_2} {M : Type u_4} {v : ιM} [Ring R] [AddCommGroup M] [Module R M] (hv : LinearIndependent R v) :

        Canonical isomorphism between linear combinations and the span of linearly independent vectors.

        Equations
        Instances For
          @[simp]
          theorem LinearIndependent.totalEquiv_apply_coe {ι : Type u'} {R : Type u_2} {M : Type u_4} {v : ιM} [Ring R] [AddCommGroup M] [Module R M] (hv : LinearIndependent R v) (l : ι →₀ R) :
          (hv.totalEquiv l) = (Finsupp.total ι M R v) l
          def LinearIndependent.repr {ι : Type u'} {R : Type u_2} {M : Type u_4} {v : ιM} [Ring R] [AddCommGroup M] [Module R M] (hv : LinearIndependent R v) :

          Linear combination representing a vector in the span of linearly independent vectors.

          Given a family of linearly independent vectors, we can represent any vector in their span as a linear combination of these vectors. These are provided by this linear map. It is simply one direction of LinearIndependent.total_equiv.

          Equations
          • hv.repr = hv.totalEquiv.symm
          Instances For
            @[simp]
            theorem LinearIndependent.total_repr {ι : Type u'} {R : Type u_2} {M : Type u_4} {v : ιM} [Ring R] [AddCommGroup M] [Module R M] (hv : LinearIndependent R v) (x : (Submodule.span R (Set.range v))) :
            (Finsupp.total ι M R v) (hv.repr x) = x
            theorem LinearIndependent.total_comp_repr {ι : Type u'} {R : Type u_2} {M : Type u_4} {v : ιM} [Ring R] [AddCommGroup M] [Module R M] (hv : LinearIndependent R v) :
            Finsupp.total ι M R v ∘ₗ hv.repr = (Submodule.span R (Set.range v)).subtype
            theorem LinearIndependent.repr_ker {ι : Type u'} {R : Type u_2} {M : Type u_4} {v : ιM} [Ring R] [AddCommGroup M] [Module R M] (hv : LinearIndependent R v) :
            theorem LinearIndependent.repr_range {ι : Type u'} {R : Type u_2} {M : Type u_4} {v : ιM} [Ring R] [AddCommGroup M] [Module R M] (hv : LinearIndependent R v) :
            theorem LinearIndependent.repr_eq {ι : Type u'} {R : Type u_2} {M : Type u_4} {v : ιM} [Ring R] [AddCommGroup M] [Module R M] (hv : LinearIndependent R v) {l : ι →₀ R} {x : (Submodule.span R (Set.range v))} (eq : (Finsupp.total ι M R v) l = x) :
            hv.repr x = l
            theorem LinearIndependent.repr_eq_single {ι : Type u'} {R : Type u_2} {M : Type u_4} {v : ιM} [Ring R] [AddCommGroup M] [Module R M] (hv : LinearIndependent R v) (i : ι) (x : (Submodule.span R (Set.range v))) (hx : x = v i) :
            hv.repr x = Finsupp.single i 1
            theorem LinearIndependent.span_repr_eq {ι : Type u'} {R : Type u_2} {M : Type u_4} {v : ιM} [Ring R] [AddCommGroup M] [Module R M] (hv : LinearIndependent R v) [Nontrivial R] (x : (Submodule.span R (Set.range v))) :
            theorem linearIndependent_iff_not_smul_mem_span {ι : Type u'} {R : Type u_2} {M : Type u_4} {v : ιM} [Ring R] [AddCommGroup M] [Module R M] :
            LinearIndependent R v ∀ (i : ι) (a : R), a v i Submodule.span R (v '' (Set.univ \ {i}))a = 0
            theorem LinearIndependent.independent_span_singleton {ι : Type u'} {R : Type u_2} {M : Type u_4} {v : ιM} [Ring R] [AddCommGroup M] [Module R M] (hv : LinearIndependent R v) :

            See also CompleteLattice.independent_iff_linearIndependent_of_ne_zero.

            theorem exists_maximal_independent' {ι : Type u'} (R : Type u_2) {M : Type u_4} [Ring R] [AddCommGroup M] [Module R M] (s : ιM) :
            ∃ (I : Set ι), (LinearIndependent R fun (x : I) => s x) ∀ (J : Set ι), I J(LinearIndependent R fun (x : J) => s x)I = J
            theorem exists_maximal_independent {ι : Type u'} (R : Type u_2) {M : Type u_4} [Ring R] [AddCommGroup M] [Module R M] (s : ιM) :
            ∃ (I : Set ι), (LinearIndependent R fun (x : I) => s x) iI, ∃ (a : R), a 0 a s i Submodule.span R (s '' I)
            theorem surjective_of_linearIndependent_of_span {ι : Type u'} {ι' : Type u_1} {R : Type u_2} {M : Type u_4} {v : ιM} [Ring R] [AddCommGroup M] [Module R M] [Nontrivial R] (hv : LinearIndependent R v) (f : ι' ι) (hss : Set.range v (Submodule.span R (Set.range (v f)))) :
            theorem eq_of_linearIndependent_of_span_subtype {R : Type u_2} {M : Type u_4} [Ring R] [AddCommGroup M] [Module R M] [Nontrivial R] {s : Set M} {t : Set M} (hs : LinearIndependent R fun (x : s) => x) (h : t s) (hst : s (Submodule.span R t)) :
            s = t
            theorem LinearIndependent.image_subtype {R : Type u_2} {M : Type u_4} {M' : Type u_5} [Ring R] [AddCommGroup M] [AddCommGroup M'] [Module R M] [Module R M'] {s : Set M} {f : M →ₗ[R] M'} (hs : LinearIndependent R fun (x : s) => x) (hf_inj : Disjoint (Submodule.span R s) (LinearMap.ker f)) :
            LinearIndependent R fun (x : (f '' s)) => x
            theorem LinearIndependent.inl_union_inr {R : Type u_2} {M : Type u_4} {M' : Type u_5} [Ring R] [AddCommGroup M] [AddCommGroup M'] [Module R M] [Module R M'] {s : Set M} {t : Set M'} (hs : LinearIndependent R fun (x : s) => x) (ht : LinearIndependent R fun (x : t) => x) :
            LinearIndependent R fun (x : ((LinearMap.inl R M M') '' s (LinearMap.inr R M M') '' t)) => x
            theorem linearIndependent_inl_union_inr' {ι : Type u'} {ι' : Type u_1} {R : Type u_2} {M : Type u_4} {M' : Type u_5} [Ring R] [AddCommGroup M] [AddCommGroup M'] [Module R M] [Module R M'] {v : ιM} {v' : ι'M'} (hv : LinearIndependent R v) (hv' : LinearIndependent R v') :
            LinearIndependent R (Sum.elim ((LinearMap.inl R M M') v) ((LinearMap.inr R M M') v'))
            theorem linearIndependent_monoidHom (G : Type u_8) [Monoid G] (L : Type u_9) [CommRing L] [NoZeroDivisors L] :
            LinearIndependent L fun (f : G →* L) => f

            Dedekind's linear independence of characters

            theorem linearIndependent_algHom_toLinearMap (K : Type u_8) (M : Type u_9) (L : Type u_10) [CommSemiring K] [Semiring M] [Algebra K M] [CommRing L] [IsDomain L] [Algebra K L] :
            LinearIndependent L AlgHom.toLinearMap
            theorem linearIndependent_algHom_toLinearMap' (K : Type u_8) (M : Type u_9) (L : Type u_10) [CommRing K] [Semiring M] [Algebra K M] [CommRing L] [IsDomain L] [Algebra K L] [NoZeroSMulDivisors K L] :
            LinearIndependent K AlgHom.toLinearMap
            theorem le_of_span_le_span {R : Type u_2} {M : Type u_4} [Ring R] [AddCommGroup M] [Module R M] [Nontrivial R] {s : Set M} {t : Set M} {u : Set M} (hl : LinearIndependent (ι := { x : M // x u }) R Subtype.val) (hsu : s u) (htu : t u) (hst : Submodule.span R s Submodule.span R t) :
            s t
            theorem span_le_span_iff {R : Type u_2} {M : Type u_4} [Ring R] [AddCommGroup M] [Module R M] [Nontrivial R] {s : Set M} {t : Set M} {u : Set M} (hl : LinearIndependent (ι := { x : M // x u }) R Subtype.val) (hsu : s u) (htu : t u) :
            theorem linearIndependent_unique_iff {ι : Type u'} {R : Type u_2} {M : Type u_4} [Ring R] [Nontrivial R] [AddCommGroup M] [Module R M] [NoZeroSMulDivisors R M] (v : ιM) [Unique ι] :
            LinearIndependent R v v default 0
            theorem linearIndependent_unique {ι : Type u'} {R : Type u_2} {M : Type u_4} [Ring R] [Nontrivial R] [AddCommGroup M] [Module R M] [NoZeroSMulDivisors R M] (v : ιM) [Unique ι] :
            v default 0LinearIndependent R v

            Alias of the reverse direction of linearIndependent_unique_iff.

            theorem linearIndependent_singleton {R : Type u_2} {M : Type u_4} [Ring R] [Nontrivial R] [AddCommGroup M] [Module R M] [NoZeroSMulDivisors R M] {x : M} (hx : x 0) :
            LinearIndependent R fun (x_1 : {x}) => x_1

            Properties which require DivisionRing K #

            These can be considered generalizations of properties of linear independence in vector spaces.

            theorem mem_span_insert_exchange {K : Type u_3} {V : Type u} [DivisionRing K] [AddCommGroup V] [Module K V] {s : Set V} {x : V} {y : V} :
            x Submodule.span K (insert y s)xSubmodule.span K sy Submodule.span K (insert x s)
            theorem linearIndependent_iff_not_mem_span {ι : Type u'} {K : Type u_3} {V : Type u} [DivisionRing K] [AddCommGroup V] [Module K V] {v : ιV} :
            LinearIndependent K v ∀ (i : ι), v iSubmodule.span K (v '' (Set.univ \ {i}))
            theorem LinearIndependent.insert {K : Type u_3} {V : Type u} [DivisionRing K] [AddCommGroup V] [Module K V] {s : Set V} {x : V} (hs : LinearIndependent K fun (b : s) => b) (hx : xSubmodule.span K s) :
            LinearIndependent K fun (b : (insert x s)) => b
            theorem linearIndependent_option' {ι : Type u'} {K : Type u_3} {V : Type u} [DivisionRing K] [AddCommGroup V] [Module K V] {v : ιV} {x : V} :
            (LinearIndependent K fun (o : Option ι) => o.casesOn' x v) LinearIndependent K v xSubmodule.span K (Set.range v)
            theorem LinearIndependent.option {ι : Type u'} {K : Type u_3} {V : Type u} [DivisionRing K] [AddCommGroup V] [Module K V] {v : ιV} {x : V} (hv : LinearIndependent K v) (hx : xSubmodule.span K (Set.range v)) :
            LinearIndependent K fun (o : Option ι) => o.casesOn' x v
            theorem linearIndependent_option {ι : Type u'} {K : Type u_3} {V : Type u} [DivisionRing K] [AddCommGroup V] [Module K V] {v : Option ιV} :
            LinearIndependent K v LinearIndependent K (v some) v noneSubmodule.span K (Set.range (v some))
            theorem linearIndependent_insert' {K : Type u_3} {V : Type u} [DivisionRing K] [AddCommGroup V] [Module K V] {ι : Type u_8} {s : Set ι} {a : ι} {f : ιV} (has : as) :
            (LinearIndependent K fun (x : (insert a s)) => f x) (LinearIndependent K fun (x : s) => f x) f aSubmodule.span K (f '' s)
            theorem linearIndependent_insert {K : Type u_3} {V : Type u} [DivisionRing K] [AddCommGroup V] [Module K V] {s : Set V} {x : V} (hxs : xs) :
            (LinearIndependent K fun (b : (insert x s)) => b) (LinearIndependent K fun (b : s) => b) xSubmodule.span K s
            theorem linearIndependent_pair {K : Type u_3} {V : Type u} [DivisionRing K] [AddCommGroup V] [Module K V] {x : V} {y : V} (hx : x 0) (hy : ∀ (a : K), a x y) :
            LinearIndependent (ι := { x_1 : V // x_1 {x, y} }) K Subtype.val
            theorem linearIndependent_fin_cons {K : Type u_3} {V : Type u} [DivisionRing K] [AddCommGroup V] [Module K V] {x : V} {n : } {v : Fin nV} :
            theorem linearIndependent_fin_snoc {K : Type u_3} {V : Type u} [DivisionRing K] [AddCommGroup V] [Module K V] {x : V} {n : } {v : Fin nV} :
            theorem LinearIndependent.fin_cons {K : Type u_3} {V : Type u} [DivisionRing K] [AddCommGroup V] [Module K V] {x : V} {n : } {v : Fin nV} (hv : LinearIndependent K v) (hx : xSubmodule.span K (Set.range v)) :

            See LinearIndependent.fin_cons' for an uglier version that works if you only have a module over a semiring.

            theorem linearIndependent_fin_succ {K : Type u_3} {V : Type u} [DivisionRing K] [AddCommGroup V] [Module K V] {n : } {v : Fin (n + 1)V} :
            theorem linearIndependent_fin_succ' {K : Type u_3} {V : Type u} [DivisionRing K] [AddCommGroup V] [Module K V] {n : } {v : Fin (n + 1)V} :
            theorem linearIndependent_fin2 {K : Type u_3} {V : Type u} [DivisionRing K] [AddCommGroup V] [Module K V] {f : Fin 2V} :
            LinearIndependent K f f 1 0 ∀ (a : K), a f 1 f 0
            theorem exists_linearIndependent_extension {K : Type u_3} {V : Type u} [DivisionRing K] [AddCommGroup V] [Module K V] {s : Set V} {t : Set V} (hs : LinearIndependent (ι := { x : V // x s }) K Subtype.val) (hst : s t) :
            bt, s b t (Submodule.span K b) LinearIndependent (ι := { x : V // x b }) K Subtype.val
            theorem exists_linearIndependent (K : Type u_3) {V : Type u} [DivisionRing K] [AddCommGroup V] [Module K V] (t : Set V) :
            bt, Submodule.span K b = Submodule.span K t LinearIndependent (ι := { x : V // x b }) K Subtype.val
            noncomputable def LinearIndependent.extend {K : Type u_3} {V : Type u} [DivisionRing K] [AddCommGroup V] [Module K V] {s : Set V} {t : Set V} (hs : LinearIndependent K fun (x : s) => x) (hst : s t) :
            Set V

            LinearIndependent.extend adds vectors to a linear independent set s ⊆ t until it spans all elements of t.

            Equations
            Instances For
              theorem LinearIndependent.extend_subset {K : Type u_3} {V : Type u} [DivisionRing K] [AddCommGroup V] [Module K V] {s : Set V} {t : Set V} (hs : LinearIndependent K fun (x : s) => x) (hst : s t) :
              hs.extend hst t
              theorem LinearIndependent.subset_extend {K : Type u_3} {V : Type u} [DivisionRing K] [AddCommGroup V] [Module K V] {s : Set V} {t : Set V} (hs : LinearIndependent K fun (x : s) => x) (hst : s t) :
              s hs.extend hst
              theorem LinearIndependent.subset_span_extend {K : Type u_3} {V : Type u} [DivisionRing K] [AddCommGroup V] [Module K V] {s : Set V} {t : Set V} (hs : LinearIndependent K fun (x : s) => x) (hst : s t) :
              t (Submodule.span K (hs.extend hst))
              theorem LinearIndependent.linearIndependent_extend {K : Type u_3} {V : Type u} [DivisionRing K] [AddCommGroup V] [Module K V] {s : Set V} {t : Set V} (hs : LinearIndependent K fun (x : s) => x) (hst : s t) :
              LinearIndependent (ι := { x : V // x hs.extend hst }) K Subtype.val
              theorem exists_of_linearIndependent_of_finite_span {K : Type u_3} {V : Type u} [DivisionRing K] [AddCommGroup V] [Module K V] {s : Set V} {t : Finset V} (hs : LinearIndependent K fun (x : s) => x) (hst : s (Submodule.span K t)) :
              ∃ (t' : Finset V), t' s t s t' t'.card = t.card
              theorem exists_finite_card_le_of_finite_of_linearIndependent_of_span {K : Type u_3} {V : Type u} [DivisionRing K] [AddCommGroup V] [Module K V] {s : Set V} {t : Set V} (ht : t.Finite) (hs : LinearIndependent K fun (x : s) => x) (hst : s (Submodule.span K t)) :
              ∃ (h : s.Finite), h.toFinset.card ht.toFinset.card