Skip to content

Commit

Permalink
feat(combinatorics/set_family/shadow): Shadow of a set family (#10223)
Browse files Browse the repository at this point in the history
This defines `shadow π’œ` where `π’œ : finset (finset Ξ±))`.
  • Loading branch information
YaelDillies committed Nov 16, 2021
1 parent 7fec401 commit 0a2a922
Showing 1 changed file with 123 additions and 0 deletions.
123 changes: 123 additions & 0 deletions src/combinatorics/set_family/shadow.lean
@@ -0,0 +1,123 @@
/-
Copyright (c) 2021 Bhavik Mehta. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Bhavik Mehta, Alena Gusakov, YaΓ«l Dillies
-/
import data.finset.lattice

/-!
# Shadows
This file defines shadows of a set family. The shadow of a set family is the set family of sets we
get by removing any element from any set of the original family. If one pictures `finset Ξ±` as a big
hypercube (each dimension being membership of a given element), then taking the shadow corresponds
to projecting each finset down once in all available directions.
## Main definitions
The `shadow` of a set family is everything we can get by removing an element from each set.
## Notation
`βˆ‚ π’œ` is notation for `shadow π’œ`. It is situated in locale `finset_family`.
We also maintain the convention that `a, b : Ξ±` are elements of the ground type, `s, t : finset Ξ±`
are finsets, and `π’œ, ℬ : finset (finset Ξ±)` are finset families.
## References
* https://github.com/b-mehta/maths-notes/blob/master/iii/mich/combinatorics.pdf
* http://discretemath.imp.fu-berlin.de/DMII-2015-16/kruskal.pdf
## Tags
shadow, set family
-/

open finset nat

variables {Ξ± : Type*}

namespace finset
variables [decidable_eq Ξ±] {π’œ : finset (finset Ξ±)} {s t : finset Ξ±} {a : Ξ±} {k : β„•}

/-- The shadow of a set family `π’œ` is all sets we can get by removing one element from any set in
`π’œ`, and the (`k` times) iterated shadow (`shadow^[k]`) is all sets we can get by removing `k`
elements from any set in `π’œ`. -/
def shadow (π’œ : finset (finset Ξ±)) : finset (finset Ξ±) := π’œ.sup (Ξ» s, s.image (erase s))

localized "notation `βˆ‚ `:90 := finset.shadow" in finset_family

/-- The shadow of the empty set is empty. -/
@[simp] lemma shadow_empty : βˆ‚ (βˆ… : finset (finset Ξ±)) = βˆ… := rfl

/-- The shadow is monotone. -/
@[mono] lemma shadow_monotone : monotone (shadow : finset (finset Ξ±) β†’ finset (finset Ξ±)) :=
Ξ» π’œ ℬ, sup_mono

/-- `s` is in the shadow of `π’œ` iff there is an `t ∈ π’œ` from which we can remove one element to
get `s`. -/
lemma mem_shadow_iff : s ∈ βˆ‚ π’œ ↔ βˆƒ t ∈ π’œ, βˆƒ a ∈ t, erase t a = s :=
by simp only [shadow, mem_sup, mem_image]

lemma erase_mem_shadow (hs : s ∈ π’œ) (ha : a ∈ s) : erase s a ∈ βˆ‚ π’œ :=
mem_shadow_iff.2 ⟨s, hs, a, ha, rfl⟩

/-- `t` is in the shadow of `π’œ` iff we can add an element to it so that the resulting finset is in
`π’œ`. -/
lemma mem_shadow_iff_insert_mem : s ∈ βˆ‚ π’œ ↔ βˆƒ a βˆ‰ s, insert a s ∈ π’œ :=
begin
refine mem_shadow_iff.trans ⟨_, _⟩,
{ rintro ⟨s, hs, a, ha, rfl⟩,
refine ⟨a, not_mem_erase a s, _⟩,
rwa insert_erase ha },
{ rintro ⟨a, ha, hs⟩,
exact ⟨insert a s, hs, a, mem_insert_self _ _, erase_insert ha⟩ }
end

/-- `s ∈ βˆ‚ π’œ` iff `s` is exactly one element less than something from `π’œ` -/
lemma mem_shadow_iff_exists_mem_card_add_one :
s ∈ βˆ‚ π’œ ↔ βˆƒ t ∈ π’œ, s βŠ† t ∧ t.card = s.card + 1 :=
begin
refine mem_shadow_iff_insert_mem.trans ⟨_, _⟩,
{ rintro ⟨a, ha, hs⟩,
exact ⟨insert a s, hs, subset_insert _ _, card_insert_of_not_mem ha⟩ },
{ rintro ⟨t, ht, hst, h⟩,
obtain ⟨a, ha⟩ : βˆƒ a, t \ s = {a} :=
card_eq_one.1 (by rw [card_sdiff hst, h, add_tsub_cancel_left]),
exact ⟨a, λ hat,
not_mem_sdiff_of_mem_right hat ((ha.ge : _ βŠ† _) $ mem_singleton_self a),
by rwa [insert_eq a s, ←ha, sdiff_union_of_subset hst]⟩ }
end

/-- Being in the shadow of `π’œ` means we have a superset in `π’œ`. -/
lemma exists_subset_of_mem_shadow (hs : s ∈ βˆ‚ π’œ) : βˆƒ t ∈ π’œ, s βŠ† t :=
let ⟨t, ht, hst⟩ := mem_shadow_iff_exists_mem_card_add_one.1 hs in ⟨t, ht, hst.1⟩

/-- `t ∈ βˆ‚^k π’œ` iff `t` is exactly `k` elements less than something in `π’œ`. -/
lemma mem_shadow_iff_exists_mem_card_add :
s ∈ (βˆ‚^[k]) π’œ ↔ βˆƒ t ∈ π’œ, s βŠ† t ∧ t.card = s.card + k :=
begin
induction k with k ih generalizing π’œ s,
{ refine ⟨λ hs, ⟨s, hs, subset.refl _, rfl⟩, _⟩,
rintro ⟨t, ht, hst, hcard⟩,
rwa eq_of_subset_of_card_le hst hcard.le },
simp only [exists_prop, function.comp_app, function.iterate_succ],
refine ih.trans _,
clear ih,
split,
{ rintro ⟨t, ht, hst, hcardst⟩,
obtain ⟨u, hu, htu, hcardtu⟩ := mem_shadow_iff_exists_mem_card_add_one.1 ht,
refine ⟨u, hu, hst.trans htu, _⟩,
rw [hcardtu, hcardst],
refl },
{ rintro ⟨t, ht, hst, hcard⟩,
obtain ⟨u, hsu, hut, hu⟩ := finset.exists_intermediate_set k
(by { rw [add_comm, hcard], exact le_succ _ }) hst,
rw add_comm at hu,
refine ⟨u, mem_shadow_iff_exists_mem_card_add_one.2 ⟨t, ht, hut, _⟩, hsu, hu⟩,
rw [hcard, hu],
refl }
end

end finset

0 comments on commit 0a2a922

Please sign in to comment.