Skip to content

Commit

Permalink
feat(linear_algebra/finite_dimensional): generalisations to division_…
Browse files Browse the repository at this point in the history
…ring (#8822)

I generalise a few results about finite dimensional modules from fields to division rings. Mostly this is me trying out @alexjbest's `generalisation_linter`. (review: it works really well, and is very helpful for finding the right home for lemmas, but it is slow).



Co-authored-by: Scott Morrison <scott.morrison@gmail.com>
  • Loading branch information
semorrison and semorrison committed Sep 7, 2021
1 parent 9c886ff commit 463e753
Show file tree
Hide file tree
Showing 5 changed files with 208 additions and 181 deletions.
2 changes: 1 addition & 1 deletion src/analysis/normed_space/finite_dimension.lean
Expand Up @@ -112,7 +112,7 @@ begin
{ left,
have : finrank 𝕜 f.range = 1,
{ refine le_antisymm _ (zero_lt_iff.mpr H),
simpa [finrank_of_field] using f.range.finrank_le },
simpa [finrank_self] using f.range.finrank_le },
rw [this, add_comm, nat.add_one] at Z,
exact nat.succ.inj Z } },
have : is_closed (f.ker : set E),
Expand Down
2 changes: 1 addition & 1 deletion src/field_theory/finite/polynomial.lean
Expand Up @@ -159,7 +159,7 @@ calc module.rank K (R σ K) =
linear_equiv.dim_eq
(finsupp.supported_equiv_finsupp {s : σ →₀ ℕ | ∀n:σ, s n ≤ fintype.card K - 1 })
... = cardinal.mk {s : σ →₀ ℕ | ∀ (n : σ), s n ≤ fintype.card K - 1} :
by rw [finsupp.dim_eq, dim_of_ring, mul_one]
by rw [finsupp.dim_eq, dim_self, mul_one]
... = cardinal.mk {s : σ → ℕ | ∀ (n : σ), s n < fintype.card K } :
begin
refine quotient.sound ⟨equiv.subtype_equiv finsupp.equiv_fun_on_fintype $ assume f, _⟩,
Expand Down
138 changes: 74 additions & 64 deletions src/linear_algebra/dimension.lean
Expand Up @@ -227,6 +227,10 @@ begin
apply dim_range_le,
end

theorem dim_quotient_le (p : submodule R M) :
module.rank R p.quotient ≤ module.rank R M :=
(mkq p).dim_le_of_surjective (surjective_quot_mk _)

variables [nontrivial R]

lemma {m} cardinal_lift_le_dim_of_linear_independent
Expand Down Expand Up @@ -717,6 +721,13 @@ theorem basis.mk_range_eq_dim (v : basis ι R M) :
cardinal.mk (range v) = module.rank R M :=
v.reindex_range.mk_eq_dim''

/-- If a vector space has a finite basis, then its dimension (seen as a cardinal) is equal to the
cardinality of the basis. -/
lemma dim_eq_card_basis {ι : Type w} [fintype ι] (h : basis ι R M) :
module.rank R M = fintype.card ι :=
by rw [←h.mk_range_eq_dim, cardinal.fintype_card,
set.card_range_of_injective h.injective]

theorem basis.mk_eq_dim (v : basis ι R M) :
cardinal.lift.{w v} (cardinal.mk ι) = cardinal.lift.{v w} (module.rank R M) :=
by rw [←v.mk_range_eq_dim, cardinal.mk_range_eq_of_injective v.injective]
Expand Down Expand Up @@ -757,7 +768,7 @@ by { rw [← @set_of_mem_eq _ s, ← subtype.range_coe_subtype], exact dim_span

variables (R)

lemma dim_of_ring : module.rank R R = 1 :=
lemma dim_self : module.rank R R = 1 :=
by rw [←cardinal.lift_inj, ← (basis.singleton punit R).mk_eq_dim, cardinal.mk_punit]

end strong_rank_condition
Expand Down Expand Up @@ -852,6 +863,68 @@ begin
⟨equiv.ulift.trans (equiv.sum_congr equiv.ulift equiv.ulift).symm ⟩),
end

-- TODO the remainder of this section should generalize beyond division rings.

lemma dim_zero_iff_forall_zero : module.rank K V = 0 ↔ ∀ x : V, x = 0 :=
begin
split,
{ intros h x,
have card_mk_range := (basis.of_vector_space K V).mk_range_eq_dim,
rw [h, cardinal.mk_emptyc_iff, coe_of_vector_space, subtype.range_coe] at card_mk_range,
simpa [card_mk_range] using (of_vector_space K V).mem_span x },
{ intro h,
have : (⊤ : submodule K V) = ⊥,
{ ext x, simp [h x] },
rw [←dim_top, this, dim_bot] }
end

lemma dim_zero_iff : module.rank K V = 0 ↔ subsingleton V :=
dim_zero_iff_forall_zero.trans (subsingleton_iff_forall_eq 0).symm

lemma dim_pos_iff_exists_ne_zero : 0 < module.rank K V ↔ ∃ x : V, x ≠ 0 :=
begin
rw ←not_iff_not,
simpa using dim_zero_iff_forall_zero
end

lemma dim_pos_iff_nontrivial : 0 < module.rank K V ↔ nontrivial V :=
dim_pos_iff_exists_ne_zero.trans (nontrivial_iff_exists_ne 0).symm

lemma dim_pos [h : nontrivial V] : 0 < module.rank K V :=
dim_pos_iff_nontrivial.2 h


section fintype
variable [fintype η]
variables [∀i, add_comm_group (φ i)] [∀i, module K (φ i)]

open linear_map

lemma dim_pi : module.rank K (Πi, φ i) = cardinal.sum (λi, module.rank K (φ i)) :=
begin
let b := assume i, basis.of_vector_space K (φ i),
let this : basis (Σ j, _) K (Π j, φ j) := pi.basis b,
rw [←cardinal.lift_inj, ← this.mk_eq_dim],
simp [λ i, (b i).mk_range_eq_dim.symm, cardinal.sum_mk]
end

lemma dim_fun {V η : Type u} [fintype η] [add_comm_group V] [module K V] :
module.rank K (η → V) = fintype.card η * module.rank K V :=
by rw [dim_pi, cardinal.sum_const, cardinal.fintype_card]

lemma dim_fun_eq_lift_mul :
module.rank K (η → V) = (fintype.card η : cardinal.{max u₁' v}) *
cardinal.lift.{v u₁'} (module.rank K V) :=
by rw [dim_pi, cardinal.sum_const_eq_lift_mul, cardinal.fintype_card, cardinal.lift_nat_cast]

lemma dim_fun' : module.rank K (η → K) = fintype.card η :=
by rw [dim_fun_eq_lift_mul, dim_self, cardinal.lift_one, mul_one, cardinal.nat_cast_inj]

lemma dim_fin_fun (n : ℕ) : module.rank K (fin n → K) = n :=
by simp [dim_fun']

end fintype

end division_ring

section field
Expand All @@ -863,10 +936,6 @@ theorem dim_quotient_add_dim (p : submodule K V) :
module.rank K p.quotient + module.rank K p = module.rank K V :=
by classical; exact let ⟨f⟩ := quotient_prod_linear_equiv p in dim_prod.symm.trans f.dim_eq

theorem dim_quotient_le (p : submodule K V) :
module.rank K p.quotient ≤ module.rank K V :=
by { rw ← dim_quotient_add_dim p, exact self_le_add_right _ _ }

/-- rank-nullity theorem -/
theorem dim_range_add_dim_ker (f : V →ₗ[K] V₁) :
module.rank K f.range + module.rank K f.ker = module.rank K V :=
Expand Down Expand Up @@ -944,37 +1013,6 @@ by { rw [← dim_sup_add_dim_inf_eq], exact self_le_add_right _ _ }

end

section fintype
variable [fintype η]
variables [∀i, add_comm_group (φ i)] [∀i, module K (φ i)]

open linear_map

lemma dim_pi : module.rank K (Πi, φ i) = cardinal.sum (λi, module.rank K (φ i)) :=
begin
let b := assume i, basis.of_vector_space K (φ i),
let this : basis (Σ j, _) K (Π j, φ j) := pi.basis b,
rw [←cardinal.lift_inj, ← this.mk_eq_dim],
simp [λ i, (b i).mk_range_eq_dim.symm, cardinal.sum_mk]
end

lemma dim_fun {V η : Type u} [fintype η] [add_comm_group V] [module K V] :
module.rank K (η → V) = fintype.card η * module.rank K V :=
by rw [dim_pi, cardinal.sum_const, cardinal.fintype_card]

lemma dim_fun_eq_lift_mul :
module.rank K (η → V) = (fintype.card η : cardinal.{max u₁' v}) *
cardinal.lift.{v u₁'} (module.rank K V) :=
by rw [dim_pi, cardinal.sum_const_eq_lift_mul, cardinal.fintype_card, cardinal.lift_nat_cast]

lemma dim_fun' : module.rank K (η → K) = fintype.card η :=
by rw [dim_fun_eq_lift_mul, dim_of_ring, cardinal.lift_one, mul_one, cardinal.nat_cast_inj]

lemma dim_fin_fun (n : ℕ) : module.rank K (fin n → K) = n :=
by simp [dim_fun']

end fintype

lemma exists_mem_ne_zero_of_ne_bot {s : submodule K V} (h : s ≠ ⊥) : ∃ b : V, b ∈ s ∧ b ≠ 0 :=
begin
classical,
Expand Down Expand Up @@ -1035,21 +1073,6 @@ end rank

-- TODO The remainder of this file could be generalized to arbitrary rings.

lemma dim_zero_iff_forall_zero : module.rank K V = 0 ↔ ∀ x : V, x = 0 :=
begin
split,
{ intros h x,
have card_mk_range := (basis.of_vector_space K V).mk_range_eq_dim,
rw [h, cardinal.mk_emptyc_iff, coe_of_vector_space, subtype.range_coe] at card_mk_range,
simpa [card_mk_range] using (of_vector_space K V).mem_span x },
{ intro h,
have : (⊤ : submodule K V) = ⊥,
{ ext x, simp [h x] },
rw [←dim_top, this, dim_bot] }
end

lemma dim_zero_iff : module.rank K V = 0 ↔ subsingleton V :=
dim_zero_iff_forall_zero.trans (subsingleton_iff_forall_eq 0).symm

/-- The `ι` indexed basis on `V`, where `ι` is an empty type and `V` is zero-dimensional.
Expand All @@ -1067,19 +1090,6 @@ end
basis.of_dim_eq_zero hV i = 0 :=
rfl


lemma dim_pos_iff_exists_ne_zero : 0 < module.rank K V ↔ ∃ x : V, x ≠ 0 :=
begin
rw ←not_iff_not,
simpa using dim_zero_iff_forall_zero
end

lemma dim_pos_iff_nontrivial : 0 < module.rank K V ↔ nontrivial V :=
dim_pos_iff_exists_ne_zero.trans (nontrivial_iff_exists_ne 0).symm

lemma dim_pos [h : nontrivial V] : 0 < module.rank K V :=
dim_pos_iff_nontrivial.2 h

lemma le_dim_iff_exists_linear_independent {c : cardinal} :
c ≤ module.rank K V ↔ ∃ s : set V, cardinal.mk s = c ∧ linear_independent K (coe : s → V) :=
begin
Expand Down

0 comments on commit 463e753

Please sign in to comment.