Skip to content

Commit cabd20e

Browse files
committed
feat: Rank-nullity theorem for commutative domains (#9412)
Co-authored-by: Andrew Yang <36414270+erdOne@users.noreply.github.com>
1 parent ac9e42e commit cabd20e

File tree

8 files changed

+307
-99
lines changed

8 files changed

+307
-99
lines changed

Mathlib.lean

Lines changed: 2 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -334,6 +334,7 @@ import Mathlib.Algebra.Module.Submodule.Basic
334334
import Mathlib.Algebra.Module.Submodule.Bilinear
335335
import Mathlib.Algebra.Module.Submodule.Lattice
336336
import Mathlib.Algebra.Module.Submodule.LinearMap
337+
import Mathlib.Algebra.Module.Submodule.Localization
337338
import Mathlib.Algebra.Module.Submodule.Map
338339
import Mathlib.Algebra.Module.Submodule.Pointwise
339340
import Mathlib.Algebra.Module.Torsion
@@ -2394,6 +2395,7 @@ import Mathlib.LinearAlgebra.Dimension.Finite
23942395
import Mathlib.LinearAlgebra.Dimension.Finrank
23952396
import Mathlib.LinearAlgebra.Dimension.Free
23962397
import Mathlib.LinearAlgebra.Dimension.LinearMap
2398+
import Mathlib.LinearAlgebra.Dimension.Localization
23972399
import Mathlib.LinearAlgebra.Dimension.StrongRankCondition
23982400
import Mathlib.LinearAlgebra.DirectSum.Finsupp
23992401
import Mathlib.LinearAlgebra.DirectSum.TensorProduct

Mathlib/Algebra/Module/LocalizedModule.lean

Lines changed: 13 additions & 5 deletions
Original file line numberDiff line numberDiff line change
@@ -551,7 +551,11 @@ variable {R : Type*} [CommSemiring R] (S : Submonoid R)
551551

552552
variable {M M' M'' : Type*} [AddCommMonoid M] [AddCommMonoid M'] [AddCommMonoid M'']
553553

554-
variable [Module R M] [Module R M'] [Module R M''] (f : M →ₗ[R] M') (g : M →ₗ[R] M'')
554+
variable {A : Type*} [CommSemiring A] [Algebra R A] [Module A M'] [IsLocalization S A]
555+
556+
variable [Module R M] [Module R M'] [Module R M''] [IsScalarTower R A M']
557+
558+
variable (f : M →ₗ[R] M') (g : M →ₗ[R] M'')
555559

556560
/-- The characteristic predicate for localized module.
557561
`IsLocalizedModule S f` describes that `f : M ⟶ M'` is the localization map identifying `M'` as
@@ -1045,6 +1049,14 @@ theorem mk_eq_mk' (s : S) (m : M) :
10451049
LocalizedModule.mk_cancel, LocalizedModule.mkLinearMap_apply]
10461050
#align is_localized_module.mk_eq_mk' IsLocalizedModule.mk_eq_mk'
10471051

1052+
variable (A) in
1053+
lemma mk'_smul_mk' (x : R) (m : M) (s t : S) :
1054+
IsLocalization.mk' A x s • mk' f m t = mk' f (x • m) (s * t) := by
1055+
apply smul_injective f (s * t)
1056+
conv_lhs => simp only [smul_assoc, mul_smul, smul_comm t]
1057+
simp only [mk'_cancel', map_smul, Submonoid.smul_def s]
1058+
rw [← smul_assoc, IsLocalization.smul_mk'_self, algebraMap_smul]
1059+
10481060
variable (S)
10491061

10501062
theorem eq_zero_iff {m : M} : f m = 0 ↔ ∃ s' : S, s' • m = 0 :=
@@ -1089,10 +1101,6 @@ theorem mkOfAlgebra {R S S' : Type*} [CommRing R] [CommRing S] [CommRing S'] [Al
10891101

10901102
end Algebra
10911103

1092-
variable {A : Type*}
1093-
[CommSemiring A] [Algebra R A] [Module A M'] [IsScalarTower R A M'] [IsLocalization S A]
1094-
1095-
10961104
/-- If `(f : M →ₗ[R] M')` is a localization of modules, then the map
10971105
`(localization S) × M → N, (s, m) ↦ s • f m` is the tensor product (insomuch as it is the universal
10981106
bilinear map).
Lines changed: 115 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,115 @@
1+
/-
2+
Copyright (c) 2024 Andrew Yang. All rights reserved.
3+
Released under Apache 2.0 license as described in the file LICENSE.
4+
Authors: Andrew Yang
5+
-/
6+
import Mathlib.Algebra.Module.LocalizedModule
7+
8+
/-!
9+
# Localization of Submodules
10+
11+
Results about localizations of submodules and quotient modules are provided in this file.
12+
13+
## Main result
14+
- `Submodule.localized`:
15+
The localization of an `R`-submodule of `M` at `p` viewed as an `Rₚ`-submodule of `Mₚ`.
16+
- `Submodule.toLocalized`:
17+
The localization map of a submodule `M' →ₗ[R] M'.localized p`.
18+
- `Submodule.toLocalizedQuotient`:
19+
The localization map of a quotient module `M ⧸ M' →ₗ[R] LocalizedModule p M ⧸ M'.localized p`.
20+
21+
## TODO
22+
- Statements regarding the exactness of localization.
23+
- Connection with flatness.
24+
25+
-/
26+
27+
open nonZeroDivisors
28+
29+
universe u u' v v'
30+
31+
variable {R : Type u} (S : Type u') {M : Type v} {N : Type v'}
32+
variable [CommRing R] [CommRing S] [AddCommGroup M] [AddCommGroup N] [Module R M]
33+
variable [Module R M] [Module R N] [Algebra R S] [Module S N] [IsScalarTower R S N]
34+
variable (p : Submonoid R) [IsLocalization p S] (f : M →ₗ[R] N) [IsLocalizedModule p f]
35+
variable (hp : p ≤ R⁰)
36+
37+
variable (M' : Submodule R M)
38+
39+
/-- Let `S` be the localization of `R` at `p` and `N` be the localization of `M` at `p`.
40+
This is the localization of an `R`-submodule of `M` viewed as an `S`-submodule of `N`. -/
41+
def Submodule.localized' : Submodule S N where
42+
carrier := { x | ∃ m ∈ M', ∃ s : p, IsLocalizedModule.mk' f m s = x }
43+
add_mem' := fun {x} {y} ⟨m, hm, s, hx⟩ ⟨n, hn, t, hy⟩ ↦ ⟨t • m + s • n, add_mem (M'.smul_mem t hm)
44+
(M'.smul_mem s hn), s * t, by rw [← hx, ← hy, IsLocalizedModule.mk'_add_mk']⟩
45+
zero_mem' := ⟨0, zero_mem _, 1, by simp⟩
46+
smul_mem' := fun r x h ↦ by
47+
have ⟨m, hm, s, hx⟩ := h
48+
have ⟨y, t, hyt⟩ := IsLocalization.mk'_surjective p r
49+
exact ⟨y • m, M'.smul_mem y hm, t * s, by simp [← hyt, ← hx, IsLocalizedModule.mk'_smul_mk']⟩
50+
51+
/-- The localization of an `R`-submodule of `M` at `p` viewed as an `Rₚ`-submodule of `Mₚ`. -/
52+
abbrev Submodule.localized : Submodule (Localization p) (LocalizedModule p M) :=
53+
M'.localized' (Localization p) p (LocalizedModule.mkLinearMap p M)
54+
55+
/-- The localization map of a submodule. -/
56+
@[simps!]
57+
def Submodule.toLocalized' : M' →ₗ[R] M'.localized' S p f :=
58+
f.restrict (q := (M'.localized' S p f).restrictScalars R) (fun x hx ↦ ⟨x, hx, 1, by simp⟩)
59+
60+
/-- The localization map of a submodule. -/
61+
abbrev Submodule.toLocalized : M' →ₗ[R] M'.localized p :=
62+
M'.toLocalized' (Localization p) p (LocalizedModule.mkLinearMap p M)
63+
64+
instance Submodule.isLocalizedModule : IsLocalizedModule p (M'.toLocalized' S p f) where
65+
map_units x := by
66+
simp_rw [Module.End_isUnit_iff]
67+
constructor
68+
· exact fun _ _ e ↦ Subtype.ext
69+
(IsLocalizedModule.smul_injective f x (congr_arg Subtype.val e))
70+
· rintro m
71+
use (IsLocalization.mk' S 1 x) • m
72+
rw [Module.algebraMap_end_apply, ← smul_assoc, IsLocalization.smul_mk'_one,
73+
IsLocalization.mk'_self', one_smul]
74+
surj' := by
75+
rintro ⟨y, x, hx, s, rfl⟩
76+
exact ⟨⟨⟨x, hx⟩, s⟩, by ext; simp⟩
77+
exists_of_eq e := by simpa [Subtype.ext_iff] using
78+
IsLocalizedModule.exists_of_eq (S := p) (f := f) (congr_arg Subtype.val e)
79+
80+
/-- The localization map of a quotient module. -/
81+
def Submodule.toLocalizedQuotient' : M ⧸ M' →ₗ[R] N ⧸ M'.localized' S p f :=
82+
Submodule.mapQ M' ((M'.localized' S p f).restrictScalars R) f (fun x hx ↦ ⟨x, hx, 1, by simp⟩)
83+
84+
/-- The localization map of a quotient module. -/
85+
abbrev Submodule.toLocalizedQuotient : M ⧸ M' →ₗ[R] LocalizedModule p M ⧸ M'.localized p :=
86+
M'.toLocalizedQuotient' (Localization p) p (LocalizedModule.mkLinearMap p M)
87+
88+
@[simp]
89+
lemma Submodule.toLocalizedQuotient'_mk (x : M) :
90+
M'.toLocalizedQuotient' S p f (Submodule.Quotient.mk x) = Submodule.Quotient.mk (f x) := rfl
91+
92+
open Submodule Submodule.Quotient IsLocalization in
93+
instance IsLocalizedModule.toLocalizedQuotient' (M' : Submodule R M) :
94+
IsLocalizedModule p (M'.toLocalizedQuotient' S p f) where
95+
map_units x := by
96+
refine (Module.End_isUnit_iff _).mpr ⟨fun m n e ↦ ?_, fun m ↦ ⟨(IsLocalization.mk' S 1 x) • m,
97+
by rw [Module.algebraMap_end_apply, ← smul_assoc, smul_mk'_one, mk'_self', one_smul]⟩⟩
98+
obtain ⟨⟨m, rfl⟩, n, rfl⟩ := PProd.mk (mk_surjective _ m) (mk_surjective _ n)
99+
simp only [Module.algebraMap_end_apply, ← mk_smul, Submodule.Quotient.eq, ← smul_sub] at e
100+
replace e := Submodule.smul_mem _ (IsLocalization.mk' S 1 x) e
101+
rwa [smul_comm, ← smul_assoc, smul_mk'_one, mk'_self', one_smul, ← Submodule.Quotient.eq] at e
102+
surj' y := by
103+
obtain ⟨y, rfl⟩ := mk_surjective _ y
104+
obtain ⟨⟨y, s⟩, rfl⟩ := IsLocalizedModule.mk'_surjective p f y
105+
exact ⟨⟨Submodule.Quotient.mk y, s⟩,
106+
by simp only [Function.uncurry_apply_pair, toLocalizedQuotient'_mk, ← mk_smul, mk'_cancel']⟩
107+
exists_of_eq {m n} e := by
108+
obtain ⟨⟨m, rfl⟩, n, rfl⟩ := PProd.mk (mk_surjective _ m) (mk_surjective _ n)
109+
obtain ⟨x, hx, s, hs⟩ : f (m - n) ∈ _ := by simpa [Submodule.Quotient.eq] using e
110+
obtain ⟨c, hc⟩ := exists_of_eq (S := p) (show f (s • (m - n)) = f x by simp [-map_sub, ← hs])
111+
exact ⟨c * s, by simpa only [← Quotient.mk_smul, Submodule.Quotient.eq,
112+
← smul_sub, mul_smul, hc] using M'.smul_mem c hx⟩
113+
114+
instance (M' : Submodule R M) : IsLocalizedModule p (M'.toLocalizedQuotient p) :=
115+
IsLocalizedModule.toLocalizedQuotient' _ _ _ _

Mathlib/Algebra/Module/Zlattice.lean

Lines changed: 2 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -464,7 +464,7 @@ theorem Zlattice.rank : finrank ℤ L = finrank K E := by
464464
Finset.sdiff_eq_empty_iff_subset] at h
465465
replace h := Finset.card_le_card h
466466
rwa [not_lt, h_card, ← topEquiv.finrank_eq, ← h_spanE, ← ht_span,
467-
finrank_span_set_eq_card _ ht_lin]
467+
finrank_span_set_eq_card ht_lin]
468468
-- Assume that `e ∪ {v}` is not `ℤ`-linear independent then we get the contradiction
469469
suffices ¬ LinearIndependent ℤ (fun x : ↥(insert v (Set.range e)) => (x : E)) by
470470
contrapose! this
@@ -500,6 +500,6 @@ theorem Zlattice.rank : finrank ℤ L = finrank K E := by
500500
· -- To prove that `finrank K E ≤ finrank ℤ L`, we use the fact `b` generates `E` over `K`
501501
-- and thus `finrank K E ≤ card b = finrank ℤ L`
502502
rw [← topEquiv.finrank_eq, ← h_spanE]
503-
convert finrank_span_le_card (K := K) (Set.range b)
503+
convert finrank_span_le_card (R := K) (Set.range b)
504504

505505
end Zlattice

Mathlib/LinearAlgebra/Dimension/Constructions.lean

Lines changed: 85 additions & 22 deletions
Original file line numberDiff line numberDiff line change
@@ -374,27 +374,6 @@ theorem FiniteDimensional.finrank_tensorProduct :
374374

375375
end TensorProduct
376376

377-
section Span
378-
379-
variable [StrongRankCondition R]
380-
381-
theorem rank_span_le_of_finite {s : Set M} (hs : s.Finite) : Module.rank R (span R s) ≤ #s := by
382-
rw [Module.rank_def]
383-
apply ciSup_le'
384-
rintro ⟨t, ht⟩
385-
letI := hs.fintype
386-
simpa [Cardinal.mk_image_eq Subtype.val_injective] using linearIndependent_le_span' _
387-
(ht.map (f := Submodule.subtype _) (by simp)).image s (fun x ↦ by aesop)
388-
389-
theorem rank_span_finset_le (s : Finset M) : Module.rank R (span R (s : Set M)) ≤ s.card := by
390-
simpa using rank_span_le_of_finite s.finite_toSet
391-
392-
theorem rank_span_of_finset (s : Finset M) : Module.rank R (span R (s : Set M)) < ℵ₀ :=
393-
(rank_span_finset_le s).trans_lt (Cardinal.nat_lt_aleph0 _)
394-
#align rank_span_of_finset rank_span_of_finset
395-
396-
end Span
397-
398377
section SubmoduleRank
399378

400379
section
@@ -449,9 +428,93 @@ theorem Submodule.finrank_le_finrank_of_le {s t : Submodule R M} [Module.Finite
449428

450429
end
451430

452-
453431
end SubmoduleRank
454432

433+
section Span
434+
435+
variable [StrongRankCondition R]
436+
437+
theorem rank_span_le (s : Set M) : Module.rank R (span R s) ≤ #s := by
438+
rw [Finsupp.span_eq_range_total, ← lift_strictMono.le_iff_le]
439+
refine (lift_rank_range_le _).trans ?_
440+
rw [rank_finsupp_self]
441+
simp only [lift_lift, ge_iff_le, le_refl]
442+
#align rank_span_le rank_span_le
443+
444+
theorem rank_span_finset_le (s : Finset M) : Module.rank R (span R (s : Set M)) ≤ s.card := by
445+
simpa using rank_span_le s.toSet
446+
447+
theorem rank_span_of_finset (s : Finset M) : Module.rank R (span R (s : Set M)) < ℵ₀ :=
448+
(rank_span_finset_le s).trans_lt (Cardinal.nat_lt_aleph0 _)
449+
#align rank_span_of_finset rank_span_of_finset
450+
451+
open Submodule FiniteDimensional
452+
453+
variable (R)
454+
455+
/-- The rank of a set of vectors as a natural number. -/
456+
protected noncomputable def Set.finrank (s : Set M) : ℕ :=
457+
finrank R (span R s)
458+
#align set.finrank Set.finrank
459+
460+
variable {R}
461+
462+
theorem finrank_span_le_card (s : Set M) [Fintype s] : finrank R (span R s) ≤ s.toFinset.card :=
463+
finrank_le_of_rank_le (by simpa using rank_span_le (R := R) s)
464+
#align finrank_span_le_card finrank_span_le_card
465+
466+
theorem finrank_span_finset_le_card (s : Finset M) : (s : Set M).finrank R ≤ s.card :=
467+
calc
468+
(s : Set M).finrank R ≤ (s : Set M).toFinset.card := finrank_span_le_card (M := M) s
469+
_ = s.card := by simp
470+
#align finrank_span_finset_le_card finrank_span_finset_le_card
471+
472+
theorem finrank_range_le_card {ι : Type*} [Fintype ι] (b : ι → M) :
473+
(Set.range b).finrank R ≤ Fintype.card ι := by
474+
classical
475+
refine (finrank_span_le_card _).trans ?_
476+
rw [Set.toFinset_range]
477+
exact Finset.card_image_le
478+
#align finrank_range_le_card finrank_range_le_card
479+
480+
theorem finrank_span_eq_card [Nontrivial R] {ι : Type*} [Fintype ι] {b : ι → M}
481+
(hb : LinearIndependent R b) :
482+
finrank R (span R (Set.range b)) = Fintype.card ι :=
483+
finrank_eq_of_rank_eq
484+
(by
485+
have : Module.rank R (span R (Set.range b)) = #(Set.range b) := rank_span hb
486+
rwa [← lift_inj, mk_range_eq_of_injective hb.injective, Cardinal.mk_fintype, lift_natCast,
487+
lift_eq_nat_iff] at this)
488+
#align finrank_span_eq_card finrank_span_eq_card
489+
490+
theorem finrank_span_set_eq_card {s : Set M} [Fintype s] (hs : LinearIndependent R ((↑) : s → M)) :
491+
finrank R (span R s) = s.toFinset.card :=
492+
finrank_eq_of_rank_eq
493+
(by
494+
have : Module.rank R (span R s) = #s := rank_span_set hs
495+
rwa [Cardinal.mk_fintype, ← Set.toFinset_card] at this)
496+
#align finrank_span_set_eq_card finrank_span_set_eq_card
497+
498+
theorem finrank_span_finset_eq_card {s : Finset M} (hs : LinearIndependent R ((↑) : s → M)) :
499+
finrank R (span R (s : Set M)) = s.card := by
500+
convert finrank_span_set_eq_card (s := (s : Set M)) hs
501+
ext
502+
simp
503+
#align finrank_span_finset_eq_card finrank_span_finset_eq_card
504+
505+
theorem span_lt_of_subset_of_card_lt_finrank {s : Set M} [Fintype s] {t : Submodule R M}
506+
(subset : s ⊆ t) (card_lt : s.toFinset.card < finrank R t) : span R s < t :=
507+
lt_of_le_of_finrank_lt_finrank (span_le.mpr subset)
508+
(lt_of_le_of_lt (finrank_span_le_card _) card_lt)
509+
#align span_lt_of_subset_of_card_lt_finrank span_lt_of_subset_of_card_lt_finrank
510+
511+
theorem span_lt_top_of_card_lt_finrank {s : Set M} [Fintype s]
512+
(card_lt : s.toFinset.card < finrank R M) : span R s < ⊤ :=
513+
lt_top_of_finrank_lt_finrank (lt_of_le_of_lt (finrank_span_le_card _) card_lt)
514+
#align span_lt_top_of_card_lt_finrank span_lt_top_of_card_lt_finrank
515+
516+
end Span
517+
455518
section SubalgebraRank
456519

457520
open Module

0 commit comments

Comments
 (0)