Lemmas about rank and finrank in rings satisfying strong rank condition. #
Main statements #
For modules over rings satisfying the rank condition
- Basis.le_span: the cardinality of a basis is bounded by the cardinality of any spanning set
For modules over rings satisfying the strong rank condition
- linearIndependent_le_span: For any linearly independent family- v : ι → Mand any finite spanning set- w : Set M, the cardinality of- ιis bounded by the cardinality of- w.
- linearIndependent_le_basis: If- bis a basis for a module- M, and- sis a linearly independent set, then the cardinality of- sis bounded by the cardinality of- b.
For modules over rings with invariant basis number (including all commutative rings and all noetherian rings)
- mk_eq_mk_of_basis: the dimension theorem, any two bases of the same vector space have the same cardinality.
The dimension theorem: if v and v' are two bases, their index types
have the same cardinalities.
Given two bases indexed by ι and ι' of an R-module, where R satisfies the invariant
basis number property, an equiv ι ≃ ι'.
Equations
- v.indexEquiv v' = ⋯.some
Instances For
An auxiliary lemma for Basis.le_span.
If R satisfies the rank condition,
then for any finite basis b : Basis ι R M,
and any finite spanning set w : Set M,
the cardinality of ι is bounded by the cardinality of w.
Another auxiliary lemma for Basis.le_span, which does not require assuming the basis is finite,
but still assumes we have a finite spanning set.
If R satisfies the rank condition,
then the cardinality of any basis is bounded by the cardinality of any spanning set.
If R satisfies the strong rank condition,
then any linearly independent family v : ι → M
contained in the span of some finite w : Set M,
is itself finite.
If R satisfies the strong rank condition,
then for any linearly independent family v : ι → M
contained in the span of some finite w : Set M,
the cardinality of ι is bounded by the cardinality of w.
If R satisfies the strong rank condition,
then for any linearly independent family v : ι → M
and any finite spanning set w : Set M,
the cardinality of ι is bounded by the cardinality of w.
A version of linearIndependent_le_span for Finset.
An auxiliary lemma for linearIndependent_le_basis:
we handle the case where the basis b is infinite.
Over any ring R satisfying the strong rank condition,
if b is a basis for a module M,
and s is a linearly independent set,
then the cardinality of s is bounded by the cardinality of b.
Let R satisfy the strong rank condition. If m elements of a free rank n R-module are
linearly independent, then m ≤ n.
Over any ring R satisfying the strong rank condition,
if b is an infinite basis for a module M,
then every maximal linearly independent set has the same cardinality as b.
This proof (along with some of the lemmas above) comes from [Les familles libres maximales d'un module ont-elles le meme cardinal?][lazarus1973]
If a vector space has a finite basis, then its dimension (seen as a cardinal) is equal to the cardinality of the basis.
An induction (and recursion) principle for proving results about all submodules of a fixed
finite free module M. A property is true for all submodules of M if it satisfies the following
"inductive step": the property is true for a submodule N if it's true for all submodules N'
of N with the property that there exists 0 ≠ x ∈ N such that the sum N' + Rx is direct.
Equations
- Submodule.inductionOnRank b P ih N = Submodule.inductionOnRankAux b P ih (Fintype.card ι) N ⋯
Instances For
If S a module-finite free R-algebra, then the R-rank of a nonzero R-free
ideal I of S is the same as the rank of S.
If a vector space (or module) has a finite basis, then its dimension (or rank) is equal to the cardinality of the basis.
If a free module is of finite rank, then the cardinality of any basis is equal to its
finrank.
If a vector space (or module) has a finite basis, then its dimension (or rank) is equal to the
cardinality of the basis. This lemma uses a Finset instead of indexed types.
A ring satisfying StrongRankCondition (such as a DivisionRing) is one-dimensional as a
module over itself.
Given a basis of a ring over itself indexed by a type ι, then ι is Unique.
Equations
- Basis.unique R b = let_fun A := ⋯; ⋯.some
Instances For
The rank of a finite module is finite.
Alias of rank_lt_aleph0.
The rank of a finite module is finite.
If M is finite, finrank M = rank M.
Alias of finrank_eq_rank.
If M is finite, finrank M = rank M.