Documentation

Mathlib.Combinatorics.SimpleGraph.Regularity.Increment

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 the partition obtained by gluing the parts partitions together (the increment partition) and shows that the energy globally increases.

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]

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

The increment partition in Szemerédi's Regularity Lemma.

If an equipartition is not uniform, then the increment partition is a (much bigger) equipartition with a slightly higher energy. This is helpful since the energy is bounded by a constant (see Finpartition.energy_le_one), so this process eventually terminates and yields a not-too-big uniform equipartition.

Equations
Instances For
    theorem SzemerediRegularity.card_increment {α : Type u_1} [Fintype α] [DecidableEq α] {P : Finpartition Finset.univ} {hP : P.IsEquipartition} {G : SimpleGraph α} [DecidableRel G.Adj] {ε : } (hPα : P.parts.card * 16 ^ P.parts.card Fintype.card α) (hPG : ¬P.IsUniform G ε) :

    The increment partition has a prescribed (very big) size in terms of the original partition.

    theorem SzemerediRegularity.increment_isEquipartition {α : Type u_1} [Fintype α] [DecidableEq α] {P : Finpartition Finset.univ} (hP : P.IsEquipartition) (G : SimpleGraph α) [DecidableRel G.Adj] (ε : ) :
    (SzemerediRegularity.increment hP G ε).IsEquipartition
    theorem SzemerediRegularity.le_sum_distinctPairs_edgeDensity_sq {α : Type u_1} [Fintype α] [DecidableEq α] {P : Finpartition Finset.univ} {hP : P.IsEquipartition} {G : SimpleGraph α} [DecidableRel G.Adj] {ε : } [Nonempty α] (x : { i : Finset α × Finset α // i P.parts.offDiag }) (hε₁ : ε 1) (hPα : P.parts.card * 16 ^ P.parts.card Fintype.card α) (hPε : 100 4 ^ P.parts.card * ε ^ 5) :
    (G.edgeDensity (x).1 (x).2) ^ 2 + ((if G.IsUniform ε (x).1 (x).2 then 0 else ε ^ 4 / 3) - ε ^ 5 / 25) ((SzemerediRegularity.distinctPairs hP G ε x).sum fun (i : Finset α × Finset α) => (G.edgeDensity i.1 i.2) ^ 2) / 16 ^ P.parts.card
    theorem SzemerediRegularity.energy_increment {α : Type u_1} [Fintype α] [DecidableEq α] {P : Finpartition Finset.univ} {G : SimpleGraph α} [DecidableRel G.Adj] {ε : } [Nonempty α] (hP : P.IsEquipartition) (hP₇ : 7 P.parts.card) (hPε : 100 4 ^ P.parts.card * ε ^ 5) (hPα : P.parts.card * 16 ^ P.parts.card Fintype.card α) (hPG : ¬P.IsUniform G ε) (hε₀ : 0 ε) (hε₁ : ε 1) :
    (P.energy G) + ε ^ 5 / 4 ((SzemerediRegularity.increment hP G ε).energy G)

    The increment partition has energy greater than the original one by a known fixed amount.