doc-next-gen

Mathlib.LinearAlgebra.Span.Defs

Module docstring

{"# The span of a set of vectors, as a submodule

  • Submodule.span s is defined to be the smallest submodule containing the set s.

Notations

  • We introduce the notation R ∙ v for the span of a singleton, Submodule.span R {v}. This is \\span, not the same as the scalar multiplication /\\bub.

"}

Submodule.span definition
(s : Set M) : Submodule R M
Full source
/-- The span of a set `s ⊆ M` is the smallest submodule of M that contains `s`. -/
def span (s : Set M) : Submodule R M :=
  sInf { p | s ⊆ p }
Span of a set of vectors in a module
Informal description
The span of a subset $s$ of an $R$-module $M$ is the smallest submodule of $M$ containing $s$. It consists of all finite $R$-linear combinations of elements from $s$.
Submodule.IsPrincipal structure
(S : Submodule R M)
Full source
/-- An `R`-submodule of `M` is principal if it is generated by one element. -/
@[mk_iff]
class IsPrincipal (S : Submodule R M) : Prop where
  principal (S) : ∃ a, S = span R {a}
Principal Submodule
Informal description
A submodule $S$ of an $R$-module $M$ is called *principal* if there exists an element $v \in M$ such that $S$ is generated by $v$, i.e., $S$ is the smallest submodule containing $v$. In other words, $S$ consists of all $R$-linear combinations of $v$.
Submodule.mem_span theorem
: x ∈ span R s ↔ ∀ p : Submodule R M, s ⊆ p → x ∈ p
Full source
theorem mem_span : x ∈ span R sx ∈ span R s ↔ ∀ p : Submodule R M, s ⊆ p → x ∈ p :=
  mem_iInter₂
Characterization of Span Membership via Submodule Containment
Informal description
An element $x$ belongs to the span of a set $s$ in an $R$-module $M$ if and only if $x$ is contained in every submodule of $M$ that contains $s$. In other words: $$x \in \operatorname{span}_R s \leftrightarrow \forall p \leq M, s \subseteq p \implies x \in p$$
Submodule.subset_span theorem
: s ⊆ span R s
Full source
@[aesop safe 20 apply (rule_sets := [SetLike])]
theorem subset_span : s ⊆ span R s := fun _ h => mem_span.2 fun _ hp => hp h
Span Contains Original Set
Informal description
For any subset $s$ of an $R$-module $M$, the span of $s$ contains $s$, i.e., $s \subseteq \operatorname{span}_R s$.
Submodule.span_le theorem
{p} : span R s ≤ p ↔ s ⊆ p
Full source
theorem span_le {p} : spanspan R s ≤ p ↔ s ⊆ p :=
  ⟨Subset.trans subset_span, fun ss _ h => mem_span.1 h _ ss⟩
Span containment criterion: $\operatorname{span}_R s \leq p \leftrightarrow s \subseteq p$
Informal description
For any subset $s$ of an $R$-module $M$ and any submodule $p$ of $M$, the span of $s$ is contained in $p$ if and only if $s$ is a subset of $p$. In other words: $$\operatorname{span}_R s \leq p \leftrightarrow s \subseteq p$$
Submodule.span_mono theorem
(h : s ⊆ t) : span R s ≤ span R t
Full source
@[gcongr] theorem span_mono (h : s ⊆ t) : span R s ≤ span R t :=
  span_le.2 <| Subset.trans h subset_span
Monotonicity of Span: $s \subseteq t \Rightarrow \operatorname{span}_R s \leq \operatorname{span}_R t$
Informal description
For any subsets $s$ and $t$ of an $R$-module $M$, if $s$ is contained in $t$ (i.e., $s \subseteq t$), then the span of $s$ is contained in the span of $t$ (i.e., $\operatorname{span}_R s \leq \operatorname{span}_R t$).
Submodule.span_monotone theorem
: Monotone (span R : Set M → Submodule R M)
Full source
theorem span_monotone : Monotone (span R : Set M → Submodule R M) := fun _ _ => span_mono
Monotonicity of Span: $s \subseteq t \Rightarrow \operatorname{span}_R s \leq \operatorname{span}_R t$
Informal description
The span function $\operatorname{span}_R$ is monotone with respect to subset inclusion, meaning that for any subsets $s$ and $t$ of an $R$-module $M$, if $s \subseteq t$, then $\operatorname{span}_R s \leq \operatorname{span}_R t$.
Submodule.span_eq_of_le theorem
(h₁ : s ⊆ p) (h₂ : p ≤ span R s) : span R s = p
Full source
theorem span_eq_of_le (h₁ : s ⊆ p) (h₂ : p ≤ span R s) : span R s = p :=
  le_antisymm (span_le.2 h₁) h₂
Span Equality under Submodule Containment: $\operatorname{span}_R s = p$ when $s \subseteq p \leq \operatorname{span}_R s$
Informal description
Let $s$ be a subset of an $R$-module $M$ and $p$ be a submodule of $M$. If $s \subseteq p$ and $p \leq \operatorname{span}_R s$, then $\operatorname{span}_R s = p$.
Submodule.span_eq theorem
: span R (p : Set M) = p
Full source
theorem span_eq : span R (p : Set M) = p :=
  span_eq_of_le _ (Subset.refl _) subset_span
Span of a Submodule Equals Itself
Informal description
For any submodule $p$ of an $R$-module $M$, the span of $p$ (viewed as a subset of $M$) is equal to $p$ itself, i.e., $\operatorname{span}_R p = p$.
Submodule.span_eq_span theorem
(hs : s ⊆ span R t) (ht : t ⊆ span R s) : span R s = span R t
Full source
theorem span_eq_span (hs : s ⊆ span R t) (ht : t ⊆ span R s) : span R s = span R t :=
  le_antisymm (span_le.2 hs) (span_le.2 ht)
Span Equality under Mutual Containment: $\operatorname{span}_R s = \operatorname{span}_R t$ when $s \subseteq \operatorname{span}_R t$ and $t \subseteq \operatorname{span}_R s$
Informal description
For any subsets $s$ and $t$ of an $R$-module $M$, if $s$ is contained in the span of $t$ and $t$ is contained in the span of $s$, then the spans of $s$ and $t$ are equal, i.e., $\operatorname{span}_R s = \operatorname{span}_R t$.
Submodule.coe_span_eq_self theorem
[SetLike S M] [AddSubmonoidClass S M] [SMulMemClass S R M] (s : S) : (span R (s : Set M) : Set M) = s
Full source
/-- A version of `Submodule.span_eq` for subobjects closed under addition and scalar multiplication
and containing zero. In general, this should not be used directly, but can be used to quickly
generate proofs for specific types of subobjects. -/
lemma coe_span_eq_self [SetLike S M] [AddSubmonoidClass S M] [SMulMemClass S R M] (s : S) :
    (span R (s : Set M) : Set M) = s := by
  refine le_antisymm ?_ subset_span
  let s' : Submodule R M :=
    { carrier := s
      add_mem' := add_mem
      zero_mem' := zero_mem _
      smul_mem' := SMulMemClass.smul_mem }
  exact span_le (p := s') |>.mpr le_rfl
Span of a Closed Subset Equals Itself
Informal description
Let $M$ be a module over a ring $R$, and let $S$ be a type of subsets of $M$ that forms a `SetLike` structure. Suppose $S$ is equipped with instances of `AddSubmonoidClass` (i.e., closed under addition and contains zero) and `SMulMemClass` (i.e., closed under scalar multiplication). Then for any subset $s \in S$, the underlying set of the span of $s$ is equal to $s$ itself, i.e., $$(\operatorname{span}_R s) = s.$$
Submodule.span_insert_zero theorem
: span R (insert (0 : M) s) = span R s
Full source
@[simp]
theorem span_insert_zero : span R (insert (0 : M) s) = span R s := by
  refine le_antisymm ?_ (Submodule.span_mono (Set.subset_insert 0 s))
  rw [span_le, Set.insert_subset_iff]
  exact ⟨by simp only [SetLike.mem_coe, Submodule.zero_mem], Submodule.subset_span⟩
Span Invariance under Zero Insertion: $\operatorname{span}_R (\{0\} \cup s) = \operatorname{span}_R s$
Informal description
For any subset $s$ of an $R$-module $M$, the span of the set obtained by inserting the zero vector into $s$ is equal to the span of $s$ itself, i.e., $$\operatorname{span}_R (\{0\} \cup s) = \operatorname{span}_R s.$$
Submodule.span_sdiff_singleton_zero theorem
: span R (s \ {0}) = span R s
Full source
@[simp]
lemma span_sdiff_singleton_zero : span R (s \ {0}) = span R s := by
  by_cases h : 0 ∈ s
  · rw [← span_insert_zero, show insert 0 (s \ {0}) = s by simp [h]]
  · simp [h]
Span Invariance under Zero Removal: $\operatorname{span}_R (s \setminus \{0\}) = \operatorname{span}_R s$
Informal description
For any subset $s$ of an $R$-module $M$, the span of the set difference $s \setminus \{0\}$ is equal to the span of $s$ itself, i.e., $$\operatorname{span}_R (s \setminus \{0\}) = \operatorname{span}_R s.$$
Submodule.closure_subset_span theorem
{s : Set M} : (AddSubmonoid.closure s : Set M) ⊆ span R s
Full source
theorem closure_subset_span {s : Set M} : (AddSubmonoid.closure s : Set M) ⊆ span R s :=
  (@AddSubmonoid.closure_le _ _ _ (span R s).toAddSubmonoid).mpr subset_span
Additive Submonoid Closure is Contained in Linear Span
Informal description
For any subset $s$ of an $R$-module $M$, the additive submonoid closure of $s$ is contained in the $R$-linear span of $s$, i.e., $\operatorname{AddSubmonoid.closure}(s) \subseteq \operatorname{span}_R s$.
Submodule.closure_le_toAddSubmonoid_span theorem
{s : Set M} : AddSubmonoid.closure s ≤ (span R s).toAddSubmonoid
Full source
theorem closure_le_toAddSubmonoid_span {s : Set M} :
    AddSubmonoid.closure s ≤ (span R s).toAddSubmonoid :=
  closure_subset_span
Additive Submonoid Closure is Contained in Span's Additive Submonoid
Informal description
For any subset $s$ of an $R$-module $M$, the additive submonoid closure of $s$ is contained in the additive submonoid associated with the $R$-linear span of $s$, i.e., $\operatorname{AddSubmonoid.closure}(s) \leq (\operatorname{span}_R s).\text{toAddSubmonoid}$.
Submodule.span_closure theorem
{s : Set M} : span R (AddSubmonoid.closure s : Set M) = span R s
Full source
@[simp]
theorem span_closure {s : Set M} : span R (AddSubmonoid.closure s : Set M) = span R s :=
  le_antisymm (span_le.mpr closure_subset_span) (span_mono AddSubmonoid.subset_closure)
Span Invariance under Additive Submonoid Closure
Informal description
For any subset $s$ of an $R$-module $M$, the $R$-linear span of the additive submonoid closure of $s$ is equal to the $R$-linear span of $s$ itself. That is: $$\operatorname{span}_R (\operatorname{AddSubmonoid.closure}(s)) = \operatorname{span}_R s$$
Submodule.span_induction theorem
{p : (x : M) → x ∈ span R s → Prop} (mem : ∀ (x) (h : x ∈ s), p x (subset_span h)) (zero : p 0 (Submodule.zero_mem _)) (add : ∀ x y hx hy, p x hx → p y hy → p (x + y) (Submodule.add_mem _ ‹_› ‹_›)) (smul : ∀ (a : R) (x hx), p x hx → p (a • x) (Submodule.smul_mem _ _ ‹_›)) {x} (hx : x ∈ span R s) : p x hx
Full source
/-- 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`. -/
@[elab_as_elim]
theorem span_induction {p : (x : M) → x ∈ span R s → Prop}
    (mem : ∀ (x) (h : x ∈ s), p x (subset_span h))
    (zero : p 0 (Submodule.zero_mem _))
    (add : ∀ x y hx hy, p x hx → p y hy → p (x + y) (Submodule.add_mem _ ‹_› ‹_›))
    (smul : ∀ (a : R) (x hx), p x hx → p (a • x) (Submodule.smul_mem _ _ ‹_›)) {x}
    (hx : x ∈ span R s) : p x hx := by
  let p : Submodule R M :=
    { carrier := { x | ∃ hx, p x hx }
      add_mem' := fun ⟨_, hpx⟩ ⟨_, hpy⟩ ↦ ⟨_, add _ _ _ _ hpx hpy⟩
      zero_mem' := ⟨_, zero⟩
      smul_mem' := fun r ↦ fun ⟨_, hpx⟩ ↦ ⟨_, smul r _ _ hpx⟩ }
  exact span_le (p := p) |>.mpr (fun y hy ↦ ⟨subset_span hy, mem y hy⟩) hx |>.elim fun _ ↦ id
Induction Principle for Span Membership in Modules
Informal description
Let $M$ be a module over a ring $R$ and let $s$ be a subset of $M$. For a predicate $p$ on elements of $M$ and their membership in $\operatorname{span}_R s$, suppose the following hold: 1. For every $x \in s$, $p(x)$ holds. 2. $p(0)$ holds. 3. For any $x, y \in \operatorname{span}_R s$, if $p(x)$ and $p(y)$ hold, then $p(x + y)$ holds. 4. For any $a \in R$ and $x \in \operatorname{span}_R s$, if $p(x)$ holds, then $p(a \cdot x)$ holds. Then, for every $x \in \operatorname{span}_R s$, $p(x)$ holds.
Submodule.span_induction₂ theorem
{N : Type*} [AddCommMonoid N] [Module R N] {t : Set N} {p : (x : M) → (y : N) → x ∈ span R s → y ∈ span R t → Prop} (mem_mem : ∀ (x) (y) (hx : x ∈ s) (hy : y ∈ t), p x y (subset_span hx) (subset_span hy)) (zero_left : ∀ y hy, p 0 y (zero_mem _) hy) (zero_right : ∀ x hx, p x 0 hx (zero_mem _)) (add_left : ∀ x y z hx hy hz, p x z hx hz → p y z hy hz → p (x + y) z (add_mem hx hy) hz) (add_right : ∀ x y z hx hy hz, p x y hx hy → p x z hx hz → p x (y + z) hx (add_mem hy hz)) (smul_left : ∀ (r : R) x y hx hy, p x y hx hy → p (r • x) y (smul_mem _ r hx) hy) (smul_right : ∀ (r : R) x y hx hy, p x y hx hy → p x (r • y) hx (smul_mem _ r hy)) {a : M} {b : N} (ha : a ∈ Submodule.span R s) (hb : b ∈ Submodule.span R t) : p a b ha hb
Full source
/-- An induction principle for span membership. This is a version of `Submodule.span_induction`
for binary predicates. -/
theorem span_induction₂ {N : Type*} [AddCommMonoid N] [Module R N] {t : Set N}
    {p : (x : M) → (y : N) → x ∈ span R sy ∈ span R t → Prop}
    (mem_mem : ∀ (x) (y) (hx : x ∈ s) (hy : y ∈ t), p x y (subset_span hx) (subset_span hy))
    (zero_left : ∀ y hy, p 0 y (zero_mem _) hy) (zero_right : ∀ x hx, p x 0 hx (zero_mem _))
    (add_left : ∀ x y z hx hy hz, p x z hx hz → p y z hy hz → p (x + y) z (add_mem hx hy) hz)
    (add_right : ∀ x y z hx hy hz, p x y hx hy → p x z hx hz → p x (y + z) hx (add_mem hy hz))
    (smul_left : ∀ (r : R) x y hx hy, p x y hx hy → p (r • x) y (smul_mem _ r hx) hy)
    (smul_right : ∀ (r : R) x y hx hy, p x y hx hy → p x (r • y) hx (smul_mem _ r hy))
    {a : M} {b : N} (ha : a ∈ Submodule.span R s)
    (hb : b ∈ Submodule.span R t) : p a b ha hb := by
  induction hb using span_induction with
  | mem z hz => induction ha using span_induction with
    | mem _ h => exact mem_mem _ _ h hz
    | zero => exact zero_left _ _
    | add _ _ _ _ h₁ h₂ => exact add_left _ _ _ _ _ _ h₁ h₂
    | smul _ _ _ h => exact smul_left _ _ _ _ _ h
  | zero => exact zero_right a ha
  | add _ _ _ _ h₁ h₂ => exact add_right _ _ _ _ _ _ h₁ h₂
  | smul _ _ _ h => exact smul_right _ _ _ _ _ h
Double Induction Principle for Span Membership in Modules
Informal description
Let $M$ and $N$ be modules over a ring $R$, and let $s \subseteq M$ and $t \subseteq N$ be subsets. Given a predicate $p : M \times N \to \mathrm{Prop}$ such that: 1. For all $x \in s$ and $y \in t$, $p(x,y)$ holds; 2. For all $y \in \operatorname{span}_R t$, $p(0,y)$ holds; 3. For all $x \in \operatorname{span}_R s$, $p(x,0)$ holds; 4. For all $x, y \in \operatorname{span}_R s$ and $z \in \operatorname{span}_R t$, if $p(x,z)$ and $p(y,z)$ hold, then $p(x+y,z)$ holds; 5. For all $x \in \operatorname{span}_R s$ and $y, z \in \operatorname{span}_R t$, if $p(x,y)$ and $p(x,z)$ hold, then $p(x,y+z)$ holds; 6. For all $r \in R$, $x \in \operatorname{span}_R s$, and $y \in \operatorname{span}_R t$, if $p(x,y)$ holds, then $p(r \cdot x, y)$ holds; 7. For all $r \in R$, $x \in \operatorname{span}_R s$, and $y \in \operatorname{span}_R t$, if $p(x,y)$ holds, then $p(x, r \cdot y)$ holds. Then, for all $a \in \operatorname{span}_R s$ and $b \in \operatorname{span}_R t$, the predicate $p(a,b)$ holds.
Submodule.span_eq_closure theorem
{s : Set M} : (span R s).toAddSubmonoid = closure (@univ R • s)
Full source
theorem span_eq_closure {s : Set M} : (span R s).toAddSubmonoid = closure (@univ R • s) := by
  refine le_antisymm (fun x (hx : x ∈ span R s) ↦ ?of_mem_span) (fun x hx ↦ ?of_mem_closure)
  case of_mem_span =>
    induction hx using span_induction with
    | mem x hx => exact subset_closure ⟨1, trivial, x, hx, one_smul R x⟩
    | zero => exact zero_mem _
    | add _ _ _ _ h₁ h₂ => exact add_mem h₁ h₂
    | smul r₁ y _h hy =>
      clear _h
      induction hy using AddSubmonoid.closure_induction with
      | mem _ h =>
        obtain ⟨r₂, -, x, hx, rfl⟩ := h
        exact subset_closure ⟨r₁ * r₂, trivial, x, hx, mul_smul ..⟩
      | one => simpa only [smul_zero] using zero_mem _
      | mul _ _ _ _ h₁ h₂ => simpa only [smul_add] using add_mem h₁ h₂
  case of_mem_closure =>
    refine closure_le.2 ?_ hx
    rintro - ⟨r, -, x, hx, rfl⟩
    exact smul_mem _ _ (subset_span hx)
Span as Additive Closure of Scalar Multiples
Informal description
For any subset $s$ of an $R$-module $M$, the additive submonoid underlying the span of $s$ is equal to the additive submonoid closure of the set of all scalar multiples of elements of $s$ by scalars from $R$, i.e., \[ (\operatorname{span}_R s)_{\text{add}} = \langle R \cdot s \rangle_{\text{add}}. \] Here, $R \cdot s$ denotes the set $\{r \cdot x \mid r \in R, x \in s\}$ and $\langle \cdot \rangle_{\text{add}}$ denotes the additive submonoid closure.
Submodule.closure_induction theorem
{p : (x : M) → x ∈ span R s → Prop} (zero : p 0 (Submodule.zero_mem _)) (add : ∀ x y hx hy, p x hx → p y hy → p (x + y) (Submodule.add_mem _ ‹_› ‹_›)) (smul_mem : ∀ (r x) (h : x ∈ s), p (r • x) (Submodule.smul_mem _ _ <| subset_span h)) {x} (hx : x ∈ span R s) : p x hx
Full source
/-- A variant of `span_induction` that combines `∀ x ∈ s, p x` and `∀ r x, p x → p (r • x)`
into a single condition `∀ r, ∀ x ∈ s, p (r • x)`, which can be easier to verify. -/
@[elab_as_elim]
theorem closure_induction {p : (x : M) → x ∈ span R s → Prop}
    (zero : p 0 (Submodule.zero_mem _))
    (add : ∀ x y hx hy, p x hx → p y hy → p (x + y) (Submodule.add_mem _ ‹_› ‹_›))
    (smul_mem : ∀ (r x) (h : x ∈ s), p (r • x) (Submodule.smul_mem _ _ <| subset_span h)) {x}
    (hx : x ∈ span R s) : p x hx := by
  have key {v} : v ∈ span R sv ∈ span R s ↔ v ∈ closure (@univ R • s) := by simp [← span_eq_closure]
  refine AddSubmonoid.closure_induction (motive := fun x hx ↦ p x (key.mpr hx))
    ?_ zero (by simpa only [key] using add) (key.mp hx)
  rintro - ⟨r, -, x, hx, rfl⟩
  exact smul_mem r x hx
Induction Principle for the Span of a Set in a Module
Informal description
Let $M$ be a module over a ring $R$, and let $s$ be a subset of $M$. Suppose $p$ is a predicate on elements of $M$ such that: 1. $p(0)$ holds, where $0$ is the zero vector in $M$; 2. For any $x, y \in M$ with $x \in \operatorname{span}_R s$ and $y \in \operatorname{span}_R s$, if $p(x)$ and $p(y)$ hold, then $p(x + y)$ holds; 3. For any scalar $r \in R$ and any $x \in s$, $p(r \cdot x)$ holds. Then, for every $x \in \operatorname{span}_R s$, the predicate $p(x)$ holds.
Submodule.span_span_coe_preimage theorem
: span R (((↑) : span R s → M) ⁻¹' s) = ⊤
Full source
@[simp]
theorem span_span_coe_preimage : span R (((↑) : span R s → M) ⁻¹' s) =  :=
  eq_top_iff.2 fun x _ ↦ Subtype.recOn x fun _ hx' ↦
    span_induction (fun _ h ↦ subset_span h) (zero_mem _) (fun _ _ _ _ ↦ add_mem)
      (fun _ _ _ ↦ smul_mem _ _) hx'
Span of Preimage Equals Original Span in Module
Informal description
Let $M$ be a module over a ring $R$ and $s$ be a subset of $M$. The span of the preimage of $s$ under the canonical inclusion map from $\operatorname{span}_R s$ to $M$ is equal to the entire submodule $\operatorname{span}_R s$, i.e., $$\operatorname{span}_R \left( (\iota^{-1}(s)) \right) = \operatorname{span}_R s$$ where $\iota : \operatorname{span}_R s \hookrightarrow M$ is the inclusion map.
Submodule.span_setOf_mem_eq_top theorem
: span R {x : span R s | (x : M) ∈ s} = ⊤
Full source
@[simp]
lemma span_setOf_mem_eq_top :
    span R {x : span R s | (x : M) ∈ s} =  :=
  span_span_coe_preimage
Span of Elements in Original Set Equals Original Span
Informal description
Let $M$ be a module over a ring $R$ and $s$ be a subset of $M$. The span of the set $\{x \in \operatorname{span}_R s \mid x \in s\}$ is equal to the entire submodule $\operatorname{span}_R s$, i.e., $$\operatorname{span}_R \{x \in \operatorname{span}_R s \mid x \in s\} = \operatorname{span}_R s.$$
Submodule.span_nat_eq_addSubmonoid_closure theorem
(s : Set M) : (span ℕ s).toAddSubmonoid = AddSubmonoid.closure s
Full source
theorem span_nat_eq_addSubmonoid_closure (s : Set M) :
    (span  s).toAddSubmonoid = AddSubmonoid.closure s := by
  refine Eq.symm (AddSubmonoid.closure_eq_of_le subset_span ?_)
  apply (OrderIso.to_galoisConnection (AddSubmonoid.toNatSubmodule (M := M)).symm).l_le
     (a := span  s) (b := AddSubmonoid.closure s)
  rw [span_le]
  exact AddSubmonoid.subset_closure
Equality of $\mathbb{N}$-span and additive submonoid closure
Informal description
For any subset $s$ of an additive monoid $M$, the additive submonoid generated by the $\mathbb{N}$-span of $s$ is equal to the additive submonoid closure of $s$. In other words: $$(\operatorname{span}_{\mathbb{N}} s).\text{toAddSubmonoid} = \text{AddSubmonoid.closure}(s)$$
Submodule.span_nat_eq theorem
(s : AddSubmonoid M) : (span ℕ (s : Set M)).toAddSubmonoid = s
Full source
@[simp]
theorem span_nat_eq (s : AddSubmonoid M) : (span  (s : Set M)).toAddSubmonoid = s := by
  rw [span_nat_eq_addSubmonoid_closure, s.closure_eq]
$\mathbb{N}$-span of Additive Submonoid Equals Itself
Informal description
For any additive submonoid $s$ of an additive monoid $M$, the additive submonoid generated by the $\mathbb{N}$-span of $s$ (viewed as a subset of $M$) is equal to $s$ itself. In other words: $$(\operatorname{span}_{\mathbb{N}} s).\text{toAddSubmonoid} = s$$
Submodule.span_int_eq_addSubgroup_closure theorem
{M : Type*} [AddCommGroup M] (s : Set M) : (span ℤ s).toAddSubgroup = AddSubgroup.closure s
Full source
theorem span_int_eq_addSubgroup_closure {M : Type*} [AddCommGroup M] (s : Set M) :
    (span  s).toAddSubgroup = AddSubgroup.closure s :=
  Eq.symm <|
    AddSubgroup.closure_eq_of_le _ subset_span fun _ hx =>
      span_induction (fun _ hx => AddSubgroup.subset_closure hx) (AddSubgroup.zero_mem _)
        (fun _ _ _ _ => AddSubgroup.add_mem _) (fun _ _ _ _ => AddSubgroup.zsmul_mem _ ‹_› _) hx
Equality of $\mathbb{Z}$-span and additive subgroup closure
Informal description
For any subset $s$ of an additive commutative group $M$, the additive subgroup generated by the $\mathbb{Z}$-span of $s$ is equal to the additive subgroup closure of $s$. In other words: $$(\operatorname{span}_{\mathbb{Z}} s).\text{toAddSubgroup} = \text{AddSubgroup.closure}(s)$$
Submodule.span_int_eq theorem
{M : Type*} [AddCommGroup M] (s : AddSubgroup M) : (span ℤ (s : Set M)).toAddSubgroup = s
Full source
@[simp]
theorem span_int_eq {M : Type*} [AddCommGroup M] (s : AddSubgroup M) :
    (span  (s : Set M)).toAddSubgroup = s := by rw [span_int_eq_addSubgroup_closure, s.closure_eq]
$\mathbb{Z}$-span of Additive Subgroup Equals Itself
Informal description
For any additive subgroup $s$ of an additive commutative group $M$, the $\mathbb{Z}$-span of $s$ (viewed as a subset of $M$) is equal to $s$ itself. In other words: $$(\operatorname{span}_{\mathbb{Z}} s).\text{toAddSubgroup} = s$$
Disjoint.of_span theorem
(hst : Disjoint (span R s) (span R t)) : Disjoint (s \ {0}) t
Full source
theorem _root_.Disjoint.of_span (hst : Disjoint (span R s) (span R t)) :
    Disjoint (s \ {0}) t := by
  rw [disjoint_iff_forall_ne]
  rintro v ⟨hvs, hv0 : v ≠ 0⟩ _ hvt rfl
  exact hv0 <| (disjoint_def.1 hst) v (subset_span hvs) (subset_span hvt)
Disjointness of Spans Implies Disjointness of Nonzero Elements
Informal description
If the spans of two sets $s$ and $t$ in an $R$-module $M$ are disjoint, then the set $s \setminus \{0\}$ is disjoint from $t$.
Disjoint.of_span₀ theorem
(hst : Disjoint (span R s) (span R t)) (h0s : 0 ∉ s) : Disjoint s t
Full source
theorem _root_.Disjoint.of_span₀ (hst : Disjoint (span R s) (span R t)) (h0s : 0 ∉ s) :
    Disjoint s t := by
  rw [← diff_singleton_eq_self h0s]
  exact hst.of_span
Disjointness of spans implies disjointness of sets when zero is excluded
Informal description
Let $R$ be a ring and $M$ an $R$-module. For any subsets $s, t \subseteq M$ such that $0 \notin s$ and the spans $\operatorname{span}_R(s)$ and $\operatorname{span}_R(t)$ are disjoint, the sets $s$ and $t$ themselves are disjoint.
Submodule.gi definition
: GaloisInsertion (@span R M _ _ _) (↑)
Full source
/-- `span` forms a Galois insertion with the coercion from submodule to set. -/
protected def gi : GaloisInsertion (@span R M _ _ _) (↑) where
  choice s _ := span R s
  gc _ _ := span_le
  le_l_u _ := subset_span
  choice_eq _ _ := rfl
Galois insertion between span and submodule coercion
Informal description
The function `Submodule.gi` establishes a Galois insertion between the span function and the coercion from submodules to sets. Specifically, for any ring $R$ and $R$-module $M$, the span function `span R` and the coercion `(↑)` form a Galois insertion, meaning: 1. For any subset $s$ of $M$, the span of $s$ is the smallest submodule containing $s$. 2. For any submodule $p$ of $M$, the coercion of $p$ to a set is contained in $p$ itself. 3. The choice function for the Galois insertion is given by the span function. This implies that the span function is a closure operator on the lattice of submodules of $M$.
Submodule.span_empty theorem
: span R (∅ : Set M) = ⊥
Full source
@[simp]
theorem span_empty : span R ( : Set M) =  :=
  (Submodule.gi R M).gc.l_bot
Span of Empty Set is Trivial Submodule
Informal description
The span of the empty set in an $R$-module $M$ is the trivial submodule, i.e., $\operatorname{span}_R(\emptyset) = \bot$.
Submodule.span_univ theorem
: span R (univ : Set M) = ⊤
Full source
@[simp]
theorem span_univ : span R (univ : Set M) =  :=
  eq_top_iff.2 <| SetLike.le_def.2 <| subset_span
Span of Universal Set Equals Entire Module
Informal description
The span of the universal set in an $R$-module $M$ is equal to the entire module, i.e., $\operatorname{span}_R (\text{univ} : \text{Set } M) = \top$.
Submodule.span_union theorem
(s t : Set M) : span R (s ∪ t) = span R s ⊔ span R t
Full source
theorem span_union (s t : Set M) : span R (s ∪ t) = spanspan R s ⊔ span R t :=
  (Submodule.gi R M).gc.l_sup
Span of Union Equals Supremum of Spans
Informal description
For any subsets $s$ and $t$ of an $R$-module $M$, the span of their union equals the supremum of their individual spans, i.e., \[ \operatorname{span}_R (s \cup t) = \operatorname{span}_R s \sqcup \operatorname{span}_R t. \]
Submodule.span_iUnion theorem
{ι} (s : ι → Set M) : span R (⋃ i, s i) = ⨆ i, span R (s i)
Full source
theorem span_iUnion {ι} (s : ι → Set M) : span R (⋃ i, s i) = ⨆ i, span R (s i) :=
  (Submodule.gi R M).gc.l_iSup
Span of Union Equals Supremum of Spans
Informal description
For any indexed family of sets $s_i$ in an $R$-module $M$, the span of their union equals the supremum of their individual spans: \[ \operatorname{span}_R \left( \bigcup_i s_i \right) = \bigsqcup_i \operatorname{span}_R (s_i). \]
Submodule.span_iUnion₂ theorem
{ι} {κ : ι → Sort*} (s : ∀ i, κ i → Set M) : span R (⋃ (i) (j), s i j) = ⨆ (i) (j), span R (s i j)
Full source
theorem span_iUnion₂ {ι} {κ : ι → Sort*} (s : ∀ i, κ i → Set M) :
    span R (⋃ (i) (j), s i j) = ⨆ (i) (j), span R (s i j) :=
  (Submodule.gi R M).gc.l_iSup₂
Span of Union Equals Supremum of Spans for Doubly Indexed Family
Informal description
For any indexed family of sets of vectors $\{s_{i,j}\}_{i \in \iota, j \in \kappa_i}$ in an $R$-module $M$, the span of the union $\bigcup_{i,j} s_{i,j}$ is equal to the supremum of the spans of the individual sets $s_{i,j}$. That is, \[ \text{span}_R \left( \bigcup_{i,j} s_{i,j} \right) = \bigsqcup_{i,j} \text{span}_R (s_{i,j}). \]
Submodule.span_attach_biUnion theorem
[DecidableEq M] {α : Type*} (s : Finset α) (f : s → Finset M) : span R (s.attach.biUnion f : Set M) = ⨆ x, span R (f x)
Full source
theorem span_attach_biUnion [DecidableEq M] {α : Type*} (s : Finset α) (f : s → Finset M) :
    span R (s.attach.biUnion f : Set M) = ⨆ x, span R (f x) := by simp [span_iUnion]
Span of Union of Finite Subsets Equals Supremum of Spans
Informal description
Let $M$ be an $R$-module, $\alpha$ a type with decidable equality, and $s$ a finite subset of $\alpha$. For any function $f$ that assigns to each element $x \in s$ a finite subset $f(x) \subseteq M$, the span of the union $\bigcup_{x \in s} f(x)$ is equal to the supremum of the spans of the individual subsets $f(x)$. That is, \[ \operatorname{span}_R \left( \bigcup_{x \in s} f(x) \right) = \bigsqcup_{x \in s} \operatorname{span}_R (f(x)). \]
Submodule.sup_span theorem
: p ⊔ span R s = span R (p ∪ s)
Full source
theorem sup_span : p ⊔ span R s = span R (p ∪ s) := by rw [Submodule.span_union, p.span_eq]
Supremum of Submodule and Span Equals Span of Union
Informal description
For any submodule $p$ of an $R$-module $M$ and any subset $s \subseteq M$, the supremum of $p$ and the span of $s$ equals the span of the union $p \cup s$, i.e., \[ p \sqcup \operatorname{span}_R s = \operatorname{span}_R (p \cup s). \]
Submodule.span_sup theorem
: span R s ⊔ p = span R (s ∪ p)
Full source
theorem span_sup : spanspan R s ⊔ p = span R (s ∪ p) := by rw [Submodule.span_union, p.span_eq]
Span-Supremum Equality: $\operatorname{span}_R s \sqcup p = \operatorname{span}_R (s \cup p)$
Informal description
For any subset $s$ of an $R$-module $M$ and any submodule $p$ of $M$, the supremum of the span of $s$ and $p$ is equal to the span of the union $s \cup p$, i.e., \[ \operatorname{span}_R s \sqcup p = \operatorname{span}_R (s \cup p). \]
Submodule.term_∙_ definition
: Lean.TrailingParserDescr✝
Full source
notation:1000
  /- Note that the character `∙` U+2219 used below is different from the scalar multiplication
character `•` U+2022. -/
R " ∙ " x => span R (singleton x)
Notation for span of a singleton vector
Informal description
The notation \( R \cdot v \) (using the Unicode character U+2219) denotes the span of the singleton set \(\{v\}\) as a submodule over the ring \(R\). This is distinct from the scalar multiplication notation \( \bullet \) (U+2022).
Submodule.span_eq_iSup_of_singleton_spans theorem
(s : Set M) : span R s = ⨆ x ∈ s, R ∙ x
Full source
theorem span_eq_iSup_of_singleton_spans (s : Set M) : span R s = ⨆ x ∈ s, R ∙ x := by
  simp only [← span_iUnion, Set.biUnion_of_singleton s]
Span as Supremum of Singleton Spans
Informal description
For any subset $s$ of an $R$-module $M$, the span of $s$ is equal to the supremum of the spans of all singleton sets $\{x\}$ for $x \in s$. That is, $$\text{span}_R(s) = \bigsqcup_{x \in s} R \cdot x.$$
Submodule.span_range_eq_iSup theorem
{ι : Sort*} {v : ι → M} : span R (range v) = ⨆ i, R ∙ v i
Full source
theorem span_range_eq_iSup {ι : Sort*} {v : ι → M} : span R (range v) = ⨆ i, R ∙ v i := by
  rw [span_eq_iSup_of_singleton_spans, iSup_range]
Span of Range Equals Supremum of Singleton Spans
Informal description
For any family of vectors $\{v_i\}_{i \in \iota}$ in an $R$-module $M$, the span of the range of $v$ is equal to the supremum of the spans of the individual vectors $v_i$. That is, $$\text{span}_R(\text{range}\, v) = \bigsqcup_{i \in \iota} R \cdot v_i.$$
Submodule.span_smul_le theorem
(s : Set M) (r : R) : span R (r • s) ≤ span R s
Full source
theorem span_smul_le (s : Set M) (r : R) : span R (r • s) ≤ span R s := by
  rw [span_le]
  rintro _ ⟨x, hx, rfl⟩
  exact smul_mem (span R s) r (subset_span hx)
Span containment under scalar multiplication: $\operatorname{span}_R (r \cdot s) \leq \operatorname{span}_R s$
Informal description
For any subset $s$ of an $R$-module $M$ and any scalar $r \in R$, the span of the scaled set $r \cdot s$ is contained in the span of $s$, i.e., $$\operatorname{span}_R (r \cdot s) \leq \operatorname{span}_R s.$$
Submodule.subset_span_trans theorem
{U V W : Set M} (hUV : U ⊆ Submodule.span R V) (hVW : V ⊆ Submodule.span R W) : U ⊆ Submodule.span R W
Full source
theorem subset_span_trans {U V W : Set M} (hUV : U ⊆ Submodule.span R V)
    (hVW : V ⊆ Submodule.span R W) : U ⊆ Submodule.span R W :=
  (Submodule.gi R M).gc.le_u_l_trans hUV hVW
Transitivity of Span Containment
Informal description
For any subsets $U$, $V$, and $W$ of an $R$-module $M$, if $U$ is contained in the span of $V$ and $V$ is contained in the span of $W$, then $U$ is contained in the span of $W$. In other words, if $U \subseteq \operatorname{span}_R V$ and $V \subseteq \operatorname{span}_R W$, then $U \subseteq \operatorname{span}_R W$.
Submodule.coe_iSup_of_directed theorem
{ι} [Nonempty ι] (S : ι → Submodule R M) (H : Directed (· ≤ ·) S) : ((iSup S : Submodule R M) : Set M) = ⋃ i, S i
Full source
@[simp]
theorem coe_iSup_of_directed {ι} [Nonempty ι] (S : ι → Submodule R M)
    (H : Directed (· ≤ ·) S) : ((iSup S : Submodule R M) : Set M) = ⋃ i, S i :=
  let s : Submodule R M :=
    { __ := AddSubmonoid.copy _ _ (AddSubmonoid.coe_iSup_of_directed H).symm
      smul_mem' := fun r _ hx ↦ have ⟨i, hi⟩ := Set.mem_iUnion.mp hx
        Set.mem_iUnion.mpr ⟨i, (S i).smul_mem' r hi⟩ }
  have : iSup S = s := le_antisymm
    (iSup_le fun i ↦ le_iSup (fun i ↦ (S i : Set M)) i) (Set.iUnion_subset fun _ ↦ le_iSup S _)
  this.symmrfl
Supremum of Directed Family of Submodules Equals Union of Their Underlying Sets
Informal description
Let $R$ be a ring and $M$ an $R$-module. Given a nonempty index set $\iota$ and a directed family of submodules $(S_i)_{i \in \iota}$ (directed with respect to inclusion $\leq$), the underlying set of the supremum $\bigsqcup_{i} S_i$ is equal to the union $\bigcup_{i} S_i$ of the underlying sets.
Submodule.mem_iSup_of_directed theorem
{ι} [Nonempty ι] (S : ι → Submodule R M) (H : Directed (· ≤ ·) S) {x} : x ∈ iSup S ↔ ∃ i, x ∈ S i
Full source
@[simp]
theorem mem_iSup_of_directed {ι} [Nonempty ι] (S : ι → Submodule R M) (H : Directed (· ≤ ·) S) {x} :
    x ∈ iSup Sx ∈ iSup S ↔ ∃ i, x ∈ S i := by
  rw [← SetLike.mem_coe, coe_iSup_of_directed S H, mem_iUnion]
  rfl
Characterization of Membership in Directed Supremum of Submodules: $x \in \bigsqcup_i S_i \leftrightarrow \exists i, x \in S_i$
Informal description
Let $R$ be a ring and $M$ an $R$-module. Given a nonempty index set $\iota$ and a directed family of submodules $(S_i)_{i \in \iota}$ (directed with respect to inclusion $\leq$), an element $x \in M$ belongs to the supremum $\bigsqcup_{i} S_i$ if and only if there exists an index $i \in \iota$ such that $x \in S_i$.
Submodule.mem_sSup_of_directed theorem
{s : Set (Submodule R M)} {z} (hs : s.Nonempty) (hdir : DirectedOn (· ≤ ·) s) : z ∈ sSup s ↔ ∃ y ∈ s, z ∈ y
Full source
theorem mem_sSup_of_directed {s : Set (Submodule R M)} {z} (hs : s.Nonempty)
    (hdir : DirectedOn (· ≤ ·) s) : z ∈ sSup sz ∈ sSup s ↔ ∃ y ∈ s, z ∈ y := by
  have : Nonempty s := hs.to_subtype
  simp only [sSup_eq_iSup', mem_iSup_of_directed _ hdir.directed_val, SetCoe.exists, Subtype.coe_mk,
    exists_prop]
Characterization of Membership in Directed Supremum of Submodules: $z \in \bigsqcup s \leftrightarrow \exists y \in s, z \in y$
Informal description
Let $R$ be a ring and $M$ an $R$-module. Given a nonempty set $s$ of submodules of $M$ that is directed with respect to the inclusion relation $\leq$, an element $z \in M$ belongs to the supremum $\bigsqcup s$ if and only if there exists a submodule $y \in s$ such that $z \in y$.
Submodule.coe_iSup_of_chain theorem
(a : ℕ →o Submodule R M) : (↑(⨆ k, a k) : Set M) = ⋃ k, (a k : Set M)
Full source
@[norm_cast, simp]
theorem coe_iSup_of_chain (a : ℕ →o Submodule R M) : (↑(⨆ k, a k) : Set M) = ⋃ k, (a k : Set M) :=
  coe_iSup_of_directed a a.monotone.directed_le
Supremum of Chain of Submodules Equals Union of Their Underlying Sets
Informal description
Let $R$ be a ring and $M$ an $R$-module. Given a chain of submodules $(a_k)_{k \in \mathbb{N}}$ (i.e., a monotone sequence of submodules), the underlying set of the supremum $\bigsqcup_{k} a_k$ is equal to the union $\bigcup_{k} a_k$ of the underlying sets of the submodules in the chain.
Submodule.mem_iSup_of_chain theorem
(a : ℕ →o Submodule R M) (m : M) : (m ∈ ⨆ k, a k) ↔ ∃ k, m ∈ a k
Full source
@[simp]
theorem mem_iSup_of_chain (a : ℕ →o Submodule R M) (m : M) : (m ∈ ⨆ k, a k) ↔ ∃ k, m ∈ a k :=
  mem_iSup_of_directed a a.monotone.directed_le
Characterization of Membership in Supremum of Chain of Submodules: $m \in \bigsqcup_k a_k \leftrightarrow \exists k, m \in a_k$
Informal description
Let $R$ be a ring and $M$ an $R$-module. Given a chain of submodules $(a_k)_{k \in \mathbb{N}}$ (i.e., a monotone sequence of submodules), an element $m \in M$ belongs to the supremum $\bigsqcup_{k} a_k$ if and only if there exists an index $k \in \mathbb{N}$ such that $m \in a_k$.
Submodule.mem_sup theorem
: x ∈ p ⊔ p' ↔ ∃ y ∈ p, ∃ z ∈ p', y + z = x
Full source
theorem mem_sup : x ∈ p ⊔ p'x ∈ p ⊔ p' ↔ ∃ y ∈ p, ∃ z ∈ p', y + z = x :=
  ⟨fun h => by
    rw [← span_eq p, ← span_eq p', ← span_union] at h
    refine span_induction ?_ ?_ ?_ ?_ h
    · rintro y (h | h)
      · exact ⟨y, h, 0, by simp, by simp⟩
      · exact ⟨0, by simp, y, h, by simp⟩
    · exact ⟨0, by simp, 0, by simp⟩
    · rintro _ _ - - ⟨y₁, hy₁, z₁, hz₁, rfl⟩ ⟨y₂, hy₂, z₂, hz₂, rfl⟩
      exact ⟨_, add_mem hy₁ hy₂, _, add_mem hz₁ hz₂, by
        rw [add_assoc, add_assoc, ← add_assoc y₂, ← add_assoc z₁, add_comm y₂]⟩
    · rintro a - _ ⟨y, hy, z, hz, rfl⟩
      exact ⟨_, smul_mem _ a hy, _, smul_mem _ a hz, by simp [smul_add]⟩, by
    rintro ⟨y, hy, z, hz, rfl⟩
    exact add_mem ((le_sup_left : p ≤ p ⊔ p') hy) ((le_sup_right : p' ≤ p ⊔ p') hz)⟩
Membership in Supremum of Submodules via Sum Decomposition
Informal description
For any submodules $p$ and $p'$ of an $R$-module $M$, an element $x \in M$ belongs to the supremum $p \sqcup p'$ if and only if there exist elements $y \in p$ and $z \in p'$ such that $x = y + z$.
Submodule.mem_sup' theorem
: x ∈ p ⊔ p' ↔ ∃ (y : p) (z : p'), (y : M) + z = x
Full source
theorem mem_sup' : x ∈ p ⊔ p'x ∈ p ⊔ p' ↔ ∃ (y : p) (z : p'), (y : M) + z = x :=
  mem_sup.trans <| by simp only [Subtype.exists, exists_prop]
Membership in Supremum of Submodules via Sum Decomposition (Bundled Version)
Informal description
For any submodules $p$ and $p'$ of an $R$-module $M$, an element $x \in M$ belongs to the supremum $p \sqcup p'$ if and only if there exist elements $y \in p$ and $z \in p'$ such that $x = y + z$.
Submodule.exists_add_eq_of_codisjoint theorem
(h : Codisjoint p p') (x : M) : ∃ y ∈ p, ∃ z ∈ p', y + z = x
Full source
lemma exists_add_eq_of_codisjoint (h : Codisjoint p p') (x : M) :
    ∃ y ∈ p, ∃ z ∈ p', y + z = x := by
  suffices x ∈ p ⊔ p' by exact Submodule.mem_sup.mp this
  simpa only [h.eq_top] using Submodule.mem_top
Decomposition in Codisjoint Submodules via Sum
Informal description
For any codisjoint submodules $p$ and $p'$ of an $R$-module $M$, and for any element $x \in M$, there exist elements $y \in p$ and $z \in p'$ such that $x = y + z$.
Submodule.coe_sup theorem
: ↑(p ⊔ p') = (p + p' : Set M)
Full source
theorem coe_sup : ↑(p ⊔ p') = (p + p' : Set M) := by
  ext
  rw [SetLike.mem_coe, mem_sup, Set.mem_add]
  simp
Supremum of Submodules as Set Sum: $\overline{p \sqcup p'} = p + p'$
Informal description
For any submodules $p$ and $p'$ of an $R$-module $M$, the underlying set of the supremum $p \sqcup p'$ is equal to the set sum $p + p'$, i.e., $\{x + y \mid x \in p, y \in p'\}$.
Submodule.sup_toAddSubmonoid theorem
: (p ⊔ p').toAddSubmonoid = p.toAddSubmonoid ⊔ p'.toAddSubmonoid
Full source
theorem sup_toAddSubmonoid : (p ⊔ p').toAddSubmonoid = p.toAddSubmonoid ⊔ p'.toAddSubmonoid := by
  ext x
  rw [mem_toAddSubmonoid, mem_sup, AddSubmonoid.mem_sup]
  rfl
Additive Submonoid of Supremum of Submodules Equals Supremum of Additive Submonoids
Informal description
For any two submodules $p$ and $p'$ of an $R$-module $M$, the additive submonoid associated with the supremum $p \sqcup p'$ is equal to the supremum of the additive submonoids associated with $p$ and $p'$. In other words, $(p \sqcup p').\text{toAddSubmonoid} = p.\text{toAddSubmonoid} \sqcup p'.\text{toAddSubmonoid}$.
Submodule.sup_eq_top_iff theorem
: p ⊔ p' = ⊤ ↔ ∀ m : M, ∃ u ∈ p, ∃ v ∈ p', m = u + v
Full source
theorem sup_eq_top_iff : p ⊔ p'p ⊔ p' = ⊤ ↔ ∀ m : M, ∃ u ∈ p, ∃ v ∈ p', m = u + v := by
  rw [eq_top_iff']
  refine forall_congr' fun m ↦ ?_
  rw [mem_sup]
  tauto
Supremum of Submodules Equals Entire Module if and only if Every Element is a Sum of Elements from Each Submodule
Informal description
For any two submodules $p$ and $p'$ of an $R$-module $M$, the supremum $p \sqcup p'$ equals the entire module $M$ if and only if every element $m \in M$ can be expressed as a sum $m = u + v$ for some $u \in p$ and $v \in p'$.
Submodule.mem_span_singleton_self theorem
(x : M) : x ∈ R ∙ x
Full source
theorem mem_span_singleton_self (x : M) : x ∈ R ∙ x :=
  subset_span rfl
Self-Containment in Span of Singleton: $x \in R \cdot x$
Informal description
For any vector $x$ in an $R$-module $M$, the vector $x$ is contained in the span of the singleton set $\{x\}$, i.e., $x \in R \cdot x$.
Submodule.nontrivial_span_singleton theorem
{x : M} (h : x ≠ 0) : Nontrivial (R ∙ x)
Full source
theorem nontrivial_span_singleton {x : M} (h : x ≠ 0) : Nontrivial (R ∙ x) :=
  ⟨by
    use 0, ⟨x, Submodule.mem_span_singleton_self x⟩
    intro H
    rw [eq_comm, Submodule.mk_eq_zero] at H
    exact h H⟩
Nontriviality of the Span of a Nonzero Vector
Informal description
For any nonzero vector $x$ in an $R$-module $M$, the span of the singleton set $\{x\}$ (denoted $R \cdot x$) is a nontrivial submodule, meaning it contains at least two distinct elements.
Submodule.mem_span_singleton theorem
{y : M} : (x ∈ R ∙ y) ↔ ∃ a : R, a • y = x
Full source
theorem mem_span_singleton {y : M} : (x ∈ R ∙ y) ↔ ∃ a : R, a • y = x :=
  ⟨fun h => by
    refine span_induction ?_ ?_ ?_ ?_ h
    · rintro y (rfl | ⟨⟨_⟩⟩)
      exact ⟨1, by simp⟩
    · exact ⟨0, by simp⟩
    · rintro _ _ - - ⟨a, rfl⟩ ⟨b, rfl⟩
      exact ⟨a + b, by simp [add_smul]⟩
    · rintro a _ - ⟨b, rfl⟩
      exact ⟨a * b, by simp [smul_smul]⟩, by
    rintro ⟨a, y, rfl⟩; exact smul_mem _ _ (subset_span <| by simp)⟩
Characterization of Membership in Span of Singleton: $x \in R \cdot y \leftrightarrow \exists a \in R, a \cdot y = x$
Informal description
For any vector $y$ in an $R$-module $M$, a vector $x$ belongs to the span of the singleton set $\{y\}$ (denoted $R \cdot y$) if and only if there exists a scalar $a \in R$ such that $a \cdot y = x$.
Submodule.le_span_singleton_iff theorem
{s : Submodule R M} {v₀ : M} : (s ≤ R ∙ v₀) ↔ ∀ v ∈ s, ∃ r : R, r • v₀ = v
Full source
theorem le_span_singleton_iff {s : Submodule R M} {v₀ : M} :
    (s ≤ R ∙ v₀) ↔ ∀ v ∈ s, ∃ r : R, r • v₀ = v := by simp_rw [SetLike.le_def, mem_span_singleton]
Submodule Containment in Span of Singleton iff Vectors are Scalar Multiples
Informal description
For a submodule $s$ of an $R$-module $M$ and a vector $v_0 \in M$, the submodule $s$ is contained in the span of $\{v_0\}$ if and only if every vector $v \in s$ can be expressed as a scalar multiple of $v_0$, i.e., there exists $r \in R$ such that $r \cdot v_0 = v$.
Submodule.span_singleton_eq_top_iff theorem
(x : M) : (R ∙ x) = ⊤ ↔ ∀ v, ∃ r : R, r • x = v
Full source
theorem span_singleton_eq_top_iff (x : M) : (R ∙ x) = ⊤ ↔ ∀ v, ∃ r : R, r • x = v := by
  rw [eq_top_iff, le_span_singleton_iff]
  tauto
Span of Singleton Equals Entire Module iff Every Vector is a Scalar Multiple
Informal description
For a vector $x$ in an $R$-module $M$, the span of the singleton set $\{x\}$ is equal to the entire module $M$ if and only if every vector $v \in M$ can be expressed as a scalar multiple of $x$, i.e., there exists $r \in R$ such that $r \cdot x = v$.
Submodule.span_zero_singleton theorem
: (R ∙ (0 : M)) = ⊥
Full source
@[simp]
theorem span_zero_singleton : (R ∙ (0 : M)) =  := by
  ext
  simp [mem_span_singleton, eq_comm]
Span of Zero Vector is Trivial Submodule
Informal description
The span of the singleton set containing the zero vector in an $R$-module $M$ is equal to the trivial submodule, i.e., $\operatorname{span}_R \{0\} = \bot$.
Submodule.span_singleton_eq_range theorem
(y : M) : ↑(R ∙ y) = range ((· • y) : R → M)
Full source
theorem span_singleton_eq_range (y : M) : ↑(R ∙ y) = range ((· • y) : R → M) :=
  Set.ext fun _ => mem_span_singleton
Span of Singleton Equals Range of Scalar Multiplication Map: $R \cdot y = \text{range}(r \mapsto r \cdot y)$
Informal description
For any vector $y$ in an $R$-module $M$, the underlying set of the span of the singleton $\{y\}$ (denoted $R \cdot y$) is equal to the range of the scalar multiplication map $r \mapsto r \cdot y$ from $R$ to $M$.
Submodule.span_singleton_smul_le theorem
{S} [SMul S R] [SMul S M] [IsScalarTower S R M] (r : S) (x : M) : (R ∙ r • x) ≤ R ∙ x
Full source
theorem span_singleton_smul_le {S} [SMul S R] [SMul S M] [IsScalarTower S R M]
    (r : S) (x : M) : (R ∙ r • x) ≤ R ∙ x := by
  rw [span_le, Set.singleton_subset_iff, SetLike.mem_coe]
  exact smul_of_tower_mem _ _ (mem_span_singleton_self _)
Span containment under scalar multiplication: $R \cdot (r \cdot x) \subseteq R \cdot x$
Informal description
Let $S$ be a type with scalar multiplication operations on $R$ and $M$, such that the scalar multiplications are compatible via the tower condition `IsScalarTower S R M`. For any scalar $r \in S$ and any vector $x \in M$, the span of the singleton $\{r \cdot x\}$ is contained in the span of the singleton $\{x\}$, i.e., $$ R \cdot (r \cdot x) \subseteq R \cdot x. $$
Submodule.span_singleton_group_smul_eq theorem
{G} [Group G] [SMul G R] [MulAction G M] [IsScalarTower G R M] (g : G) (x : M) : (R ∙ g • x) = R ∙ x
Full source
theorem span_singleton_group_smul_eq {G} [Group G] [SMul G R] [MulAction G M] [IsScalarTower G R M]
    (g : G) (x : M) : (R ∙ g • x) = R ∙ x := by
  refine le_antisymm (span_singleton_smul_le R g x) ?_
  convert span_singleton_smul_le R g⁻¹ (g • x)
  exact (inv_smul_smul g x).symm
Span Equality under Group Action: $R \cdot (g \cdot x) = R \cdot x$
Informal description
Let $G$ be a group with a scalar multiplication operation on $R$ and a multiplicative action on $M$, such that the scalar multiplications are compatible via the tower condition `IsScalarTower G R M`. For any group element $g \in G$ and any vector $x \in M$, the span of the singleton $\{g \cdot x\}$ is equal to the span of the singleton $\{x\}$, i.e., $$ R \cdot (g \cdot x) = R \cdot x. $$
Submodule.span_singleton_smul_eq theorem
{r : R} (hr : IsUnit r) (x : M) : (R ∙ r • x) = R ∙ x
Full source
theorem span_singleton_smul_eq {r : R} (hr : IsUnit r) (x : M) : (R ∙ r • x) = R ∙ x := by
  lift r to  using hr
  rw [← Units.smul_def]
  exact span_singleton_group_smul_eq R r x
Span Equality under Scalar Multiplication by a Unit: $R \cdot (r \cdot x) = R \cdot x$
Informal description
For any scalar $r$ in a ring $R$ that is a unit (i.e., invertible), and any vector $x$ in an $R$-module $M$, the span of the singleton $\{r \cdot x\}$ is equal to the span of the singleton $\{x\}$, i.e., $$ R \cdot (r \cdot x) = R \cdot x. $$
Submodule.mem_span_singleton_trans theorem
{x y z : M} (hxy : x ∈ R ∙ y) (hyz : y ∈ R ∙ z) : x ∈ R ∙ z
Full source
theorem mem_span_singleton_trans {x y z : M} (hxy : x ∈ R ∙ y) (hyz : y ∈ R ∙ z) : x ∈ R ∙ z := by
  rw [← SetLike.mem_coe, ← singleton_subset_iff] at *
  exact Submodule.subset_span_trans hxy hyz
Transitivity of Membership in Span of Singletons: $x \in R \cdot y \land y \in R \cdot z \to x \in R \cdot z$
Informal description
For any vectors $x, y, z$ in an $R$-module $M$, if $x$ is in the span of $\{y\}$ and $y$ is in the span of $\{z\}$, then $x$ is in the span of $\{z\}$. In other words, if $x \in R \cdot y$ and $y \in R \cdot z$, then $x \in R \cdot z$.
Submodule.span_insert theorem
(x) (s : Set M) : span R (insert x s) = (R ∙ x) ⊔ span R s
Full source
theorem span_insert (x) (s : Set M) : span R (insert x s) = (R ∙ x) ⊔ span R s := by
  rw [insert_eq, span_union]
Span of Insertion Equals Supremum of Singleton Span and Set Span: $\operatorname{span}_R (\{x\} \cup s) = (R \cdot x) \sqcup \operatorname{span}_R s$
Informal description
For any element $x$ in an $R$-module $M$ and any subset $s \subseteq M$, the span of the set obtained by inserting $x$ into $s$ is equal to the supremum of the span of the singleton $\{x\}$ and the span of $s$, i.e., \[ \operatorname{span}_R (\{x\} \cup s) = (R \cdot x) \sqcup \operatorname{span}_R s. \]
Submodule.span_insert_eq_span theorem
(h : x ∈ span R s) : span R (insert x s) = span R s
Full source
theorem span_insert_eq_span (h : x ∈ span R s) : span R (insert x s) = span R s :=
  span_eq_of_le _ (Set.insert_subset_iff.mpr ⟨h, subset_span⟩) (span_mono <| subset_insert _ _)
Span Unchanged by Insertion of Dependent Element: $\operatorname{span}_R (\{x\} \cup s) = \operatorname{span}_R s$ when $x \in \operatorname{span}_R s$
Informal description
For any element $x$ in an $R$-module $M$ and any subset $s \subseteq M$, if $x$ is in the span of $s$ (i.e., $x \in \operatorname{span}_R s$), then the span of the set obtained by inserting $x$ into $s$ is equal to the span of $s$, i.e., $\operatorname{span}_R (\{x\} \cup s) = \operatorname{span}_R s$.
Submodule.span_span theorem
: span R (span R s : Set M) = span R s
Full source
theorem span_span : span R (span R s : Set M) = span R s :=
  span_eq _
Idempotence of Span Operation: $\operatorname{span}_R (\operatorname{span}_R s) = \operatorname{span}_R s$
Informal description
For any subset $s$ of an $R$-module $M$, the span of the span of $s$ (viewed as a subset of $M$) is equal to the span of $s$, i.e., $\operatorname{span}_R (\operatorname{span}_R s) = \operatorname{span}_R s$.
Submodule.mem_span_insert theorem
{y} : x ∈ span R (insert y s) ↔ ∃ a : R, ∃ z ∈ span R s, x = a • y + z
Full source
theorem mem_span_insert {y} :
    x ∈ span R (insert y s)x ∈ span R (insert y s) ↔ ∃ a : R, ∃ z ∈ span R s, x = a • y + z := by
  simp [span_insert, mem_sup, mem_span_singleton, eq_comm (a := x)]
Characterization of Vectors in the Span of an Inserted Set: $x \in \text{span}_R(\{y\} \cup s) \leftrightarrow \exists a \in R, \exists z \in \text{span}_R s, x = a y + z$
Informal description
For any vectors $x, y$ in an $R$-module $M$ and any subset $s \subseteq M$, the vector $x$ belongs to the span of the set $\{y\} \cup s$ if and only if there exists a scalar $a \in R$ and a vector $z$ in the span of $s$ such that $x = a \cdot y + z$.
Submodule.mem_span_pair theorem
{x y z : M} : z ∈ span R ({ x, y } : Set M) ↔ ∃ a b : R, a • x + b • y = z
Full source
theorem mem_span_pair {x y z : M} :
    z ∈ span R ({x, y} : Set M)z ∈ span R ({x, y} : Set M) ↔ ∃ a b : R, a • x + b • y = z := by
  simp_rw [mem_span_insert, mem_span_singleton, exists_exists_eq_and, eq_comm]
Characterization of Vectors in the Span of a Pair: $z \in \text{span}_R\{x, y\} \leftrightarrow \exists a, b \in R, z = a x + b y$
Informal description
For any vectors $x, y, z$ in an $R$-module $M$, the vector $z$ lies in the span of the set $\{x, y\}$ if and only if there exist scalars $a, b \in R$ such that $z = a \cdot x + b \cdot y$.
Submodule.span_eq_bot theorem
: span R (s : Set M) = ⊥ ↔ ∀ x ∈ s, (x : M) = 0
Full source
@[simp]
theorem span_eq_bot : spanspan R (s : Set M) = ⊥ ↔ ∀ x ∈ s, (x : M) = 0 :=
  eq_bot_iff.trans
    ⟨fun H _ h => (mem_bot R).1 <| H <| subset_span h, fun H =>
      span_le.2 fun x h => (mem_bot R).2 <| H x h⟩
Span is Trivial if and only if All Vectors are Zero
Informal description
For any subset $s$ of an $R$-module $M$, the span of $s$ is equal to the trivial submodule $\{0\}$ if and only if every element $x \in s$ is the zero vector, i.e., $$\operatorname{span}_R s = 0 \leftrightarrow \forall x \in s, x = 0.$$
Submodule.span_singleton_eq_bot theorem
: (R ∙ x) = ⊥ ↔ x = 0
Full source
theorem span_singleton_eq_bot : (R ∙ x) = ⊥ ↔ x = 0 := by simp
Span of Singleton is Trivial if and only if Element is Zero
Informal description
For any element $x$ in an $R$-module $M$, the span of the singleton set $\{x\}$ is equal to the trivial submodule $\{0\}$ if and only if $x$ is the zero element, i.e., $\mathrm{span}_R(\{x\}) = 0 \leftrightarrow x = 0$.
Submodule.span_zero theorem
: span R (0 : Set M) = ⊥
Full source
@[simp]
theorem span_zero : span R (0 : Set M) =  := by rw [← singleton_zero, span_singleton_eq_bot]
Span of Zero Vector is Trivial Submodule
Informal description
The span of the singleton set containing the zero vector in an $R$-module $M$ is equal to the trivial submodule, i.e., $\mathrm{span}_R(\{0\}) = \{0\}$.
Submodule.span_singleton_le_iff_mem theorem
(m : M) (p : Submodule R M) : (R ∙ m) ≤ p ↔ m ∈ p
Full source
@[simp]
theorem span_singleton_le_iff_mem (m : M) (p : Submodule R M) : (R ∙ m) ≤ p ↔ m ∈ p := by
  rw [span_le, singleton_subset_iff, SetLike.mem_coe]
Span of Singleton is Contained in Submodule if and only if Element Belongs to Submodule
Informal description
For any element $m$ in an $R$-module $M$ and any submodule $p$ of $M$, the span of the singleton set $\{m\}$ is contained in $p$ if and only if $m$ is an element of $p$. In other words: $$R \cdot m \leq p \leftrightarrow m \in p$$
Submodule.iSup_span theorem
{ι : Sort*} (p : ι → Set M) : ⨆ i, span R (p i) = span R (⋃ i, p i)
Full source
theorem iSup_span {ι : Sort*} (p : ι → Set M) : ⨆ i, span R (p i) = span R (⋃ i, p i) :=
  le_antisymm (iSup_le fun i => span_mono <| subset_iUnion _ i) <|
    span_le.mpr <| iUnion_subset fun i _ hm => mem_iSup_of_mem i <| subset_span hm
Supremum of Spans Equals Span of Union
Informal description
For any family of subsets $\{p_i\}_{i \in \iota}$ of an $R$-module $M$, the supremum of the spans of the $p_i$ equals the span of their union. That is: $$\bigsqcup_i \operatorname{span}_R p_i = \operatorname{span}_R \left( \bigcup_i p_i \right)$$
Submodule.iSup_eq_span theorem
{ι : Sort*} (p : ι → Submodule R M) : ⨆ i, p i = span R (⋃ i, ↑(p i))
Full source
theorem iSup_eq_span {ι : Sort*} (p : ι → Submodule R M) : ⨆ i, p i = span R (⋃ i, ↑(p i)) := by
  simp_rw [← iSup_span, span_eq]
Supremum of Submodules Equals Span of Their Union
Informal description
For any family of submodules $\{p_i\}_{i \in \iota}$ of an $R$-module $M$, the supremum of the submodules $p_i$ equals the span of the union of their underlying sets. That is: $$\bigsqcup_i p_i = \operatorname{span}_R \left( \bigcup_i p_i \right)$$
Submodule.submodule_eq_sSup_le_nonzero_spans theorem
(p : Submodule R M) : p = sSup {T : Submodule R M | ∃ m ∈ p, m ≠ 0 ∧ T = span R { m }}
Full source
/-- A submodule is equal to the supremum of the spans of the submodule's nonzero elements. -/
theorem submodule_eq_sSup_le_nonzero_spans (p : Submodule R M) :
    p = sSup { T : Submodule R M | ∃ m ∈ p, m ≠ 0 ∧ T = span R {m} } := by
  let S := { T : Submodule R M | ∃ m ∈ p, m ≠ 0 ∧ T = span R {m} }
  apply le_antisymm
  · intro m hm
    by_cases h : m = 0
    · rw [h]
      simp
    · exact @le_sSup _ _ S _ ⟨m, ⟨hm, ⟨h, rfl⟩⟩⟩ m (mem_span_singleton_self m)
  · rw [sSup_le_iff]
    rintro S ⟨_, ⟨_, ⟨_, rfl⟩⟩⟩
    rwa [span_singleton_le_iff_mem]
Submodule as Supremum of Spans of its Nonzero Elements
Informal description
For any submodule $p$ of an $R$-module $M$, $p$ is equal to the supremum of the spans of all nonzero elements in $p$. In other words: $$ p = \sup \{ \text{span}_R \{m\} \mid m \in p, m \neq 0 \} $$
Submodule.lt_sup_iff_not_mem theorem
{I : Submodule R M} {a : M} : (I < I ⊔ R ∙ a) ↔ a ∉ I
Full source
theorem lt_sup_iff_not_mem {I : Submodule R M} {a : M} : (I < I ⊔ R ∙ a) ↔ a ∉ I := by simp
Strict Supremum Condition for Submodule and Span of a Singleton
Informal description
For a submodule $I$ of an $R$-module $M$ and an element $a \in M$, the submodule $I$ is strictly contained in the join of $I$ and the span of $\{a\}$ if and only if $a$ does not belong to $I$. In other words, $I < I \sqcup R \cdot a$ if and only if $a \notin I$.
Submodule.mem_iSup theorem
{ι : Sort*} (p : ι → Submodule R M) {m : M} : (m ∈ ⨆ i, p i) ↔ ∀ N, (∀ i, p i ≤ N) → m ∈ N
Full source
theorem mem_iSup {ι : Sort*} (p : ι → Submodule R M) {m : M} :
    (m ∈ ⨆ i, p i) ↔ ∀ N, (∀ i, p i ≤ N) → m ∈ N := by
  rw [← span_singleton_le_iff_mem, le_iSup_iff]
  simp only [span_singleton_le_iff_mem]
Characterization of Membership in Supremum of Submodules via Containment
Informal description
For any family of submodules $(p_i)_{i \in \iota}$ of an $R$-module $M$ and any element $m \in M$, the element $m$ belongs to the supremum $\bigsqcup_i p_i$ of the family if and only if for every submodule $N$ of $M$ that contains all $p_i$, we have $m \in N$.
Submodule.mem_sSup theorem
{s : Set (Submodule R M)} {m : M} : (m ∈ sSup s) ↔ ∀ N, (∀ p ∈ s, p ≤ N) → m ∈ N
Full source
theorem mem_sSup {s : Set (Submodule R M)} {m : M} :
    (m ∈ sSup s) ↔ ∀ N, (∀ p ∈ s, p ≤ N) → m ∈ N := by
  simp_rw [sSup_eq_iSup, Submodule.mem_iSup, iSup_le_iff]
Characterization of Membership in Supremum of Submodules
Informal description
For any set $s$ of submodules of a module $M$ over a ring $R$, and any element $m \in M$, we have $m \in \bigsqcup s$ if and only if for every submodule $N$ of $M$ that contains all submodules in $s$, it follows that $m \in N$.
Submodule.mem_span_finite_of_mem_span theorem
{S : Set M} {x : M} (hx : x ∈ span R S) : ∃ T : Finset M, ↑T ⊆ S ∧ x ∈ span R (T : Set M)
Full source
/-- For every element in the span of a set, there exists a finite subset of the set
such that the element is contained in the span of the subset. -/
theorem mem_span_finite_of_mem_span {S : Set M} {x : M} (hx : x ∈ span R S) :
    ∃ T : Finset M, ↑T ⊆ S ∧ x ∈ span R (T : Set M) := by
  classical
  refine span_induction (fun x hx => ?_) ?_ ?_ ?_ hx
  · refine ⟨{x}, ?_, ?_⟩
    · rwa [Finset.coe_singleton, Set.singleton_subset_iff]
    · rw [Finset.coe_singleton]
      exact Submodule.mem_span_singleton_self x
  · use ∅
    simp
  · rintro x y - - ⟨X, hX, hxX⟩ ⟨Y, hY, hyY⟩
    refine ⟨X ∪ Y, ?_, ?_⟩
    · rw [Finset.coe_union]
      exact Set.union_subset hX hY
    rw [Finset.coe_union, span_union, mem_sup]
    exact ⟨x, hxX, y, hyY, rfl⟩
  · rintro a x - ⟨T, hT, h2⟩
    exact ⟨T, hT, smul_mem _ _ h2⟩
Finite Span Property: Every Element in a Span is in the Span of a Finite Subset
Informal description
For any subset $S$ of an $R$-module $M$ and any element $x \in \operatorname{span}_R S$, there exists a finite subset $T \subseteq S$ such that $x \in \operatorname{span}_R T$.
Submodule.sup_toAddSubgroup theorem
(p p' : Submodule R M) : (p ⊔ p').toAddSubgroup = p.toAddSubgroup ⊔ p'.toAddSubgroup
Full source
lemma sup_toAddSubgroup (p p' : Submodule R M) :
    (p ⊔ p').toAddSubgroup = p.toAddSubgroup ⊔ p'.toAddSubgroup := by
  ext x
  rw [mem_toAddSubgroup, mem_sup, AddSubgroup.mem_sup]
  rfl
Supremum of Submodules Preserves Underlying Additive Subgroup Structure
Informal description
For any two submodules $p$ and $p'$ of an $R$-module $M$, the underlying additive subgroup of their supremum $p \sqcup p'$ is equal to the supremum of their underlying additive subgroups, i.e., \[ (p \sqcup p').\text{toAddSubgroup} = p.\text{toAddSubgroup} \sqcup p'.\text{toAddSubgroup}. \]
Submodule.mem_span_insert' theorem
{x y} {s : Set M} : x ∈ span R (insert y s) ↔ ∃ a : R, x + a • y ∈ span R s
Full source
theorem mem_span_insert' {x y} {s : Set M} :
    x ∈ span R (insert y s)x ∈ span R (insert y s) ↔ ∃ a : R, x + a • y ∈ span R s := by
  rw [mem_span_insert]; constructor
  · rintro ⟨a, z, hz, rfl⟩
    exact ⟨-a, by simp [hz, add_assoc]⟩
  · rintro ⟨a, h⟩
    exact ⟨-a, _, h, by simp [add_comm, add_left_comm]⟩
Characterization of Vectors in the Span of an Inserted Set: $x \in \text{span}_R(\{y\} \cup s) \leftrightarrow \exists a \in R, x + a y \in \text{span}_R s$
Informal description
For any vectors $x, y$ in an $R$-module $M$ and any subset $s \subseteq M$, the vector $x$ belongs to the span of the set $\{y\} \cup s$ if and only if there exists a scalar $a \in R$ such that $x + a \cdot y$ belongs to the span of $s$.
Submodule.span_range_update_add_smul theorem
(hij : i ≠ j) (v : ι → M) (r : R) : span R (Set.range (Function.update v j (v j + r • v i))) = span R (Set.range v)
Full source
lemma span_range_update_add_smul (hij : i ≠ j) (v : ι → M) (r : R) :
    span R (Set.range (Function.update v j (v j + r • v i))) = span R (Set.range v) := by
  refine le_antisymm ?_ ?_ <;> simp only [span_le, Set.range_subset_iff, SetLike.mem_coe] <;>
    intro k <;> obtain rfl | hjk := eq_or_ne j k
  · rw [update_self]
    exact add_mem (subset_span ⟨j, rfl⟩) <| smul_mem _ _ <| subset_span ⟨i, rfl⟩
  · exact subset_span ⟨k, (update_of_ne hjk.symm ..).symm⟩
  · nth_rw 2 [← add_sub_cancel_right (v j) (r • v i)]
    exact sub_mem (subset_span ⟨j, update_self ..⟩)
      (smul_mem _ _ (subset_span ⟨i, update_of_ne hij ..⟩))
  · exact subset_span ⟨k, update_of_ne hjk.symm ..⟩
Span Preservation under Vector Update with Scalar Multiplication
Informal description
Let $i$ and $j$ be distinct indices (i.e., $i \neq j$), $v$ be a family of vectors in an $R$-module $M$, and $r$ be a scalar in $R$. Then the span of the range of the updated family $v'$, where $v'_j = v_j + r \cdot v_i$ and $v'_k = v_k$ for $k \neq j$, is equal to the span of the range of the original family $v$. That is, \[ \operatorname{span}_R \left( \operatorname{range} (v') \right) = \operatorname{span}_R \left( \operatorname{range} (v) \right), \] where $v'$ is defined as above.
Submodule.span_range_update_sub_smul theorem
(hij : i ≠ j) (v : ι → M) (r : R) : span R (Set.range (Function.update v j (v j - r • v i))) = span R (Set.range v)
Full source
lemma span_range_update_sub_smul (hij : i ≠ j) (v : ι → M) (r : R) :
    span R (Set.range (Function.update v j (v j - r • v i))) = span R (Set.range v) := by
  rw [sub_eq_add_neg, ← neg_smul, span_range_update_add_smul hij]
Span Preservation under Vector Update with Negative Scalar Multiplication
Informal description
Let $i$ and $j$ be distinct indices (i.e., $i \neq j$), $v$ be a family of vectors in an $R$-module $M$, and $r$ be a scalar in $R$. Then the span of the range of the updated family $v'$, where $v'_j = v_j - r \cdot v_i$ and $v'_k = v_k$ for $k \neq j$, is equal to the span of the range of the original family $v$. That is, \[ \operatorname{span}_R \left( \operatorname{range} (v') \right) = \operatorname{span}_R \left( \operatorname{range} (v) \right), \] where $v'$ is defined as above.