Documentation

Mathlib.Order.Partition.Finpartition

Finite partitions #

In this file, we define finite partitions. A finpartition of a : α is a finite set of pairwise disjoint parts parts : Finset α which does not contain and whose supremum is a.

Finpartitions of a finset are at the heart of Szemerédi's regularity lemma. They are also studied purely order theoretically in Sperner theory.

Constructions #

We provide many ways to build finpartitions:

Finpartition.indiscrete and Finpartition.bind together form the monadic structure of Finpartition.

Implementation notes #

Forbidding as a part follows mathematical tradition and is a pragmatic choice concerning operations on Finpartition. Not caring about being a part or not breaks extensionality (it's not because the parts of P and the parts of Q have the same elements that P = Q). Enforcing to be a part makes Finpartition.bind uglier and doesn't rid us of the need of Finpartition.ofErase.

TODO #

The order is the wrong way around to make Finpartition a a graded order. Is it bad to depart from the literature and turn the order around?

The specialisation to Finset α could be generalised to atomistic orders.

structure Finpartition {α : Type u_1} [Lattice α] [OrderBot α] (a : α) :
Type u_1

A finite partition of a : α is a pairwise disjoint finite set of elements whose supremum is a. We forbid as a part.

  • parts : Finset α

    The elements of the finite partition of a

  • supIndep : self.parts.SupIndep id

    The partition is supremum-independent

  • sup_parts : self.parts.sup id = a

    The supremum of the partition is a

  • bot_notMem : self.parts

    No element of the partition is bottom

Instances For
    theorem Finpartition.ext_iff {α : Type u_1} {inst✝ : Lattice α} {inst✝¹ : OrderBot α} {a : α} {x y : Finpartition a} :
    x = y x.parts = y.parts
    theorem Finpartition.ext {α : Type u_1} {inst✝ : Lattice α} {inst✝¹ : OrderBot α} {a : α} {x y : Finpartition a} (parts : x.parts = y.parts) :
    x = y
    instance instDecidableEqFinpartition {α✝ : Type u_2} {inst✝ : Lattice α✝} {inst✝¹ : OrderBot α✝} {a✝ : α✝} [DecidableEq α✝] :
    Equations
    def instDecidableEqFinpartition.decEq {α✝ : Type u_2} {inst✝ : Lattice α✝} {inst✝¹ : OrderBot α✝} {a✝ : α✝} [DecidableEq α✝] (x✝ x✝¹ : Finpartition a✝) :
    Decidable (x✝ = x✝¹)
    Equations
    • One or more equations did not get rendered due to their size.
    Instances For
      @[deprecated Finpartition.bot_notMem (since := "2025-05-23")]
      theorem Finpartition.not_bot_mem {α : Type u_1} [Lattice α] [OrderBot α] {a : α} (self : Finpartition a) :
      self.parts

      Alias of Finpartition.bot_notMem.


      No element of the partition is bottom

      def Finpartition.ofErase {α : Type u_1} [Lattice α] [OrderBot α] [DecidableEq α] {a : α} (parts : Finset α) (sup_indep : parts.SupIndep id) (sup_parts : parts.sup id = a) :

      A Finpartition constructor which does not insist on not being a part.

      Equations
      Instances For
        @[simp]
        theorem Finpartition.ofErase_parts {α : Type u_1} [Lattice α] [OrderBot α] [DecidableEq α] {a : α} (parts : Finset α) (sup_indep : parts.SupIndep id) (sup_parts : parts.sup id = a) :
        (ofErase parts sup_indep sup_parts).parts = parts.erase
        def Finpartition.ofSubset {α : Type u_1} [Lattice α] [OrderBot α] {a b : α} (P : Finpartition a) {parts : Finset α} (subset : parts P.parts) (sup_parts : parts.sup id = b) :

        A Finpartition constructor from a bigger existing finpartition.

        Equations
        • P.ofSubset subset sup_parts = { parts := parts, supIndep := , sup_parts := sup_parts, bot_notMem := }
        Instances For
          @[simp]
          theorem Finpartition.ofSubset_parts {α : Type u_1} [Lattice α] [OrderBot α] {a b : α} (P : Finpartition a) {parts : Finset α} (subset : parts P.parts) (sup_parts : parts.sup id = b) :
          (P.ofSubset subset sup_parts).parts = parts
          def Finpartition.copy {α : Type u_1} [Lattice α] [OrderBot α] {a b : α} (P : Finpartition a) (h : a = b) :

          Changes the type of a finpartition to an equal one.

          Equations
          • P.copy h = { parts := P.parts, supIndep := , sup_parts := , bot_notMem := }
          Instances For
            @[simp]
            theorem Finpartition.copy_parts {α : Type u_1} [Lattice α] [OrderBot α] {a b : α} (P : Finpartition a) (h : a = b) :
            (P.copy h).parts = P.parts
            def Finpartition.map {α : Type u_1} [Lattice α] [OrderBot α] {β : Type u_2} [Lattice β] [OrderBot β] {a : α} (e : α ≃o β) (P : Finpartition a) :

            Transfer a finpartition over an order isomorphism.

            Equations
            Instances For
              @[simp]
              theorem Finpartition.parts_map {α : Type u_1} [Lattice α] [OrderBot α] {β : Type u_2} [Lattice β] [OrderBot β] {a : α} {e : α ≃o β} {P : Finpartition a} :

              The empty finpartition.

              Equations
              Instances For
                def Finpartition.indiscrete {α : Type u_1} [Lattice α] [OrderBot α] {a : α} (ha : a ) :

                The finpartition in one part, aka indiscrete finpartition.

                Equations
                Instances For
                  @[simp]
                  theorem Finpartition.indiscrete_parts {α : Type u_1} [Lattice α] [OrderBot α] {a : α} (ha : a ) :
                  theorem Finpartition.le {α : Type u_1} [Lattice α] [OrderBot α] {a : α} (P : Finpartition a) {b : α} (hb : b P.parts) :
                  b a
                  theorem Finpartition.ne_bot {α : Type u_1} [Lattice α] [OrderBot α] {a : α} (P : Finpartition a) {b : α} (hb : b P.parts) :
                  theorem Finpartition.disjoint {α : Type u_1} [Lattice α] [OrderBot α] {a : α} (P : Finpartition a) :
                  @[simp]
                  theorem Finpartition.parts_eq_empty_iff {α : Type u_1} [Lattice α] [OrderBot α] {a : α} {P : Finpartition a} :
                  @[simp]
                  theorem Finpartition.parts_nonempty_iff {α : Type u_1} [Lattice α] [OrderBot α] {a : α} {P : Finpartition a} :
                  theorem Finpartition.parts_nonempty {α : Type u_1} [Lattice α] [OrderBot α] {a : α} (P : Finpartition a) (ha : a ) :
                  Equations
                  @[reducible, inline]
                  abbrev IsAtom.uniqueFinpartition {α : Type u_1} [Lattice α] [OrderBot α] {a : α} {P : Finpartition a} (ha : IsAtom a) :

                  There's a unique partition of an atom.

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

                    Refinement order #

                    instance Finpartition.instLE {α : Type u_1} [Lattice α] [OrderBot α] {a : α} :

                    We say that P ≤ Q if P refines Q: each part of P is less than some part of Q.

                    Equations
                    instance Finpartition.instPartialOrder {α : Type u_1} [Lattice α] [OrderBot α] {a : α} :
                    Equations
                    Equations
                    theorem Finpartition.parts_top_subset {α : Type u_1} [Lattice α] [OrderBot α] (a : α) [Decidable (a = )] :
                    Equations
                    • One or more equations did not get rendered due to their size.
                    instance Finpartition.instMin {α : Type u_1} [DistribLattice α] [OrderBot α] [DecidableEq α] {a : α} :
                    Equations
                    @[simp]
                    theorem Finpartition.parts_inf {α : Type u_1} [DistribLattice α] [OrderBot α] [DecidableEq α] {a : α} (P Q : Finpartition a) :
                    (PQ).parts = (Finset.image (fun (bc : α × α) => bc.1bc.2) (P.parts ×ˢ Q.parts)).erase
                    Equations
                    theorem Finpartition.exists_le_of_le {α : Type u_1} [DistribLattice α] [OrderBot α] {a b : α} {P Q : Finpartition a} (h : P Q) (hb : b Q.parts) :
                    cP.parts, c b
                    theorem Finpartition.card_mono {α : Type u_1} [DistribLattice α] [OrderBot α] {a : α} {P Q : Finpartition a} (h : P Q) :
                    def Finpartition.bind {α : Type u_1} [DistribLattice α] [OrderBot α] [DecidableEq α] {a : α} (P : Finpartition a) (Q : (i : α) → i P.partsFinpartition i) :

                    Given a finpartition P of a and finpartitions of each part of P, this yields the finpartition of a obtained by juxtaposing all the subpartitions.

                    Equations
                    Instances For
                      @[simp]
                      theorem Finpartition.bind_parts {α : Type u_1} [DistribLattice α] [OrderBot α] [DecidableEq α] {a : α} (P : Finpartition a) (Q : (i : α) → i P.partsFinpartition i) :
                      (P.bind Q).parts = P.parts.attach.biUnion fun (i : { x : α // x P.parts }) => (Q i ).parts
                      theorem Finpartition.mem_bind {α : Type u_1} [DistribLattice α] [OrderBot α] [DecidableEq α] {a b : α} {P : Finpartition a} {Q : (i : α) → i P.partsFinpartition i} :
                      b (P.bind Q).parts ∃ (A : α) (hA : A P.parts), b (Q A hA).parts
                      theorem Finpartition.card_bind {α : Type u_1} [DistribLattice α] [OrderBot α] [DecidableEq α] {a : α} {P : Finpartition a} (Q : (i : α) → i P.partsFinpartition i) :
                      (P.bind Q).parts.card = AP.parts.attach, (Q A ).parts.card
                      def Finpartition.extend {α : Type u_1} [DistribLattice α] [OrderBot α] [DecidableEq α] {a b c : α} (P : Finpartition a) (hb : b ) (hab : Disjoint a b) (hc : ab = c) :

                      Adds b to a finpartition of a to make a finpartition of a ⊔ b.

                      Equations
                      • P.extend hb hab hc = { parts := insert b P.parts, supIndep := , sup_parts := , bot_notMem := }
                      Instances For
                        @[simp]
                        theorem Finpartition.extend_parts {α : Type u_1} [DistribLattice α] [OrderBot α] [DecidableEq α] {a b c : α} (P : Finpartition a) (hb : b ) (hab : Disjoint a b) (hc : ab = c) :
                        (P.extend hb hab hc).parts = insert b P.parts
                        theorem Finpartition.card_extend {α : Type u_1} [DistribLattice α] [OrderBot α] [DecidableEq α] {a : α} (P : Finpartition a) (b c : α) {hb : b } {hab : Disjoint a b} {hc : ab = c} :
                        (P.extend hb hab hc).parts.card = P.parts.card + 1
                        def Finpartition.avoid {α : Type u_1} [GeneralizedBooleanAlgebra α] [DecidableEq α] {a : α} (P : Finpartition a) (b : α) :

                        Restricts a finpartition to avoid a given element.

                        Equations
                        Instances For
                          @[simp]
                          theorem Finpartition.avoid_parts_val {α : Type u_1} [GeneralizedBooleanAlgebra α] [DecidableEq α] {a : α} (P : Finpartition a) (b : α) :
                          (P.avoid b).parts.val = (Multiset.map (fun (x : α) => x \ b) P.parts.val).dedup.erase
                          @[simp]
                          theorem Finpartition.mem_avoid {α : Type u_1} [GeneralizedBooleanAlgebra α] [DecidableEq α] {a b c : α} (P : Finpartition a) :
                          c (P.avoid b).parts dP.parts, ¬d b d \ b = c

                          Finite partitions of finsets #

                          theorem Finpartition.subset {α : Type u_1} [DecidableEq α] {s : Finset α} (P : Finpartition s) {a : Finset α} (ha : a P.parts) :
                          a s
                          theorem Finpartition.nonempty_of_mem_parts {α : Type u_1} [DecidableEq α] {s : Finset α} (P : Finpartition s) {a : Finset α} (ha : a P.parts) :
                          @[simp]
                          theorem Finpartition.empty_notMem_parts {α : Type u_1} [DecidableEq α] {s : Finset α} (P : Finpartition s) :
                          P.parts
                          @[deprecated Finpartition.empty_notMem_parts (since := "2025-05-23")]
                          theorem Finpartition.not_empty_mem_parts {α : Type u_1} [DecidableEq α] {s : Finset α} (P : Finpartition s) :
                          P.parts

                          Alias of Finpartition.empty_notMem_parts.

                          theorem Finpartition.ne_empty {α : Type u_1} [DecidableEq α] {s t : Finset α} (P : Finpartition s) (h : t P.parts) :
                          theorem Finpartition.eq_of_mem_parts {α : Type u_1} [DecidableEq α] {s t u : Finset α} (P : Finpartition s) {a : α} (ht : t P.parts) (hu : u P.parts) (hat : a t) (hau : a u) :
                          t = u
                          theorem Finpartition.exists_mem {α : Type u_1} [DecidableEq α] {s : Finset α} (P : Finpartition s) {a : α} (ha : a s) :
                          tP.parts, a t
                          theorem Finpartition.biUnion_parts {α : Type u_1} [DecidableEq α] {s : Finset α} (P : Finpartition s) :
                          theorem Finpartition.existsUnique_mem {α : Type u_1} [DecidableEq α] {s : Finset α} (P : Finpartition s) {a : α} (ha : a s) :
                          ∃! t : Finset α, t P.parts a t
                          def Finpartition.ofExistsUnique {α : Type u_1} [DecidableEq α] {s : Finset α} (parts : Finset (Finset α)) (h : pparts, p s) (h' : as, ∃! t : Finset α, t parts a t) (h'' : parts) :

                          Construct a Finpartition s from a finset of finsets parts such that each element of s is in exactly one member of parts. This provides a converse to Finpartition.subset, Finpartition.not_empty_mem_parts and Finpartition.existsUnique_mem.

                          Equations
                          Instances For
                            @[simp]
                            theorem Finpartition.ofExistsUnique_parts {α : Type u_1} [DecidableEq α] {s : Finset α} (parts : Finset (Finset α)) (h : pparts, p s) (h' : as, ∃! t : Finset α, t parts a t) (h'' : parts) :
                            (ofExistsUnique parts h h' h'').parts = parts
                            def Finpartition.part {α : Type u_1} [DecidableEq α] {s : Finset α} (P : Finpartition s) (a : α) :

                            The part of the finpartition that a lies in.

                            Equations
                            Instances For
                              @[simp]
                              theorem Finpartition.part_mem {α : Type u_1} [DecidableEq α] {s : Finset α} (P : Finpartition s) {a : α} :
                              P.part a P.parts a s
                              @[simp]
                              theorem Finpartition.part_eq_empty {α : Type u_1} [DecidableEq α] {s : Finset α} (P : Finpartition s) {a : α} :
                              P.part a = as
                              @[simp]
                              theorem Finpartition.part_nonempty {α : Type u_1} [DecidableEq α] {s : Finset α} (P : Finpartition s) {a : α} :
                              (P.part a).Nonempty a s
                              @[simp]
                              theorem Finpartition.part_subset {α : Type u_1} [DecidableEq α] {s : Finset α} (P : Finpartition s) (a : α) :
                              P.part a s
                              @[simp]
                              theorem Finpartition.mem_part_self {α : Type u_1} [DecidableEq α] {s : Finset α} (P : Finpartition s) {a : α} :
                              a P.part a a s
                              theorem Finpartition.mem_part {α : Type u_1} [DecidableEq α] {s : Finset α} (P : Finpartition s) {a : α} :
                              a sa P.part a

                              Alias of the reverse direction of Finpartition.mem_part_self.

                              theorem Finpartition.part_eq_iff_mem {α : Type u_1} [DecidableEq α] {s t : Finset α} (P : Finpartition s) {a : α} (ht : t P.parts) :
                              P.part a = t a t
                              theorem Finpartition.part_eq_of_mem {α : Type u_1} [DecidableEq α] {s t : Finset α} (P : Finpartition s) {a : α} (ht : t P.parts) (hat : a t) :
                              P.part a = t
                              theorem Finpartition.mem_part_iff_part_eq_part {α : Type u_1} [DecidableEq α] {s : Finset α} (P : Finpartition s) {a b : α} (ha : a s) (hb : b s) :
                              a P.part b P.part a = P.part b
                              theorem Finpartition.part_surjOn {α : Type u_1} [DecidableEq α] {s : Finset α} (P : Finpartition s) :
                              Set.SurjOn P.part s P.parts
                              theorem Finpartition.exists_subset_part_bijOn {α : Type u_1} [DecidableEq α] {s : Finset α} (P : Finpartition s) :
                              rs, Set.BijOn P.part r P.parts
                              theorem Finpartition.mem_part_iff_exists {α : Type u_1} [DecidableEq α] {s : Finset α} (P : Finpartition s) {a b : α} :
                              a P.part b pP.parts, a p b p
                              def Finpartition.equivSigmaParts {α : Type u_1} [DecidableEq α] {s : Finset α} (P : Finpartition s) :
                              { x : α // x s } (t : { x : Finset α // x P.parts }) × { x : α // x t }

                              Equivalence between a finpartition's parts as a dependent sum and the partitioned set.

                              Equations
                              • One or more equations did not get rendered due to their size.
                              Instances For
                                theorem Finpartition.exists_enumeration {α : Type u_1} [DecidableEq α] {s : Finset α} (P : Finpartition s) :
                                ∃ (f : { x : α // x s } (t : { x : Finset α // x P.parts }) × Fin (↑t).card), ∀ (a b : { x : α // x s }), P.part a = P.part b (f a).fst = (f b).fst
                                theorem Finpartition.sum_card_parts {α : Type u_1} [DecidableEq α] {s : Finset α} (P : Finpartition s) :
                                iP.parts, i.card = s.card
                                instance Finpartition.instBotFinset {α : Type u_1} [DecidableEq α] (s : Finset α) :

                                is the partition in singletons, aka discrete partition.

                                Equations
                                @[simp]
                                theorem Finpartition.parts_bot {α : Type u_1} [DecidableEq α] (s : Finset α) :
                                .parts = Finset.map { toFun := singleton, inj' := } s
                                theorem Finpartition.card_bot {α : Type u_1} [DecidableEq α] (s : Finset α) :
                                theorem Finpartition.mem_bot_iff {α : Type u_1} [DecidableEq α] {s t : Finset α} :
                                t .parts as, {a} = t
                                def Finpartition.ofSetSetoid {α : Type u_1} [DecidableEq α] (s : Setoid α) (x : Finset α) [DecidableRel s] :

                                A setoid over a finite type induces a finpartition of the type's elements, where the parts are the setoid's equivalence classes.

                                Equations
                                Instances For
                                  theorem Finpartition.ofSetSetoid_parts {α : Type u_1} [DecidableEq α] (s : Setoid α) (x : Finset α) [DecidableRel s] :
                                  (ofSetSetoid s x).parts = Finset.image (fun (a : α) => {bx | s a b}) x
                                  theorem Finpartition.mem_part_ofSetSetoid_iff_rel {α : Type u_1} [DecidableEq α] {a : α} {s : Setoid α} (x : Finset α) [DecidableRel s] {b : α} :
                                  b (ofSetSetoid s x).part a a x b x s a b

                                  A setoid over a finite type induces a finpartition of the type's elements, where the parts are the setoid's equivalence classes.

                                  Equations
                                  Instances For
                                    theorem Finpartition.ofSetoid_parts_val {α : Type u_1} [DecidableEq α] [Fintype α] (s : Setoid α) [DecidableRel s] :
                                    (ofSetoid s).parts.val = (Multiset.map (fun (a : α) => {b : α | s a b}) Finset.univ.val).dedup
                                    theorem Finpartition.mem_part_ofSetoid_iff_rel {α : Type u_1} [DecidableEq α] {a : α} [Fintype α] {s : Setoid α} [DecidableRel s] {b : α} :
                                    b (ofSetoid s).part a s a b
                                    def Finpartition.atomise {α : Type u_1} [DecidableEq α] (s : Finset α) (F : Finset (Finset α)) :

                                    Cuts s along the finsets in F: Two elements of s will be in the same part if they are in the same finsets of F.

                                    Equations
                                    Instances For
                                      theorem Finpartition.mem_atomise {α : Type u_1} [DecidableEq α] {s t : Finset α} {F : Finset (Finset α)} :
                                      t (atomise s F).parts t.Nonempty QF, {is | uF, u Q i u} = t
                                      theorem Finpartition.atomise_empty {α : Type u_1} [DecidableEq α] {s : Finset α} (hs : s.Nonempty) :
                                      theorem Finpartition.card_atomise_le {α : Type u_1} [DecidableEq α] {s : Finset α} {F : Finset (Finset α)} :
                                      theorem Finpartition.biUnion_filter_atomise {α : Type u_1} [DecidableEq α] {s t : Finset α} {F : Finset (Finset α)} (ht : t F) (hts : t s) :
                                      {u(atomise s F).parts | u t u.Nonempty}.biUnion id = t
                                      theorem Finpartition.card_filter_atomise_le_two_pow {α : Type u_1} [DecidableEq α] {s t : Finset α} {F : Finset (Finset α)} (ht : t F) :
                                      {u(atomise s F).parts | u t u.Nonempty}.card 2 ^ (F.card - 1)