Skip to content
This repository was archived by the owner on Jul 24, 2024. It is now read-only.

Commit 299984b

Browse files
committed
feat(combinatorics/simple_graph/uniform): Graph uniformity and uniform partitions (#12957)
Define uniformity of a pair of vertices in a graph and uniformity of a partition of vertices of a graph. Co-authored-by: Bhavik Mehta <bhavik.mehta8@gmail.com>
1 parent 47b1d78 commit 299984b

File tree

1 file changed

+139
-0
lines changed

1 file changed

+139
-0
lines changed
Lines changed: 139 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,139 @@
1+
/-
2+
Copyright (c) 2022 Yaël Dillies, Bhavik Mehta. All rights reserved.
3+
Released under Apache 2.0 license as described in the file LICENSE.
4+
Authors: Yaël Dillies, Bhavik Mehta
5+
-/
6+
import combinatorics.simple_graph.density
7+
8+
/-!
9+
# Graph uniformity and uniform partitions
10+
11+
In this file we define uniformity of a pair of vertices in a graph and uniformity of a partition of
12+
vertices of a graph. Both are also known as ε-regularity.
13+
14+
Finsets of vertices `s` and `t` are `ε`-uniform in a graph `G` if their edge density is at most
15+
`ε`-far from the density of any big enough `s'` and `t'` where `s' ⊆ s`, `t' ⊆ t`.
16+
The definition is pretty technical, but it amounts to the edges between `s` and `t` being "random"
17+
The literature contains several definitions which are equivalent up to scaling `ε` by some constant
18+
when the partition is equitable.
19+
20+
A partition `P` of the vertices is `ε`-uniform if the proportion of non `ε`-uniform pairs of parts
21+
is less than `ε`.
22+
23+
## Main declarations
24+
25+
* `simple_graph.is_uniform`: Graph uniformity of a pair of finsets of vertices.
26+
* `finpartition.non_uniforms`: Non uniform pairs of parts of a partition.
27+
* `finpartition.is_uniform`: Uniformity of a partition.
28+
-/
29+
30+
open finset
31+
32+
variables {α 𝕜 : Type*} [linear_ordered_field 𝕜]
33+
34+
/-! ### Graph uniformity -/
35+
36+
namespace simple_graph
37+
variables (G : simple_graph α) [decidable_rel G.adj] (ε : 𝕜) {s t : finset α} {a b : α}
38+
39+
/-- A pair of finsets of vertices is `ε`-uniform (aka `ε`-regular) iff their edge density is close
40+
to the density of any big enough pair of subsets. Intuitively, the edges between them are
41+
random-like. -/
42+
def is_uniform (s t : finset α) : Prop :=
43+
∀ ⦃s'⦄, s' ⊆ s → ∀ ⦃t'⦄, t' ⊆ t → (s.card : 𝕜) * ε ≤ s'.card → (t.card : 𝕜) * ε ≤ t'.card →
44+
|(G.edge_density s' t' : 𝕜) - (G.edge_density s t : 𝕜)| < ε
45+
46+
variables {G ε}
47+
48+
lemma is_uniform.mono {ε' : 𝕜} (h : ε ≤ ε') (hε : is_uniform G ε s t) : is_uniform G ε' s t :=
49+
λ s' hs' t' ht' hs ht, by refine (hε hs' ht' (le_trans _ hs) (le_trans _ ht)).trans_le h;
50+
exact mul_le_mul_of_nonneg_left h (nat.cast_nonneg _)
51+
52+
lemma is_uniform.symm : symmetric (is_uniform G ε) :=
53+
λ s t h t' ht' s' hs' ht hs,
54+
by { rw [edge_density_comm _ t', edge_density_comm _ t], exact h hs' ht' hs ht }
55+
56+
variables (G)
57+
58+
lemma is_uniform_comm : is_uniform G ε s t ↔ is_uniform G ε t s := ⟨λ h, h.symm, λ h, h.symm⟩
59+
60+
lemma is_uniform_singleton (hε : 0 < ε) : G.is_uniform ε {a} {b} :=
61+
begin
62+
intros s' hs' t' ht' hs ht,
63+
rw [card_singleton, nat.cast_one, one_mul] at hs ht,
64+
obtain rfl | rfl := finset.subset_singleton_iff.1 hs',
65+
{ exact (hε.not_le hs).elim },
66+
obtain rfl | rfl := finset.subset_singleton_iff.1 ht',
67+
{ exact (hε.not_le ht).elim },
68+
{ rwa [sub_self, abs_zero] }
69+
end
70+
71+
lemma not_is_uniform_zero : ¬ G.is_uniform (0 : 𝕜) s t :=
72+
λ h, (abs_nonneg _).not_lt $ h (empty_subset _) (empty_subset _) (by simp) (by simp)
73+
74+
lemma is_uniform_one : G.is_uniform (1 : 𝕜) s t :=
75+
begin
76+
intros s' hs' t' ht' hs ht,
77+
rw mul_one at hs ht,
78+
rw [eq_of_subset_of_card_le hs' (nat.cast_le.1 hs),
79+
eq_of_subset_of_card_le ht' (nat.cast_le.1 ht), sub_self, abs_zero],
80+
exact zero_lt_one,
81+
end
82+
83+
end simple_graph
84+
85+
/-! ### Uniform partitions -/
86+
87+
variables [decidable_eq α] {s : finset α} (P : finpartition s) (G : simple_graph α)
88+
[decidable_rel G.adj] {ε : 𝕜}
89+
90+
namespace finpartition
91+
open_locale classical
92+
93+
/-- The pairs of parts of a partition `P` which are not `ε`-uniform in a graph `G`. Note that we
94+
dismiss the diagonal. We do not care whether `s` is `ε`-uniform with itself. -/
95+
noncomputable def non_uniforms (ε : 𝕜) : finset (finset α × finset α) :=
96+
P.parts.off_diag.filter $ λ uv, ¬G.is_uniform ε uv.1 uv.2
97+
98+
lemma mk_mem_non_uniforms_iff (u v : finset α) (ε : 𝕜) :
99+
(u, v) ∈ P.non_uniforms G ε ↔ u ∈ P.parts ∧ v ∈ P.parts ∧ u ≠ v ∧ ¬G.is_uniform ε u v :=
100+
by rw [non_uniforms, mem_filter, mem_off_diag, and_assoc, and_assoc]
101+
102+
/-- A finpartition is `ε`-uniform (aka `ε`-regular) iff at most a proportion of `ε` of its pairs of
103+
parts are not `ε-uniform`. -/
104+
def is_uniform (ε : 𝕜) : Prop :=
105+
((P.non_uniforms G ε).card : 𝕜) ≤ (P.parts.card * (P.parts.card - 1) : ℕ) * ε
106+
107+
lemma non_uniforms_bot (hε : 0 < ε) : (⊥ : finpartition s).non_uniforms G ε = ∅ :=
108+
begin
109+
rw eq_empty_iff_forall_not_mem,
110+
rintro ⟨u, v⟩,
111+
simp only [finpartition.mk_mem_non_uniforms_iff, finpartition.parts_bot, mem_map, not_and,
112+
not_not, exists_imp_distrib],
113+
rintro x hx rfl y hy rfl h,
114+
exact G.is_uniform_singleton hε,
115+
end
116+
117+
lemma bot_is_uniform (hε : 0 < ε) : (⊥ : finpartition s).is_uniform G ε :=
118+
begin
119+
rw [finpartition.is_uniform, finpartition.card_bot, non_uniforms_bot _ hε,
120+
finset.card_empty, nat.cast_zero],
121+
exact mul_nonneg (nat.cast_nonneg _) hε.le,
122+
end
123+
124+
lemma is_uniform_one : P.is_uniform G (1 : 𝕜) :=
125+
begin
126+
rw [is_uniform, mul_one, nat.cast_le],
127+
refine (card_filter_le _ _).trans _,
128+
rw [off_diag_card, nat.mul_sub_left_distrib, mul_one],
129+
end
130+
131+
variables {P G}
132+
133+
lemma is_uniform_of_empty (hP : P.parts = ∅) : P.is_uniform G ε :=
134+
by simp [is_uniform, hP, non_uniforms]
135+
136+
lemma nonempty_of_not_uniform (h : ¬ P.is_uniform G ε) : P.parts.nonempty :=
137+
nonempty_of_ne_empty $ λ h₁, h $ is_uniform_of_empty h₁
138+
139+
end finpartition

0 commit comments

Comments
 (0)