Some constructions of matroids #
This file defines some very elementary examples of matroids, namely those with at most one base.
Main definitions #
emptyOn α
is the matroid onα
with empty ground set.
For E : Set α
, ...
loopyOn E
is the matroid onE
whose elements are all loops, or equivalently in which∅
is the only base.freeOn E
is the 'free matroid' whose ground setE
is the only base.- For
I ⊆ E
,uniqueBaseOn I E
is the matroid with ground setE
in whichI
is the only base.
Implementation details #
To avoid the tedious process of certifying the matroid axioms for each of these easy examples,
we bootstrap the definitions starting with emptyOn α
(which simp
can prove is a matroid)
and then construct the other examples using duality and restriction.
The Matroid α
with empty ground set.
Equations
- One or more equations did not get rendered due to their size.
Instances For
The Matroid α
with ground set E
whose only base is ∅
Equations
- Matroid.loopyOn E = (Matroid.emptyOn α).restrict E
Instances For
The Matroid α
with ground set E
whose only base is E
.
Equations
- Matroid.freeOn E = (Matroid.loopyOn E)✶
Instances For
The matroid on E
whose unique base is the subset I
of E
.
Intended for use when I ⊆ E
; if this not not the case, then the base is I ∩ E
.
Equations
- Matroid.uniqueBaseOn I E = (Matroid.freeOn I).restrict E
Instances For
@[simp]
@[simp]
theorem
Matroid.uniqueBaseOn_inter_isBasis
{α : Type u_1}
{E I X : Set α}
(hX : X ⊆ E)
:
(uniqueBaseOn I E).IsBasis (X ∩ I) X
@[simp]
@[simp]
@[simp]
theorem
Matroid.uniqueBaseOn_rankFinite
{α : Type u_1}
{E I : Set α}
(hI : I.Finite)
:
(uniqueBaseOn I E).RankFinite
theorem
Matroid.uniqueBaseOn_rankPos
{α : Type u_1}
{E I : Set α}
(hIE : I ⊆ E)
(hI : I.Nonempty)
:
(uniqueBaseOn I E).RankPos