Documentation

Mathlib.Combinatorics.Additive.AP.Three.Defs

Sets without arithmetic progressions of length three and Roth numbers #

This file defines sets without arithmetic progressions of length three, aka 3AP-free sets, and the Roth number of a set.

The corresponding notion, sets without geometric progressions of length three, are called 3GP-free sets.

The Roth number of a finset is the size of its biggest 3AP-free subset. This is a more general definition than the one often found in mathematical literature, where the n-th Roth number is the size of the biggest 3AP-free subset of {0, ..., n - 1}.

Main declarations #

TODO #

References #

Tags #

3AP-free, Salem-Spencer, Roth, arithmetic progression, average, three-free

def ThreeGPFree {α : Type u_2} [Monoid α] (s : Set α) :

A set is 3GP-free if it does not contain any non-trivial geometric progression of length three.

Equations
    Instances For
      def ThreeAPFree {α : Type u_2} [AddMonoid α] (s : Set α) :

      A set is 3AP-free if it does not contain any non-trivial arithmetic progression of length three.

      This is also sometimes called a non averaging set or Salem-Spencer set.

      Equations
        Instances For
          instance ThreeGPFree.instDecidable {α : Type u_2} [Monoid α] [DecidableEq α] {s : Finset α} :

          Whether a given finset is 3GP-free is decidable.

          Equations
            instance ThreeAPFree.instDecidable {α : Type u_2} [AddMonoid α] [DecidableEq α] {s : Finset α} :

            Whether a given finset is 3AP-free is decidable.

            Equations
              theorem ThreeGPFree.mono {α : Type u_2} [Monoid α] {s t : Set α} (h : t s) (hs : ThreeGPFree s) :
              theorem ThreeAPFree.mono {α : Type u_2} [AddMonoid α] {s t : Set α} (h : t s) (hs : ThreeAPFree s) :
              @[simp]
              theorem threeGPFree_empty {α : Type u_2} [Monoid α] :
              @[simp]
              theorem Set.Subsingleton.threeGPFree {α : Type u_2} [Monoid α] {s : Set α} (hs : s.Subsingleton) :
              theorem Set.Subsingleton.threeAPFree {α : Type u_2} [AddMonoid α] {s : Set α} (hs : s.Subsingleton) :
              @[simp]
              theorem threeGPFree_singleton {α : Type u_2} [Monoid α] (a : α) :
              @[simp]
              theorem threeAPFree_singleton {α : Type u_2} [AddMonoid α] (a : α) :
              theorem ThreeGPFree.prod {α : Type u_2} {β : Type u_3} [Monoid α] [Monoid β] {s : Set α} {t : Set β} (hs : ThreeGPFree s) (ht : ThreeGPFree t) :
              theorem ThreeAPFree.prod {α : Type u_2} {β : Type u_3} [AddMonoid α] [AddMonoid β] {s : Set α} {t : Set β} (hs : ThreeAPFree s) (ht : ThreeAPFree t) :
              theorem threeGPFree_pi {ι : Type u_4} {α : ιType u_5} [(i : ι) → Monoid (α i)] {s : (i : ι) → Set (α i)} (hs : ∀ (i : ι), ThreeGPFree (s i)) :
              theorem threeAPFree_pi {ι : Type u_4} {α : ιType u_5} [(i : ι) → AddMonoid (α i)] {s : (i : ι) → Set (α i)} (hs : ∀ (i : ι), ThreeAPFree (s i)) :
              theorem ThreeGPFree.of_image {α : Type u_2} {β : Type u_3} [CommMonoid α] [CommMonoid β] {s A : Set α} {t : Set β} {f : αβ} (hf : IsMulFreimanHom 2 s t f) (hf' : Set.InjOn f s) (hAs : A s) (hA : ThreeGPFree (f '' A)) :

              Geometric progressions of length three are reflected under 2-Freiman homomorphisms.

              theorem ThreeAPFree.of_image {α : Type u_2} {β : Type u_3} [AddCommMonoid α] [AddCommMonoid β] {s A : Set α} {t : Set β} {f : αβ} (hf : IsAddFreimanHom 2 s t f) (hf' : Set.InjOn f s) (hAs : A s) (hA : ThreeAPFree (f '' A)) :

              Arithmetic progressions of length three are reflected under 2-Freiman homomorphisms.

              theorem threeGPFree_image {α : Type u_2} {β : Type u_3} [CommMonoid α] [CommMonoid β] {s A : Set α} {t : Set β} {f : αβ} (hf : IsMulFreimanIso 2 s t f) (hAs : A s) :

              Geometric progressions of length three are unchanged under 2-Freiman isomorphisms.

              theorem threeAPFree_image {α : Type u_2} {β : Type u_3} [AddCommMonoid α] [AddCommMonoid β] {s A : Set α} {t : Set β} {f : αβ} (hf : IsAddFreimanIso 2 s t f) (hAs : A s) :

              Arithmetic progressions of length three are unchanged under 2-Freiman isomorphisms.

              theorem ThreeGPFree.image {α : Type u_2} {β : Type u_3} [CommMonoid α] [CommMonoid β] {s A : Set α} {t : Set β} {f : αβ} (hf : IsMulFreimanIso 2 s t f) (hAs : A s) :

              Alias of the reverse direction of threeGPFree_image.


              Geometric progressions of length three are unchanged under 2-Freiman isomorphisms.

              theorem ThreeAPFree.image {α : Type u_2} {β : Type u_3} [AddCommMonoid α] [AddCommMonoid β] {s A : Set α} {t : Set β} {f : αβ} (hf : IsAddFreimanIso 2 s t f) (hAs : A s) :
              theorem IsMulFreimanHom.threeGPFree {α : Type u_2} {β : Type u_3} [CommMonoid α] [CommMonoid β] {s : Set α} {t : Set β} {f : αβ} (hf : IsMulFreimanHom 2 s t f) (hf' : Set.InjOn f s) (ht : ThreeGPFree t) :

              Geometric progressions of length three are reflected under 2-Freiman homomorphisms.

              theorem IsAddFreimanHom.threeAPFree {α : Type u_2} {β : Type u_3} [AddCommMonoid α] [AddCommMonoid β] {s : Set α} {t : Set β} {f : αβ} (hf : IsAddFreimanHom 2 s t f) (hf' : Set.InjOn f s) (ht : ThreeAPFree t) :

              Arithmetic progressions of length three are reflected under 2-Freiman homomorphisms.

              theorem IsMulFreimanIso.threeGPFree_congr {α : Type u_2} {β : Type u_3} [CommMonoid α] [CommMonoid β] {s : Set α} {t : Set β} {f : αβ} (hf : IsMulFreimanIso 2 s t f) :

              Geometric progressions of length three are unchanged under 2-Freiman isomorphisms.

              theorem IsAddFreimanIso.threeAPFree_congr {α : Type u_2} {β : Type u_3} [AddCommMonoid α] [AddCommMonoid β] {s : Set α} {t : Set β} {f : αβ} (hf : IsAddFreimanIso 2 s t f) :

              Arithmetic progressions of length three are unchanged under 2-Freiman isomorphisms.

              theorem ThreeGPFree.image' {F : Type u_1} {α : Type u_2} {β : Type u_3} [CommMonoid α] [CommMonoid β] {s : Set α} [FunLike F α β] [MulHomClass F α β] (f : F) (hf : Set.InjOn (⇑f) (s * s)) (h : ThreeGPFree s) :
              ThreeGPFree (f '' s)

              Geometric progressions of length three are preserved under semigroup homomorphisms.

              theorem ThreeAPFree.image' {F : Type u_1} {α : Type u_2} {β : Type u_3} [AddCommMonoid α] [AddCommMonoid β] {s : Set α} [FunLike F α β] [AddHomClass F α β] (f : F) (hf : Set.InjOn (⇑f) (s + s)) (h : ThreeAPFree s) :
              ThreeAPFree (f '' s)

              Arithmetic progressions of length three are preserved under semigroup homomorphisms.

              theorem ThreeGPFree.eq_right {α : Type u_2} [CommMonoid α] [IsCancelMul α] {s : Set α} (hs : ThreeGPFree s) a : α :
              a s∀ ⦃b : α⦄, b s∀ ⦃c : α⦄, c sa * c = b * bb = c
              theorem ThreeAPFree.eq_right {α : Type u_2} [AddCommMonoid α] [IsCancelAdd α] {s : Set α} (hs : ThreeAPFree s) a : α :
              a s∀ ⦃b : α⦄, b s∀ ⦃c : α⦄, c sa + c = b + bb = c
              theorem threeGPFree_insert {α : Type u_2} [CommMonoid α] [IsCancelMul α] {s : Set α} {a : α} :
              ThreeGPFree (insert a s) ThreeGPFree s (∀ ⦃b : α⦄, b s∀ ⦃c : α⦄, c sa * c = b * ba = b) ∀ ⦃b : α⦄, b s∀ ⦃c : α⦄, c sb * c = a * ab = a
              theorem threeAPFree_insert {α : Type u_2} [AddCommMonoid α] [IsCancelAdd α] {s : Set α} {a : α} :
              ThreeAPFree (insert a s) ThreeAPFree s (∀ ⦃b : α⦄, b s∀ ⦃c : α⦄, c sa + c = b + ba = b) ∀ ⦃b : α⦄, b s∀ ⦃c : α⦄, c sb + c = a + ab = a
              theorem ThreeGPFree.smul_set {α : Type u_2} [CommMonoid α] [IsCancelMul α] {s : Set α} {a : α} (hs : ThreeGPFree s) :
              theorem ThreeAPFree.vadd_set {α : Type u_2} [AddCommMonoid α] [IsCancelAdd α] {s : Set α} {a : α} (hs : ThreeAPFree s) :
              theorem threeGPFree_smul_set {α : Type u_2} [CommMonoid α] [IsCancelMul α] {s : Set α} {a : α} :
              theorem threeAPFree_vadd_set {α : Type u_2} [AddCommMonoid α] [IsCancelAdd α] {s : Set α} {a : α} :
              theorem threeGPFree_insert_of_lt {α : Type u_2} [CommMonoid α] [PartialOrder α] [IsOrderedCancelMonoid α] {s : Set α} {a : α} (hs : is, i < a) :
              ThreeGPFree (insert a s) ThreeGPFree s ∀ ⦃b : α⦄, b s∀ ⦃c : α⦄, c sa * c = b * ba = b
              theorem threeAPFree_insert_of_lt {α : Type u_2} [AddCommMonoid α] [PartialOrder α] [IsOrderedCancelAddMonoid α] {s : Set α} {a : α} (hs : is, i < a) :
              ThreeAPFree (insert a s) ThreeAPFree s ∀ ⦃b : α⦄, b s∀ ⦃c : α⦄, c sa + c = b + ba = b
              theorem ThreeGPFree.smul_set₀ {α : Type u_2} [CancelCommMonoidWithZero α] [NoZeroDivisors α] {s : Set α} {a : α} (hs : ThreeGPFree s) (ha : a 0) :
              theorem threeGPFree_smul_set₀ {α : Type u_2} [CancelCommMonoidWithZero α] [NoZeroDivisors α] {s : Set α} {a : α} (ha : a 0) :
              theorem threeAPFree_iff_eq_right {s : Set } :
              ThreeAPFree s ∀ ⦃a : ⦄, a s∀ ⦃b : ⦄, b s∀ ⦃c : ⦄, c sa + c = b + ba = c
              def mulRothNumber {α : Type u_2} [DecidableEq α] [Monoid α] :

              The multiplicative Roth number of a finset is the cardinality of its biggest 3GP-free subset.

              Equations
                Instances For
                  def addRothNumber {α : Type u_2} [DecidableEq α] [AddMonoid α] :

                  The additive Roth number of a finset is the cardinality of its biggest 3AP-free subset.

                  The usual Roth number corresponds to addRothNumber (Finset.range n), see rothNumberNat.

                  Equations
                    Instances For
                      theorem mulRothNumber_le {α : Type u_2} [DecidableEq α] [Monoid α] (s : Finset α) :
                      theorem addRothNumber_le {α : Type u_2} [DecidableEq α] [AddMonoid α] (s : Finset α) :
                      theorem mulRothNumber_spec {α : Type u_2} [DecidableEq α] [Monoid α] (s : Finset α) :
                      ts, t.card = mulRothNumber s ThreeGPFree t
                      theorem addRothNumber_spec {α : Type u_2} [DecidableEq α] [AddMonoid α] (s : Finset α) :
                      ts, t.card = addRothNumber s ThreeAPFree t
                      theorem ThreeGPFree.le_mulRothNumber {α : Type u_2} [DecidableEq α] [Monoid α] {s t : Finset α} (hs : ThreeGPFree s) (h : s t) :
                      theorem ThreeAPFree.le_addRothNumber {α : Type u_2} [DecidableEq α] [AddMonoid α] {s t : Finset α} (hs : ThreeAPFree s) (h : s t) :
                      theorem ThreeGPFree.mulRothNumber_eq {α : Type u_2} [DecidableEq α] [Monoid α] {s : Finset α} (hs : ThreeGPFree s) :
                      theorem ThreeAPFree.addRothNumber_eq {α : Type u_2} [DecidableEq α] [AddMonoid α] {s : Finset α} (hs : ThreeAPFree s) :
                      @[simp]
                      theorem mulRothNumber_empty {α : Type u_2} [DecidableEq α] [Monoid α] :
                      @[simp]
                      @[simp]
                      theorem mulRothNumber_singleton {α : Type u_2} [DecidableEq α] [Monoid α] (a : α) :
                      @[simp]
                      theorem addRothNumber_singleton {α : Type u_2} [DecidableEq α] [AddMonoid α] (a : α) :
                      theorem le_mulRothNumber_product {α : Type u_2} {β : Type u_3} [DecidableEq α] [Monoid α] [DecidableEq β] [Monoid β] (s : Finset α) (t : Finset β) :
                      theorem le_addRothNumber_product {α : Type u_2} {β : Type u_3} [DecidableEq α] [AddMonoid α] [DecidableEq β] [AddMonoid β] (s : Finset α) (t : Finset β) :
                      theorem mulRothNumber_lt_of_forall_not_threeGPFree {α : Type u_2} [DecidableEq α] [Monoid α] {s : Finset α} {n : } (h : tFinset.powersetCard n s, ¬ThreeGPFree t) :
                      theorem IsMulFreimanHom.mulRothNumber_mono {α : Type u_2} {β : Type u_3} [DecidableEq α] [CommMonoid α] [CommMonoid β] [DecidableEq β] {A : Finset α} {B : Finset β} {f : αβ} (hf : IsMulFreimanHom 2 (↑A) (↑B) f) (hf' : Set.BijOn f A B) :

                      Arithmetic progressions can be pushed forward along bijective 2-Freiman homs.

                      theorem IsAddFreimanHom.addRothNumber_mono {α : Type u_2} {β : Type u_3} [DecidableEq α] [AddCommMonoid α] [AddCommMonoid β] [DecidableEq β] {A : Finset α} {B : Finset β} {f : αβ} (hf : IsAddFreimanHom 2 (↑A) (↑B) f) (hf' : Set.BijOn f A B) :

                      Arithmetic progressions can be pushed forward along bijective 2-Freiman homs.

                      theorem IsMulFreimanIso.mulRothNumber_congr {α : Type u_2} {β : Type u_3} [DecidableEq α] [CommMonoid α] [CommMonoid β] [DecidableEq β] {A : Finset α} {B : Finset β} {f : αβ} (hf : IsMulFreimanIso 2 (↑A) (↑B) f) :

                      Arithmetic progressions are preserved under 2-Freiman isos.

                      theorem IsAddFreimanIso.addRothNumber_congr {α : Type u_2} {β : Type u_3} [DecidableEq α] [AddCommMonoid α] [AddCommMonoid β] [DecidableEq β] {A : Finset α} {B : Finset β} {f : αβ} (hf : IsAddFreimanIso 2 (↑A) (↑B) f) :

                      Arithmetic progressions are preserved under 2-Freiman isos.

                      The Roth number of a natural N is the largest integer m for which there is a subset of range N of size m with no arithmetic progression of length 3. Trivially, rothNumberNat N ≤ N, but Roth's theorem (proved in 1953) shows that rothNumberNat N = o(N) and the construction by Behrend gives a lower bound of the form N * exp(-C sqrt(log(N))) ≤ rothNumberNat N. A significant refinement of Roth's theorem by Bloom and Sisask announced in 2020 gives rothNumberNat N = O(N / (log N)^(1+c)) for an absolute constant c.

                      Equations
                        Instances For
                          theorem ThreeAPFree.le_rothNumberNat {k n : } (s : Finset ) (hs : ThreeAPFree s) (hsn : xs, x < n) (hsk : s.card = k) :

                          A verbose specialization of threeAPFree.le_addRothNumber, sometimes convenient in practice.

                          The Roth number is a subadditive function. Note that by Fekete's lemma this shows that the limit rothNumberNat N / N exists, but Roth's theorem gives the stronger result that this limit is actually 0.