Documentation

Mathlib.Data.Setoid.Partition

Equivalence relations: partitions #

This file comprises properties of equivalence relations viewed as partitions. There are two implementations of partitions here:

Of course both implementations are related to Quotient and Setoid.

Setoid.isPartition.partition and Finpartition.isPartition_parts furnish a link between Setoid.IsPartition and Finpartition.

TODO #

Could the design of Finpartition inform the one of Setoid.IsPartition? Maybe bundling it and changing it from Set (Set α) to Set α where [Lattice α] [OrderBot α] would make it more usable.

Tags #

setoid, equivalence, iseqv, relation, equivalence relation, partition, equivalence class

theorem Setoid.eq_of_mem_eqv_class {α : Type u_1} {c : Set (Set α)} (H : ∀ (a : α), ∃! b : Set α, b c a b) {x : α} {b b' : Set α} (hc : b c) (hb : x b) (hc' : b' c) (hb' : x b') :
b = b'

If x ∈ α is in 2 elements of a set of sets partitioning α, those 2 sets are equal.

def Setoid.mkClasses {α : Type u_1} (c : Set (Set α)) (H : ∀ (a : α), ∃! b : Set α, b c a b) :

Makes an equivalence relation from a set of sets partitioning α.

Equations
def Setoid.classes {α : Type u_1} (r : Setoid α) :
Set (Set α)

Makes the equivalence classes of an equivalence relation.

Equations
theorem Setoid.mem_classes {α : Type u_1} (r : Setoid α) (y : α) :
{x : α | r x y} r.classes
theorem Setoid.classes_ker_subset_fiber_set {α : Type u_1} {β : Type u_2} (f : αβ) :
(ker f).classes Set.range fun (y : β) => {x : α | f x = y}
theorem Setoid.finite_classes_ker {α : Type u_2} {β : Type u_3} [Finite β] (f : αβ) :
theorem Setoid.card_classes_ker_le {α : Type u_2} {β : Type u_3} [Fintype β] (f : αβ) [Fintype (ker f).classes] :
theorem Setoid.eq_iff_classes_eq {α : Type u_1} {r₁ r₂ : Setoid α} :
r₁ = r₂ ∀ (x : α), {y : α | r₁ x y} = {y : α | r₂ x y}

Two equivalence relations are equal iff all their equivalence classes are equal.

theorem Setoid.rel_iff_exists_classes {α : Type u_1} (r : Setoid α) {x y : α} :
r x y cr.classes, x c y c
theorem Setoid.classes_inj {α : Type u_1} {r₁ r₂ : Setoid α} :
r₁ = r₂ r₁.classes = r₂.classes

Two equivalence relations are equal iff their equivalence classes are equal.

theorem Setoid.empty_notMem_classes {α : Type u_1} {r : Setoid α} :
r.classes

The empty set is not an equivalence class.

@[deprecated Setoid.empty_notMem_classes (since := "2025-05-23")]
theorem Setoid.empty_not_mem_classes {α : Type u_1} {r : Setoid α} :
r.classes

Alias of Setoid.empty_notMem_classes.


The empty set is not an equivalence class.

theorem Setoid.classes_eqv_classes {α : Type u_1} {r : Setoid α} (a : α) :
∃! b : Set α, b r.classes a b

Equivalence classes partition the type.

theorem Setoid.eq_of_mem_classes {α : Type u_1} {r : Setoid α} {x : α} {b : Set α} (hc : b r.classes) (hb : x b) {b' : Set α} (hc' : b' r.classes) (hb' : x b') :
b = b'

If x ∈ α is in 2 equivalence classes, the equivalence classes are equal.

theorem Setoid.eq_eqv_class_of_mem {α : Type u_1} {c : Set (Set α)} (H : ∀ (a : α), ∃! b : Set α, b c a b) {s : Set α} {y : α} (hs : s c) (hy : y s) :
s = {x : α | (mkClasses c H) x y}

The elements of a set of sets partitioning α are the equivalence classes of the equivalence relation defined by the set of sets.

theorem Setoid.eqv_class_mem {α : Type u_1} {c : Set (Set α)} (H : ∀ (a : α), ∃! b : Set α, b c a b) {y : α} :
{x : α | (mkClasses c H) x y} c

The equivalence classes of the equivalence relation defined by a set of sets partitioning α are elements of the set of sets.

theorem Setoid.eqv_class_mem' {α : Type u_1} {c : Set (Set α)} (H : ∀ (a : α), ∃! b : Set α, b c a b) {x : α} :
{y : α | (mkClasses c H) x y} c
theorem Setoid.eqv_classes_disjoint {α : Type u_1} {c : Set (Set α)} (H : ∀ (a : α), ∃! b : Set α, b c a b) :

Distinct elements of a set of sets partitioning α are disjoint.

theorem Setoid.eqv_classes_of_disjoint_union {α : Type u_1} {c : Set (Set α)} (hu : ⋃₀ c = Set.univ) (H : c.PairwiseDisjoint id) (a : α) :
∃! b : Set α, b c a b

A set of disjoint sets covering α partition α (classical).

def Setoid.setoidOfDisjointUnion {α : Type u_1} {c : Set (Set α)} (hu : ⋃₀ c = Set.univ) (H : c.PairwiseDisjoint id) :

Makes an equivalence relation from a set of disjoints sets covering α.

Equations
theorem Setoid.mkClasses_classes {α : Type u_1} (r : Setoid α) :

The equivalence relation made from the equivalence classes of an equivalence relation r equals r.

@[simp]
theorem Setoid.sUnion_classes {α : Type u_1} (r : Setoid α) :
noncomputable def Setoid.quotientEquivClasses {α : Type u_1} (r : Setoid α) :

The equivalence between the quotient by an equivalence relation and its type of equivalence classes.

Equations
@[simp]
theorem Setoid.quotientEquivClasses_mk_eq {α : Type u_1} (r : Setoid α) (a : α) :
(r.quotientEquivClasses a) = {x : α | r x a}
def Setoid.IsPartition {α : Type u_1} (c : Set (Set α)) :

A collection c : Set (Set α) of sets is a partition of α into pairwise disjoint sets if ∅ ∉ c and each element a : α belongs to a unique set b ∈ c.

Equations
theorem Setoid.nonempty_of_mem_partition {α : Type u_1} {c : Set (Set α)} (hc : IsPartition c) {s : Set α} (h : s c) :

A partition of α does not contain the empty set.

theorem Set.PairwiseDisjoint.isPartition_of_exists_of_ne_empty {α : Type u_2} {s : Set (Set α)} (h₁ : s.PairwiseDisjoint id) (h₂ : ∀ (a : α), xs, a x) (h₃ : s) :
theorem Setoid.exists_of_mem_partition {α : Type u_1} {c : Set (Set α)} (hc : IsPartition c) {s : Set α} (hs : s c) :
∃ (y : α), s = {x : α | (mkClasses c ) x y}

All elements of a partition of α are the equivalence class of some y ∈ α.

theorem Setoid.classes_mkClasses {α : Type u_1} (c : Set (Set α)) (hc : IsPartition c) :
(mkClasses c ).classes = c

The equivalence classes of the equivalence relation defined by a partition of α equal the original partition.

Defining on partitions as the defined on their induced equivalence relations.

Equations

Defining a partial order on partitions as the partial order on their induced equivalence relations.

Equations

The order-preserving bijection between equivalence relations on a type α, and partitions of α into subsets.

Equations
  • One or more equations did not get rendered due to their size.

A complete lattice instance for partitions; there is more infrastructure for the equivalent complete lattice on equivalence relations.

Equations

A finite setoid partition furnishes a finpartition

Equations
  • hc.finpartition = { parts := c, supIndep := , sup_parts := , bot_notMem := }
@[simp]
theorem Setoid.IsPartition.finpartition_parts {α : Type u_1} {c : Finset (Set α)} (hc : IsPartition c) :

A finpartition gives rise to a setoid partition

structure IndexedPartition {ι : Type u_1} {α : Type u_2} (s : ιSet α) :
Type (max u_1 u_2)

Constructive information associated with a partition of a type α indexed by another type ι, s : ι → Set α.

IndexedPartition.index sends an element to its index, while IndexedPartition.some sends an index to an element of the corresponding set.

This type is primarily useful for definitional control of s - if this is not needed, then Setoid.ker index by itself may be sufficient.

  • eq_of_mem {x : α} {i j : ι} : x s ix s ji = j

    two indexes are equal if they are equal in membership

  • some : ια

    sends an index to an element of the corresponding set

  • some_mem (i : ι) : self.some i s i

    membership invariance for some

  • index : αι

    index for type α

  • mem_index (x : α) : x s (self.index x)

    membership invariance for index

noncomputable def IndexedPartition.mk' {ι : Type u_1} {α : Type u_2} (s : ιSet α) (dis : Pairwise (Function.onFun Disjoint s)) (nonempty : ∀ (i : ι), (s i).Nonempty) (ex : ∀ (x : α), ∃ (i : ι), x s i) :

The non-constructive constructor for IndexedPartition.

Equations
  • IndexedPartition.mk' s dis nonempty ex = { eq_of_mem := , some := fun (i : ι) => .some, some_mem := , index := fun (x : α) => .choose, mem_index := }
instance IndexedPartition.instInhabitedUnivOfUnique {ι : Type u_1} {α : Type u_2} [Unique ι] [Inhabited α] :

On a unique index set there is the obvious trivial partition

Equations
theorem IndexedPartition.exists_mem {ι : Type u_1} {α : Type u_2} {s : ιSet α} (hs : IndexedPartition s) (x : α) :
∃ (i : ι), x s i
theorem IndexedPartition.iUnion {ι : Type u_1} {α : Type u_2} {s : ιSet α} (hs : IndexedPartition s) :
⋃ (i : ι), s i = Set.univ
theorem IndexedPartition.disjoint {ι : Type u_1} {α : Type u_2} {s : ιSet α} (hs : IndexedPartition s) :
theorem IndexedPartition.mem_iff_index_eq {ι : Type u_1} {α : Type u_2} {s : ιSet α} (hs : IndexedPartition s) {x : α} {i : ι} :
x s i hs.index x = i
theorem IndexedPartition.eq {ι : Type u_1} {α : Type u_2} {s : ιSet α} (hs : IndexedPartition s) (i : ι) :
s i = {x : α | hs.index x = i}
@[reducible, inline]
abbrev IndexedPartition.setoid {ι : Type u_1} {α : Type u_2} {s : ιSet α} (hs : IndexedPartition s) :

The equivalence relation associated to an indexed partition. Two elements are equivalent if they belong to the same set of the partition.

Equations
@[simp]
theorem IndexedPartition.index_some {ι : Type u_1} {α : Type u_2} {s : ιSet α} (hs : IndexedPartition s) (i : ι) :
hs.index (hs.some i) = i
theorem IndexedPartition.some_index {ι : Type u_1} {α : Type u_2} {s : ιSet α} (hs : IndexedPartition s) (x : α) :
hs.setoid (hs.some (hs.index x)) x
def IndexedPartition.Quotient {ι : Type u_1} {α : Type u_2} {s : ιSet α} (hs : IndexedPartition s) :
Type u_2

The quotient associated to an indexed partition.

Equations
def IndexedPartition.proj {ι : Type u_1} {α : Type u_2} {s : ιSet α} (hs : IndexedPartition s) :
αhs.Quotient

The projection onto the quotient associated to an indexed partition.

Equations
instance IndexedPartition.instInhabitedQuotient {ι : Type u_1} {α : Type u_2} {s : ιSet α} (hs : IndexedPartition s) [Inhabited α] :
Equations
theorem IndexedPartition.proj_eq_iff {ι : Type u_1} {α : Type u_2} {s : ιSet α} (hs : IndexedPartition s) {x y : α} :
hs.proj x = hs.proj y hs.index x = hs.index y
@[simp]
theorem IndexedPartition.proj_some_index {ι : Type u_1} {α : Type u_2} {s : ιSet α} (hs : IndexedPartition s) (x : α) :
hs.proj (hs.some (hs.index x)) = hs.proj x
def IndexedPartition.equivQuotient {ι : Type u_1} {α : Type u_2} {s : ιSet α} (hs : IndexedPartition s) :

The obvious equivalence between the quotient associated to an indexed partition and the indexing type.

Equations
@[simp]
theorem IndexedPartition.equivQuotient_index_apply {ι : Type u_1} {α : Type u_2} {s : ιSet α} (hs : IndexedPartition s) (x : α) :
hs.equivQuotient (hs.index x) = hs.proj x
@[simp]
theorem IndexedPartition.equivQuotient_symm_proj_apply {ι : Type u_1} {α : Type u_2} {s : ιSet α} (hs : IndexedPartition s) (x : α) :
hs.equivQuotient.symm (hs.proj x) = hs.index x
theorem IndexedPartition.equivQuotient_index {ι : Type u_1} {α : Type u_2} {s : ιSet α} (hs : IndexedPartition s) :
def IndexedPartition.out {ι : Type u_1} {α : Type u_2} {s : ιSet α} (hs : IndexedPartition s) :

A map choosing a representative for each element of the quotient associated to an indexed partition. This is a computable version of Quotient.out using IndexedPartition.some.

Equations
@[simp]
theorem IndexedPartition.out_proj {ι : Type u_1} {α : Type u_2} {s : ιSet α} (hs : IndexedPartition s) (x : α) :
hs.out (hs.proj x) = hs.some (hs.index x)

This lemma is analogous to Quotient.mk_out'.

theorem IndexedPartition.index_out {ι : Type u_1} {α : Type u_2} {s : ιSet α} (hs : IndexedPartition s) (x : hs.Quotient) :
hs.index (Quotient.out x) = hs.index (hs.out x)

The indices of Quotient.out and IndexedPartition.out are equal.

@[simp]
theorem IndexedPartition.proj_out {ι : Type u_1} {α : Type u_2} {s : ιSet α} (hs : IndexedPartition s) (x : hs.Quotient) :
hs.proj (hs.out x) = x

This lemma is analogous to Quotient.out_eq'.

theorem IndexedPartition.class_of {ι : Type u_1} {α : Type u_2} {s : ιSet α} (hs : IndexedPartition s) {x : α} :
setOf (hs.setoid x) = s (hs.index x)
theorem IndexedPartition.proj_fiber {ι : Type u_1} {α : Type u_2} {s : ιSet α} (hs : IndexedPartition s) (x : hs.Quotient) :
def IndexedPartition.piecewise {ι : Type u_1} {α : Type u_2} {s : ιSet α} (hs : IndexedPartition s) {β : Type u_3} (f : ιαβ) :
αβ

Combine functions with disjoint domains into a new function. You can use the regular expression def.*piecewise to search for other ways to define piecewise functions in mathlib4.

Equations
theorem IndexedPartition.piecewise_apply {ι : Type u_1} {α : Type u_2} {s : ιSet α} (hs : IndexedPartition s) {β : Type u_3} {f : ιαβ} (x : α) :
hs.piecewise f x = f (hs.index x) x
theorem IndexedPartition.piecewise_inj {ι : Type u_1} {α : Type u_2} {s : ιSet α} (hs : IndexedPartition s) {β : Type u_3} {f : ιαβ} (h_injOn : ∀ (i : ι), Set.InjOn (f i) (s i)) (h_disjoint : Set.univ.PairwiseDisjoint fun (i : ι) => f i '' s i) :

A family of injective functions with pairwise disjoint domains and pairwise disjoint ranges can be glued together to form an injective function.

theorem IndexedPartition.piecewise_bij {ι : Type u_1} {α : Type u_2} {s : ιSet α} (hs : IndexedPartition s) {β : Type u_3} {f : ιαβ} {t : ιSet β} (ht : IndexedPartition t) (hf : ∀ (i : ι), Set.BijOn (f i) (s i) (t i)) :

A family of bijective functions with pairwise disjoint domains and pairwise disjoint ranges can be glued together to form a bijective function.