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 ground set E
whose only base is ∅
Equations
- Matroid.loopyOn E = (Matroid.emptyOn α).restrict E
Instances For
theorem
Matroid.Finite.loopyOn_finite
{α : Type u_1}
{E : Set α}
(hE : E.Finite)
:
(loopyOn E).Finite
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]
theorem
Matroid.uniqueBaseOn_base_iff
{α : Type u_1}
{E B I : Set α}
(hIE : I ⊆ E)
:
(uniqueBaseOn I E).Base B ↔ B = I
theorem
Matroid.uniqueBaseOn_inter_ground_eq
{α : Type u_1}
(I E : Set α)
:
uniqueBaseOn (I ∩ E) E = uniqueBaseOn I E
@[simp]
theorem
Matroid.uniqueBaseOn_indep_iff'
{α : Type u_1}
{E I J : Set α}
:
(uniqueBaseOn I E).Indep J ↔ J ⊆ I ∩ E
theorem
Matroid.uniqueBaseOn_indep_iff
{α : Type u_1}
{E I J : Set α}
(hIE : I ⊆ E)
:
(uniqueBaseOn I E).Indep J ↔ J ⊆ I
theorem
Matroid.uniqueBaseOn_basis_iff
{α : Type u_1}
{E I X J : Set α}
(hX : X ⊆ E)
:
(uniqueBaseOn I E).Basis J X ↔ J = X ∩ I
theorem
Matroid.uniqueBaseOn_inter_basis
{α : Type u_1}
{E I X : Set α}
(hX : X ⊆ E)
:
(uniqueBaseOn I E).Basis (X ∩ I) X
@[simp]
theorem
Matroid.uniqueBaseOn_dual_eq
{α : Type u_1}
(I E : Set α)
:
(uniqueBaseOn I E)✶ = uniqueBaseOn (E \ I) E
@[simp]
@[simp]
theorem
Matroid.uniqueBaseOn_restrict'
{α : Type u_1}
(I E R : Set α)
:
(uniqueBaseOn I E).restrict R = uniqueBaseOn (I ∩ R ∩ E) R
theorem
Matroid.uniqueBaseOn_restrict
{α : Type u_1}
{E I : Set α}
(h : I ⊆ E)
(R : Set α)
:
(uniqueBaseOn I E).restrict R = uniqueBaseOn (I ∩ R) R