Skip to content

Commit 1a59e69

Browse files
committed
Start on SCD
1 parent fcb078e commit 1a59e69

File tree

6 files changed

+203
-9
lines changed

6 files changed

+203
-9
lines changed

LeanCamCombi.lean

Lines changed: 4 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -30,8 +30,10 @@ import LeanCamCombi.Mathlib.Combinatorics.SimpleGraph.Density
3030
import LeanCamCombi.Mathlib.Combinatorics.SimpleGraph.Finite
3131
import LeanCamCombi.Mathlib.Combinatorics.SimpleGraph.Maps
3232
import LeanCamCombi.Mathlib.Combinatorics.SimpleGraph.Subgraph
33+
import LeanCamCombi.Mathlib.Data.Finset.Image
3334
import LeanCamCombi.Mathlib.Data.List.DropRight
3435
import LeanCamCombi.Mathlib.Data.Multiset.Basic
36+
import LeanCamCombi.Mathlib.Data.Set.Function
3537
import LeanCamCombi.Mathlib.Data.Set.Pointwise.SMul
3638
import LeanCamCombi.Mathlib.GroupTheory.OrderOfElement
3739
import LeanCamCombi.Mathlib.LinearAlgebra.AffineSpace.FiniteDimensional
@@ -46,6 +48,8 @@ import LeanCamCombi.Multipartite
4648
import LeanCamCombi.PhD.VCDim.AddVCDim
4749
import LeanCamCombi.PhD.VCDim.HausslerPacking
4850
import LeanCamCombi.PhD.VCDim.HypercubeEdges
51+
import LeanCamCombi.PlainCombi.Chap1.Sec1.SCD
52+
import LeanCamCombi.PlainCombi.Chap1.Sec1.SDSS
4953
import LeanCamCombi.PlainCombi.LittlewoodOfford
5054
import LeanCamCombi.PlainCombi.OrderShatter
5155
import LeanCamCombi.PlainCombi.VanDenBergKesten
Lines changed: 10 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,10 @@
1+
import Mathlib.Data.Finset.Image
2+
import LeanCamCombi.Mathlib.Data.Set.Function
3+
4+
namespace Finset
5+
variable {α β : Type*} [DecidableEq β] {s t : Finset α} {f : α → β}
6+
7+
lemma disjoint_image' (hf : Set.InjOn f (s ∪ t)) :
8+
Disjoint (image f s) (image f t) ↔ Disjoint s t := mod_cast Set.disjoint_image' hf
9+
10+
end Finset
Lines changed: 11 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,11 @@
1+
import Mathlib.Data.Set.Function
2+
3+
namespace Set
4+
variable {α β : Type*} {s t : Set α} {f : α → β}
5+
6+
lemma disjoint_image' (hf : (s ∪ t).InjOn f) : Disjoint (f '' s) (f '' t) ↔ Disjoint s t where
7+
mp := .of_image
8+
mpr hst := disjoint_image_image fun _a ha _b hb ↦
9+
hf.ne (.inl ha) (.inr hb) fun H ↦ Set.disjoint_left.1 hst ha (H ▸ hb)
10+
11+
end Set
Lines changed: 103 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,103 @@
1+
/-
2+
Copyright (c) 2022 Yaël Dillies. All rights reserved.
3+
Released under Apache 2.0 license as described in the file LICENSE.
4+
Authors: Yaël Dillies
5+
-/
6+
import Mathlib.Order.Partition.Finpartition
7+
import LeanCamCombi.Mathlib.Data.Finset.Image
8+
9+
/-!
10+
# Symmetric Chain Decompositions
11+
12+
This file shows that if `α` is finite then `Finset α` has a decomposition into symmetric chains,
13+
namely chains of the form `{Cᵢ, ..., Cₙ₋ᵢ}` where `card α = n` and `|C_j| = j`.
14+
-/
15+
16+
open Finset
17+
18+
namespace Finpartition
19+
variable {α : Type*} [DecidableEq α] {𝒜 : Finset (Finset α)} {𝒞 𝒟 : Finset (Finset α)}
20+
{C s : Finset α} {a : α}
21+
22+
def extendExchange (a : α) (𝒞 : Finset (Finset α)) (C : Finset α) : Finset (Finset (Finset α)) :=
23+
if 𝒞.Nontrivial
24+
then {(𝒞 \ {C}).image (insert a), 𝒞 ∪ {insert a C}}
25+
else {𝒞 ∪ {insert a C}}
26+
27+
lemma eq_or_eq_of_mem_extendExchange :
28+
𝒟 ∈ extendExchange a 𝒞 C → 𝒟 = (𝒞 \ {C}).image (insert a) ∨ 𝒟 = 𝒞 ∪ {insert a C} := by
29+
unfold extendExchange; split_ifs with h𝒞 <;> simp (config := { contextual := true })
30+
31+
@[simp] lemma not_empty_mem_extendExchange : ∅ ∉ extendExchange a 𝒞 C := by
32+
unfold extendExchange
33+
split_ifs with h𝒞
34+
· simp [eq_comm (a := ∅), ← nonempty_iff_ne_empty, h𝒞.nonempty.ne_empty, h𝒞.ne_singleton]
35+
· simp [eq_comm (a := ∅), ← nonempty_iff_ne_empty]
36+
37+
@[simp] lemma sup_extendExchange (hC : C ∈ 𝒞) :
38+
(extendExchange a 𝒞 C).sup id = 𝒞 ∪ 𝒞.image (insert a) := by
39+
unfold extendExchange
40+
split_ifs with h𝒞
41+
· simp only [sup_insert, id_eq, sup_singleton, sup_eq_union, union_left_comm]
42+
rw [image_sdiff_of_injOn, image_singleton, sdiff_union_of_subset]
43+
simpa using mem_image_of_mem _ hC
44+
sorry
45+
simpa
46+
· obtain rfl := (eq_singleton_or_nontrivial hC).resolve_right h𝒞
47+
simp
48+
49+
def extendPowersetExchange (P : Finpartition s.powerset) (f : ∀ 𝒞 ∈ P.parts, 𝒞) (a : α)
50+
(ha : a ∉ s) : Finpartition (s.cons a ha).powerset where
51+
52+
-- ofErase
53+
-- (P.parts.attach.biUnion fun ⟨𝒞, h𝒞⟩ ↦ extendExchange a 𝒞 (f 𝒞 h𝒞).1)
54+
-- (by
55+
-- sorry
56+
-- )
57+
-- (by
58+
-- simp
59+
-- )
60+
61+
parts := P.parts.attach.biUnion fun ⟨𝒞, h𝒞⟩ ↦ extendExchange a 𝒞 (f 𝒞 h𝒞).1
62+
supIndep := by
63+
rw [Finset.supIndep_iff_pairwiseDisjoint]
64+
simp only [Set.PairwiseDisjoint, Set.Pairwise, Finset.coe_biUnion, Finset.mem_coe,
65+
Finset.mem_attach, Set.iUnion_true, Set.mem_iUnion, Subtype.exists, ne_eq, Function.onFun,
66+
id_eq, forall_exists_index, not_imp_comm]
67+
rintro x 𝒞 h𝒞 hx y 𝒟 h𝒟 hy hxy
68+
obtain rfl | rfl := eq_or_eq_of_mem_extendExchange hx <;>
69+
obtain rfl | rfl := eq_or_eq_of_mem_extendExchange hy
70+
· rw [disjoint_image'] at hxy
71+
obtain rfl := P.disjoint.elim h𝒞 h𝒟 fun h ↦ hxy $ h.mono sdiff_le sdiff_le
72+
rfl
73+
sorry
74+
· sorry
75+
· sorry
76+
· simp_rw [disjoint_union_left, disjoint_union_right, and_assoc] at hxy
77+
obtain rfl := P.disjoint.elim h𝒞 h𝒟 fun h ↦ hxy $ ⟨h, sorry, sorry, sorry
78+
rfl
79+
sup_parts := by
80+
ext C
81+
simp only [sup_biUnion, coe_mem, sup_extendExchange, mem_sup, mem_attach, mem_union, mem_image,
82+
true_and, Subtype.exists, exists_prop, cons_eq_insert, mem_powerset]
83+
constructor
84+
· rintro ⟨𝒞, h𝒞, hC | ⟨C, hC, rfl⟩⟩
85+
· exact Subset.trans (mem_powerset.1 $ P.le h𝒞 hC) (subset_insert ..)
86+
· exact insert_subset_insert _ (mem_powerset.1 $ P.le h𝒞 hC)
87+
by_cases ha : a ∈ C
88+
· rw [subset_insert_iff]
89+
rintro hC
90+
obtain ⟨𝒞, h𝒞, hC⟩ := P.exists_mem $ mem_powerset.2 hC
91+
exact ⟨𝒞, h𝒞, .inr ⟨_, hC, insert_erase ha⟩⟩
92+
· rw [subset_insert_iff_of_not_mem ha]
93+
rintro hC
94+
obtain ⟨𝒞, h𝒞, hC⟩ := P.exists_mem $ mem_powerset.2 hC
95+
exact ⟨𝒞, h𝒞, .inl hC⟩
96+
not_bot_mem := by simp
97+
98+
/-! ### Profile of a partition -/
99+
100+
/-- The profile of -/
101+
def profile (P : Finpartition s) : Multiset ℕ := P.parts.1.map card
102+
103+
end Finpartition
Lines changed: 57 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,57 @@
1+
/-
2+
Copyright (c) 2022 Yaël Dillies. All rights reserved.
3+
Released under Apache 2.0 license as described ∈ the file LICENSE.
4+
Authors: Yaël Dillies
5+
-/
6+
import Mathlib.Analysis.NormedSpace.HahnBanach.Extension
7+
import Mathlib.Combinatorics.Enumerative.DoubleCounting
8+
import Mathlib.Order.Partition.Finpartition
9+
import LeanCamCombi.PlainCombi.Chap1.Sec1.SCD
10+
11+
/-!
12+
# Symmetric Decompositions into Sparse Sets
13+
14+
The Littlewood-Offord problem
15+
-/
16+
17+
open scoped BigOperators
18+
19+
namespace Finset
20+
variable {α E : Type*} {𝒜 : Finset (Finset α)}
21+
{s : Finset α} {f : α → E} {r : ℝ}
22+
23+
def profile (𝒜 : Finset (Finset α)) : Multiset ℕ := sorry
24+
25+
-- def subpartition (f : ):
26+
27+
variable [NormedAddCommGroup E] [NormedSpace ℝ E]
28+
29+
lemma exists_littlewood_offord_partition [DecidableEq α] (hr : 0 < r) (hf : ∀ i ∈ s, r ≤ ‖f i‖) :
30+
∃ P : Finpartition s.powerset,
31+
#P.parts = (#s).choose (s.card / 2) ∧
32+
(∀ 𝒜 ∈ P.parts, ∀ t ∈ 𝒜, t ⊆ s) ∧
33+
∀ 𝒜 ∈ P.parts,
34+
(𝒜 : Set (Finset α)).Pairwise fun u v ↦ r ≤ dist (∑ i ∈ u, f i) (∑ i ∈ v, f i) := by
35+
classical
36+
induction' s using Finset.induction with i s hi ih
37+
· exact ⟨Finpartition.indiscrete $ singleton_ne_empty _, by simp⟩
38+
obtain ⟨P, hP, hs, hPr⟩ := ih fun j hj ↦ hf _ $ mem_insert_of_mem hj
39+
clear ih
40+
obtain ⟨g, hg, hgf⟩ :=
41+
exists_dual_vector ℝ (f i) (norm_pos_iff.1 $ hr.trans_le $ hf _ $ mem_insert_self _ _)
42+
choose! t ht using fun 𝒜 (h𝒜 : 𝒜 ∈ P.parts) ↦
43+
Finset.exists_max_image _ (fun t ↦ g (∑ i ∈ t, f i)) (P.nonempty_of_mem_parts h𝒜)
44+
sorry
45+
46+
/-- **Kleitman's lemma** -/
47+
lemma card_le_of_forall_dist_sum_le (hr : 0 < r) (h𝒜 : ∀ t ∈ 𝒜, t ⊆ s) (hf : ∀ i ∈ s, r ≤ ‖f i‖)
48+
(h𝒜r : ∀ u, u ∈ 𝒜 → ∀ v, v ∈ 𝒜 → dist (∑ i ∈ u, f i) (∑ i ∈ v, f i) < r) :
49+
#𝒜 ≤ (#s).choose (#s / 2) := by
50+
classical
51+
obtain ⟨P, hP, _hs, hr⟩ := exists_littlewood_offord_partition hr hf
52+
rw [← hP]
53+
refine card_le_card_of_forall_subsingleton (· ∈ ·) (fun t ht ↦ ?_) fun ℬ hℬ t ht u hu ↦
54+
(hr _ hℬ).eq ht.2 hu.2 (h𝒜r _ ht.1 _ hu.1).not_le
55+
simpa only [exists_prop] using P.exists_mem (mem_powerset.2 $ h𝒜 _ ht)
56+
57+
end Finset

LeanCamCombi/PlainCombi/LittlewoodOfford.lean

Lines changed: 18 additions & 9 deletions
Original file line numberDiff line numberDiff line change
@@ -1,26 +1,35 @@
11
/-
22
Copyright (c) 2022 Yaël Dillies. All rights reserved.
3-
Released under Apache 2.0 license as described in the file LICENSE.
3+
Released under Apache 2.0 license as described the file LICENSE.
44
Authors: Yaël Dillies
55
-/
66
import Mathlib.Analysis.NormedSpace.HahnBanach.Extension
77
import Mathlib.Combinatorics.Enumerative.DoubleCounting
88
import Mathlib.Order.Partition.Finpartition
9+
import LeanCamCombi.PlainCombi.Chap1.Sec1.SCD
910

1011
/-!
11-
# The Littlewood-Offord problem
12+
# Symmetric Decompositions into Sparse Sets
13+
14+
The Littlewood-Offord problem
1215
-/
1316

1417
open scoped BigOperators
1518

1619
namespace Finset
17-
variable {ι E : Type*} [NormedAddCommGroup E] [NormedSpace ℝ E] {𝒜 : Finset (Finset ι)}
18-
{s : Finset ι} {f : ι → E} {r : ℝ}
20+
variable {α E : Type*} {𝒜 : Finset (Finset α)}
21+
{s : Finset α} {f : α → E} {r : ℝ}
22+
23+
-- def subpartition (f : ):
24+
25+
variable [NormedAddCommGroup E] [NormedSpace ℝ E]
1926

20-
lemma exists_littlewood_offord_partition [DecidableEq ι] (hr : 0 < r) (hf : ∀ i ∈ s, r ≤ ‖f i‖) :
27+
lemma exists_littlewood_offord_partition [DecidableEq α] (hr : 0 < r) (hf : ∀ i ∈ s, r ≤ ‖f i‖) :
2128
∃ P : Finpartition s.powerset,
22-
#P.parts = (#s).choose (#s / 2) ∧ (∀ 𝒜 ∈ P.parts, ∀ t ∈ 𝒜, t ⊆ s) ∧ ∀ 𝒜 ∈ P.parts,
23-
(𝒜 : Set (Finset ι)).Pairwise fun u v ↦ r ≤ dist (∑ i in u, f i) (∑ i in v, f i) := by
29+
#P.parts = (#s).choose (s.card / 2) ∧
30+
(∀ 𝒜 ∈ P.parts, ∀ t ∈ 𝒜, t ⊆ s) ∧
31+
∀ 𝒜 ∈ P.parts,
32+
(𝒜 : Set (Finset α)).Pairwise fun u v ↦ r ≤ dist (∑ i ∈ u, f i) (∑ i ∈ v, f i) := by
2433
classical
2534
induction' s using Finset.induction with i s hi ih
2635
· exact ⟨Finpartition.indiscrete $ singleton_ne_empty _, by simp⟩
@@ -29,12 +38,12 @@ lemma exists_littlewood_offord_partition [DecidableEq ι] (hr : 0 < r) (hf : ∀
2938
obtain ⟨g, hg, hgf⟩ :=
3039
exists_dual_vector ℝ (f i) (norm_pos_iff.1 $ hr.trans_le $ hf _ $ mem_insert_self _ _)
3140
choose! t ht using fun 𝒜 (h𝒜 : 𝒜 ∈ P.parts) ↦
32-
Finset.exists_max_image _ (fun t ↦ g (∑ i in t, f i)) (P.nonempty_of_mem_parts h𝒜)
41+
Finset.exists_max_image _ (fun t ↦ g (∑ i t, f i)) (P.nonempty_of_mem_parts h𝒜)
3342
sorry
3443

3544
/-- **Kleitman's lemma** -/
3645
lemma card_le_of_forall_dist_sum_le (hr : 0 < r) (h𝒜 : ∀ t ∈ 𝒜, t ⊆ s) (hf : ∀ i ∈ s, r ≤ ‖f i‖)
37-
(h𝒜r : ∀ u, u ∈ 𝒜 → ∀ v, v ∈ 𝒜 → dist (∑ i in u, f i) (∑ i in v, f i) < r) :
46+
(h𝒜r : ∀ u, u ∈ 𝒜 → ∀ v, v ∈ 𝒜 → dist (∑ i u, f i) (∑ i v, f i) < r) :
3847
#𝒜 ≤ (#s).choose (#s / 2) := by
3948
classical
4049
obtain ⟨P, hP, _hs, hr⟩ := exists_littlewood_offord_partition hr hf

0 commit comments

Comments
 (0)