Skip to content
New issue

Have a question about this project? Sign up for a free GitHub account to open an issue and contact its maintainers and the community.

By clicking “Sign up for GitHub”, you agree to our terms of service and privacy statement. We’ll occasionally send you account related emails.

Already on GitHub? Sign in to your account

[Merged by Bors] - feat: A lin. indep. family of vectors in a fin. dim. space is finite #9103

Closed
wants to merge 3 commits into from
Closed
Show file tree
Hide file tree
Changes from 2 commits
Commits
File filter

Filter by extension

Filter by extension

Conversations
Failed to load comments.
Jump to
Jump to file
Failed to load files.
Diff view
Diff view
4 changes: 2 additions & 2 deletions Mathlib/Algebra/Module/Zlattice.lean
Original file line number Diff line number Diff line change
Expand Up @@ -388,7 +388,7 @@ theorem Zlattice.FG : AddSubgroup.FG L := by
rw [← hs, ← h_span]
exact span_mono (by simp only [Subtype.range_coe_subtype, Set.setOf_mem_eq, subset_rfl]))
rw [show span ℤ s = span ℤ (Set.range b) by simp [Basis.coe_mk, Subtype.range_coe_subtype]]
have : Fintype s := Set.Finite.fintype h_lind.finite
have : Fintype s := h_lind.setFinite.fintype
refine Set.Finite.of_finite_image (f := ((↑) : _ → E) ∘ Zspan.quotientEquiv b) ?_
(Function.Injective.injOn (Subtype.coe_injective.comp (Zspan.quotientEquiv b).injective) _)
have : Set.Finite ((Zspan.fundamentalDomain b) ∩ L) :=
Expand All @@ -406,7 +406,7 @@ theorem Zlattice.FG : AddSubgroup.FG L := by
exact span_le.mpr h_incl
· -- `span ℤ s` is finitely generated because `s` is finite
rw [ker_mkQ, inf_of_le_right (span_le.mpr h_incl)]
exact fg_span (LinearIndependent.finite h_lind)
exact fg_span (LinearIndependent.setFinite h_lind)

theorem Zlattice.module_finite : Module.Finite ℤ L :=
Module.Finite.iff_addGroup_fg.mpr ((AddGroup.fg_iff_addSubgroup_fg L).mpr (FG K hs))
Expand Down
2 changes: 1 addition & 1 deletion Mathlib/Analysis/InnerProductSpace/PiL2.lean
Original file line number Diff line number Diff line change
Expand Up @@ -805,7 +805,7 @@ theorem Orthonormal.exists_orthonormalBasis_extension (hv : Orthonormal 𝕜 ((
∃ (u : Finset E) (b : OrthonormalBasis u 𝕜 E), v ⊆ u ∧ ⇑b = ((↑) : u → E) := by
obtain ⟨u₀, hu₀s, hu₀, hu₀_max⟩ := exists_maximal_orthonormal hv
rw [maximal_orthonormal_iff_orthogonalComplement_eq_bot hu₀] at hu₀_max
have hu₀_finite : u₀.Finite := hu₀.linearIndependent.finite
have hu₀_finite : u₀.Finite := hu₀.linearIndependent.setFinite
let u : Finset E := hu₀_finite.toFinset
let fu : ↥u ≃ ↥u₀ := hu₀_finite.subtypeEquivToFinset.symm
have hu : Orthonormal 𝕜 ((↑) : u → E) := by simpa using hu₀.comp _ fu.injective
Expand Down
20 changes: 10 additions & 10 deletions Mathlib/LinearAlgebra/FiniteDimensional.lean
Original file line number Diff line number Diff line change
Expand Up @@ -299,18 +299,18 @@ theorem lt_aleph0_of_linearIndependent {ι : Type w} [FiniteDimensional K V] {v
apply Cardinal.nat_lt_aleph0
#align finite_dimensional.lt_aleph_0_of_linear_independent FiniteDimensional.lt_aleph0_of_linearIndependent

theorem _root_.LinearIndependent.finite [FiniteDimensional K V] {b : Set V}
lemma _root_.LinearIndependent.finite {ι : Type*} [FiniteDimensional K V] {f : ι → V}
(h : LinearIndependent K f) : Finite ι :=
Cardinal.lt_aleph0_iff_finite.1 $ FiniteDimensional.lt_aleph0_of_linearIndependent h

theorem not_linearIndependent_of_infinite {ι : Type*} [Infinite ι] [FiniteDimensional K V]
(v : ι → V) : ¬LinearIndependent K v := mt LinearIndependent.finite $ @not_finite _ _
YaelDillies marked this conversation as resolved.
Show resolved Hide resolved
#align finite_dimensional.not_linear_independent_of_infinite FiniteDimensional.not_linearIndependent_of_infinite

theorem _root_.LinearIndependent.setFinite [FiniteDimensional K V] {b : Set V}
(h : LinearIndependent K fun x : b => (x : V)) : b.Finite :=
Cardinal.lt_aleph0_iff_set_finite.mp (FiniteDimensional.lt_aleph0_of_linearIndependent h)
#align linear_independent.finite LinearIndependent.finite

theorem not_linearIndependent_of_infinite {ι : Type w} [inf : Infinite ι] [FiniteDimensional K V]
(v : ι → V) : ¬LinearIndependent K v := by
intro h_lin_indep
have : ¬ℵ₀ ≤ #ι := not_le.mpr (lt_aleph0_of_linearIndependent h_lin_indep)
have : ℵ₀ ≤ #ι := infinite_iff.mp inf
contradiction
#align finite_dimensional.not_linear_independent_of_infinite FiniteDimensional.not_linearIndependent_of_infinite
#align linear_independent.finite LinearIndependent.setFinite

/-- A finite dimensional space has positive `finrank` iff it has a nonzero element. -/
theorem finrank_pos_iff_exists_ne_zero [FiniteDimensional K V] : 0 < finrank K V ↔ ∃ x : V, x ≠ 0 :=
Expand Down