Skip to content

Commit

Permalink
feat: Basic finset lemmas (#9225)
Browse files Browse the repository at this point in the history
From LeanAPAP and LeanCamCombi
  • Loading branch information
YaelDillies committed Dec 26, 2023
1 parent 2bb32b6 commit 401acd5
Show file tree
Hide file tree
Showing 3 changed files with 65 additions and 2 deletions.
45 changes: 43 additions & 2 deletions Mathlib/Data/Finset/Basic.lean
Expand Up @@ -477,8 +477,7 @@ theorem coe_coeEmb : ⇑(coeEmb : Finset α ↪o Set α) = ((↑) : Finset α
/-- The property `s.Nonempty` expresses the fact that the finset `s` is not empty. It should be used
in theorem assumptions instead of `∃ x, x ∈ s` or `s ≠ ∅` as it gives access to a nice API thanks
to the dot notation. -/
protected def Nonempty (s : Finset α) : Prop :=
∃ x : α, x ∈ s
@[pp_dot] protected def Nonempty (s : Finset α) : Prop := ∃ x : α, x ∈ s
#align finset.nonempty Finset.Nonempty

--Porting note: Much longer than in Lean3
Expand Down Expand Up @@ -738,6 +737,12 @@ theorem coe_eq_singleton {s : Finset α} {a : α} : (s : Set α) = {a} ↔ s = {
rw [← coe_singleton, coe_inj]
#align finset.coe_eq_singleton Finset.coe_eq_singleton

@[norm_cast]
lemma coe_subset_singleton : (s : Set α) ⊆ {a} ↔ s ⊆ {a} := by rw [← coe_subset, coe_singleton]

@[norm_cast]
lemma singleton_subset_coe : {a} ⊆ (s : Set α) ↔ {a} ⊆ s := by rw [← coe_subset, coe_singleton]

theorem eq_singleton_iff_unique_mem {s : Finset α} {a : α} : s = {a} ↔ a ∈ s ∧ ∀ x ∈ s, x = a := by
constructor <;> intro t
· rw [t]
Expand Down Expand Up @@ -1918,6 +1923,12 @@ theorem erase_empty (a : α) : erase ∅ a = ∅ :=
rfl
#align finset.erase_empty Finset.erase_empty

@[simp] lemma erase_nonempty (ha : a ∈ s) : (s.erase a).Nonempty ↔ s.Nontrivial := by
simp only [Finset.Nonempty, mem_erase, and_comm (b := _ ∈ _)]
refine ⟨?_, fun hs ↦ hs.exists_ne a⟩
rintro ⟨b, hb, hba⟩
exact ⟨_, hb, _, ha, hba⟩

@[simp]
theorem erase_singleton (a : α) : ({a} : Finset α).erase a = ∅ := by
ext x
Expand Down Expand Up @@ -2075,6 +2086,19 @@ theorem erase_injOn' (a : α) : { s : Finset α | a ∈ s }.InjOn fun s => erase
fun s hs t ht (h : s.erase a = _) => by rw [← insert_erase hs, ← insert_erase ht, h]
#align finset.erase_inj_on' Finset.erase_injOn'

lemma Nonempty.exists_cons_eq (hs : s.Nonempty) : ∃ t a ha, cons a t ha = s := by
classical
obtain ⟨a, ha⟩ := hs
exact ⟨s.erase a, a, not_mem_erase _ _, by simp [insert_erase ha]⟩

lemma Nontrivial.exists_cons_eq (hs : s.Nontrivial) :
∃ t a ha b hb hab, (cons b t hb).cons a (mem_cons.not.2 $ not_or_intro hab ha) = s := by
classical
obtain ⟨a, ha, b, hb, hab⟩ := hs
have : b ∈ s.erase a := mem_erase.2 ⟨hab.symm, hb⟩
refine ⟨(s.erase a).erase b, a, ?_, b, ?_, ?_, ?_⟩ <;>
simp [insert_erase this, insert_erase ha, *]

end Erase

/-! ### sdiff -/
Expand Down Expand Up @@ -2246,6 +2270,15 @@ theorem insert_sdiff_insert (s t : Finset α) (x : α) : insert x s \ insert x t
insert_sdiff_of_mem _ (mem_insert_self _ _)
#align finset.insert_sdiff_insert Finset.insert_sdiff_insert

lemma insert_sdiff_insert' (hab : a ≠ b) (ha : a ∉ s) : insert a s \ insert b s = {a} := by
ext; aesop

lemma erase_sdiff_erase (hab : a ≠ b) (hb : b ∈ s) : s.erase a \ s.erase b = {b} := by
ext; aesop

lemma cons_sdiff_cons (hab : a ≠ b) (ha hb) : s.cons a ha \ s.cons b hb = {a} := by
rw [cons_eq_insert, cons_eq_insert, insert_sdiff_insert' hab ha]

theorem sdiff_insert_of_not_mem {x : α} (h : x ∉ s) (t : Finset α) : s \ insert x t = s \ t := by
refine' Subset.antisymm (sdiff_subset_sdiff (Subset.refl _) (subset_insert _ _)) fun y hy => _
simp only [mem_sdiff, mem_insert, not_or] at hy ⊢
Expand Down Expand Up @@ -2289,6 +2322,12 @@ theorem disjoint_erase_comm : Disjoint (s.erase a) t ↔ Disjoint s (t.erase a)
simp_rw [erase_eq, disjoint_sdiff_comm]
#align finset.disjoint_erase_comm Finset.disjoint_erase_comm

lemma disjoint_insert_erase (ha : a ∉ t) : Disjoint (s.erase a) (insert a t) ↔ Disjoint s t := by
rw [disjoint_erase_comm, erase_insert ha]

lemma disjoint_erase_insert (ha : a ∉ s) : Disjoint (insert a s) (t.erase a) ↔ Disjoint s t := by
rw [← disjoint_erase_comm, erase_insert ha]

theorem disjoint_of_erase_left (ha : a ∉ t) (hst : Disjoint (s.erase a) t) : Disjoint s t := by
rw [← erase_insert ha, ← disjoint_erase_comm, disjoint_insert_right]
exact ⟨not_mem_erase _ _, hst⟩
Expand Down Expand Up @@ -2490,6 +2529,8 @@ theorem attach_nonempty_iff {s : Finset α} : s.attach.Nonempty ↔ s.Nonempty :
simp [Finset.Nonempty]
#align finset.attach_nonempty_iff Finset.attach_nonempty_iff

protected alias ⟨_, Nonempty.attach⟩ := attach_nonempty_iff

@[simp]
theorem attach_eq_empty_iff {s : Finset α} : s.attach = ∅ ↔ s = ∅ := by
simp [eq_empty_iff_forall_not_mem]
Expand Down
16 changes: 16 additions & 0 deletions Mathlib/Data/Fintype/Basic.lean
Expand Up @@ -178,6 +178,11 @@ theorem coe_compl (s : Finset α) : ↑sᶜ = (↑s : Set α)ᶜ :=
@[simp] lemma compl_subset_compl : sᶜ ⊆ tᶜ ↔ t ⊆ s := @compl_le_compl_iff_le (Finset α) _ _ _
@[simp] lemma compl_ssubset_compl : sᶜ ⊂ tᶜ ↔ t ⊂ s := @compl_lt_compl_iff_lt (Finset α) _ _ _

lemma subset_compl_comm : s ⊆ tᶜ ↔ t ⊆ sᶜ := le_compl_iff_le_compl (α := Finset α)

@[simp] lemma subset_compl_singleton : s ⊆ {a}ᶜ ↔ a ∉ s := by
rw [subset_compl_comm, singleton_subset_iff, mem_compl]

@[simp]
theorem compl_empty : (∅ : Finset α)ᶜ = univ :=
compl_bot
Expand Down Expand Up @@ -276,6 +281,11 @@ theorem image_univ_equiv [Fintype β] (f : β ≃ α) : univ.image f = univ :=

end BooleanAlgebra

-- @[simp] --Note this would loop with `Finset.univ_unique`
lemma singleton_eq_univ [Subsingleton α] (a : α) : ({a} : Finset α) = univ := by
ext b; simp [Subsingleton.elim a b]


theorem map_univ_of_surjective [Fintype β] {f : β ↪ α} (hf : Surjective f) : univ.map f = univ :=
eq_univ_of_forall <| hf.forall.2 fun _ => mem_map_of_mem _ <| mem_univ _
#align finset.map_univ_of_surjective Finset.map_univ_of_surjective
Expand Down Expand Up @@ -360,6 +370,9 @@ instance decidableMemRangeFintype [Fintype α] [DecidableEq β] (f : α → β)
DecidablePred (· ∈ Set.range f) := fun _ => Fintype.decidableExistsFintype
#align fintype.decidable_mem_range_fintype Fintype.decidableMemRangeFintype

instance decidableSubsingleton [Fintype α] [DecidableEq α] {s : Set α} [DecidablePred (· ∈ s)] :
Decidable s.Subsingleton := decidable_of_iff (∀ a ∈ s, ∀ b ∈ s, a = b) Iff.rfl

section BundledHoms

instance decidableEqEquivFintype [DecidableEq β] [Fintype α] : DecidableEq (α ≃ β) := fun a b =>
Expand Down Expand Up @@ -472,6 +485,9 @@ namespace Finset

variable [Fintype α] [DecidableEq α] {s t : Finset α}

@[simp]
lemma filter_univ_mem (s : Finset α) : univ.filter (· ∈ s) = s := by simp [filter_mem_eq_inter]

instance decidableCodisjoint : Decidable (Codisjoint s t) :=
decidable_of_iff _ codisjoint_left.symm
#align finset.decidable_codisjoint Finset.decidableCodisjoint
Expand Down
6 changes: 6 additions & 0 deletions Mathlib/Data/Fintype/Card.lean
Expand Up @@ -398,6 +398,12 @@ theorem Fintype.card_lex (α : Type*) [Fintype α] : Fintype.card (Lex α) = Fin
rfl
#align fintype.card_lex Fintype.card_lex

@[simp] lemma Fintype.card_multiplicative (α : Type*) [Fintype α] :
card (Multiplicative α) = card α := Finset.card_map _

@[simp] lemma Fintype.card_additive (α : Type*) [Fintype α] : card (Additive α) = card α :=
Finset.card_map _

/-- Given that `α ⊕ β` is a fintype, `α` is also a fintype. This is non-computable as it uses
that `Sum.inl` is an injection, but there's no clear inverse if `α` is empty. -/
noncomputable def Fintype.sumLeft {α β} [Fintype (Sum α β)] : Fintype α :=
Expand Down

0 comments on commit 401acd5

Please sign in to comment.