Documentation

Mathlib.Combinatorics.SimpleGraph.Regularity.Chunk

Chunk of the increment partition for Szemerédi Regularity Lemma #

In the proof of Szemerédi Regularity Lemma, we need to partition each part of a starting partition to increase the energy. This file defines those partitions of parts and shows that they locally increase the energy.

This entire file is internal to the proof of Szemerédi Regularity Lemma.

Main declarations #

TODO #

Once ported to mathlib4, this file will be a great golfing ground for Heather's new tactic gcongr.

References #

[Yaël Dillies, Bhavik Mehta, Formalising Szemerédi’s Regularity Lemma in Lean][srl_itp]

Definitions #

We define chunk, the partition of a part, and star, the sets of parts of chunk that are contained in the corresponding witness of non-uniformity.

noncomputable def SzemerediRegularity.chunk {α : Type u_1} [Fintype α] [DecidableEq α] {P : Finpartition Finset.univ} (hP : P.IsEquipartition) (G : SimpleGraph α) [DecidableRel G.Adj] (ε : ) {U : Finset α} (hU : U P.parts) :

The portion of SzemerediRegularity.increment which partitions U.

Equations
  • One or more equations did not get rendered due to their size.
Instances For
    noncomputable def SzemerediRegularity.star {α : Type u_1} [Fintype α] [DecidableEq α] {P : Finpartition Finset.univ} (hP : P.IsEquipartition) (G : SimpleGraph α) [DecidableRel G.Adj] (ε : ) {U : Finset α} (hU : U P.parts) (V : Finset α) :

    The portion of SzemerediRegularity.chunk which is contained in the witness of non-uniformity of U and V.

    Equations
    Instances For

      Density estimates #

      We estimate the density between parts of chunk.

      theorem SzemerediRegularity.biUnion_star_subset_nonuniformWitness {α : Type u_1} [Fintype α] [DecidableEq α] {P : Finpartition Finset.univ} (hP : P.IsEquipartition) (G : SimpleGraph α) [DecidableRel G.Adj] (ε : ) {U : Finset α} (hU : U P.parts) (V : Finset α) :
      (SzemerediRegularity.star hP G ε hU V).biUnion id G.nonuniformWitness ε U V
      theorem SzemerediRegularity.star_subset_chunk {α : Type u_1} [Fintype α] [DecidableEq α] {P : Finpartition Finset.univ} {hP : P.IsEquipartition} {G : SimpleGraph α} [DecidableRel G.Adj] {ε : } {U : Finset α} {hU : U P.parts} {V : Finset α} :

      chunk #

      theorem SzemerediRegularity.card_chunk {α : Type u_1} [Fintype α] [DecidableEq α] {P : Finpartition Finset.univ} {hP : P.IsEquipartition} {G : SimpleGraph α} [DecidableRel G.Adj] {ε : } {U : Finset α} {hU : U P.parts} (hm : Fintype.card α / SzemerediRegularity.stepBound P.parts.card 0) :
      (SzemerediRegularity.chunk hP G ε hU).parts.card = 4 ^ P.parts.card
      theorem SzemerediRegularity.card_eq_of_mem_parts_chunk {α : Type u_1} [Fintype α] [DecidableEq α] {P : Finpartition Finset.univ} {hP : P.IsEquipartition} {G : SimpleGraph α} [DecidableRel G.Adj] {ε : } {U : Finset α} {hU : U P.parts} {s : Finset α} (hs : s (SzemerediRegularity.chunk hP G ε hU).parts) :
      theorem SzemerediRegularity.m_le_card_of_mem_chunk_parts {α : Type u_1} [Fintype α] [DecidableEq α] {P : Finpartition Finset.univ} {hP : P.IsEquipartition} {G : SimpleGraph α} [DecidableRel G.Adj] {ε : } {U : Finset α} {hU : U P.parts} {s : Finset α} (hs : s (SzemerediRegularity.chunk hP G ε hU).parts) :
      theorem SzemerediRegularity.card_le_m_add_one_of_mem_chunk_parts {α : Type u_1} [Fintype α] [DecidableEq α] {P : Finpartition Finset.univ} {hP : P.IsEquipartition} {G : SimpleGraph α} [DecidableRel G.Adj] {ε : } {U : Finset α} {hU : U P.parts} {s : Finset α} (hs : s (SzemerediRegularity.chunk hP G ε hU).parts) :
      theorem SzemerediRegularity.card_biUnion_star_le_m_add_one_card_star_mul {α : Type u_1} [Fintype α] [DecidableEq α] {P : Finpartition Finset.univ} {hP : P.IsEquipartition} {G : SimpleGraph α} [DecidableRel G.Adj] {ε : } {U : Finset α} {hU : U P.parts} {V : Finset α} :
      ((SzemerediRegularity.star hP G ε hU V).biUnion id).card (SzemerediRegularity.star hP G ε hU V).card * ((Fintype.card α / SzemerediRegularity.stepBound P.parts.card) + 1)

      Final bounds #

      Those inequalities are the end result of all this hard work.

      theorem SzemerediRegularity.edgeDensity_chunk_not_uniform {α : Type u_1} [Fintype α] [DecidableEq α] {P : Finpartition Finset.univ} {hP : P.IsEquipartition} {G : SimpleGraph α} [DecidableRel G.Adj] {ε : } {U : Finset α} {V : Finset α} [Nonempty α] (hPα : P.parts.card * 16 ^ P.parts.card Fintype.card α) (hPε : 100 4 ^ P.parts.card * ε ^ 5) (hε₁ : ε 1) {hU : U P.parts} {hV : V P.parts} (hUVne : U V) (hUV : ¬G.IsUniform ε U V) :
      (G.edgeDensity U V) ^ 2 - ε ^ 5 / 25 + ε ^ 4 / 3 (((SzemerediRegularity.chunk hP G ε hU).parts.product (SzemerediRegularity.chunk hP G ε hV).parts).sum fun (ab : Finset α × Finset α) => (G.edgeDensity ab.1 ab.2) ^ 2) / 16 ^ P.parts.card

      Lower bound on the edge densities between non-uniform parts of SzemerediRegularity.increment.

      theorem SzemerediRegularity.edgeDensity_chunk_uniform {α : Type u_1} [Fintype α] [DecidableEq α] {P : Finpartition Finset.univ} {hP : P.IsEquipartition} {G : SimpleGraph α} [DecidableRel G.Adj] {ε : } {U : Finset α} {V : Finset α} [Nonempty α] (hPα : P.parts.card * 16 ^ P.parts.card Fintype.card α) (hPε : 100 4 ^ P.parts.card * ε ^ 5) (hU : U P.parts) (hV : V P.parts) :
      (G.edgeDensity U V) ^ 2 - ε ^ 5 / 25 (((SzemerediRegularity.chunk hP G ε hU).parts.product (SzemerediRegularity.chunk hP G ε hV).parts).sum fun (ab : Finset α × Finset α) => (G.edgeDensity ab.1 ab.2) ^ 2) / 16 ^ P.parts.card

      Lower bound on the edge densities between parts of SzemerediRegularity.increment. This is the blanket lower bound used the uniform parts.