Skip to content

Commit

Permalink
feat(data/finset): add lemma for empty filter (#4188)
Browse files Browse the repository at this point in the history
A little lemma, analogous to `filter_true_of_mem` to make it convenient to reduce a filter which always fails.
  • Loading branch information
b-mehta committed Sep 20, 2020
1 parent db9842c commit f77d5d6
Show file tree
Hide file tree
Showing 2 changed files with 6 additions and 3 deletions.
7 changes: 5 additions & 2 deletions src/data/finset/basic.lean
Expand Up @@ -932,11 +932,14 @@ by ext; simp
@[simp] theorem filter_false {h} (s : finset α) : @filter α (λa, false) h s = ∅ :=
ext $ assume a, by simp only [mem_filter, and_false]; refl

/-- If all elements of a `finset` satisfy the predicate `p`, `s.filter
p` is `s`. -/
/-- If all elements of a `finset` satisfy the predicate `p`, `s.filter p` is `s`. -/
@[simp] lemma filter_true_of_mem {s : finset α} (h : ∀ x ∈ s, p x) : s.filter p = s :=
ext $ λ x, ⟨λ h, (mem_filter.1 h).1, λ hx, mem_filter.2 ⟨hx, h x hx⟩⟩

/-- If all elements of a `finset` fail to satisfy the predicate `p`, `s.filter p` is `∅`. -/
lemma filter_false_of_mem {s : finset α} (h : ∀ x ∈ s, ¬ p x) : s.filter p = ∅ :=
eq_empty_of_forall_not_mem (by simpa)

lemma filter_congr {s : finset α} (H : ∀ x ∈ s, p x ↔ q x) : filter p s = filter q s :=
eq_of_veq $ filter_congr H

Expand Down
2 changes: 1 addition & 1 deletion src/data/monoid_algebra.lean
Expand Up @@ -90,7 +90,7 @@ calc (f * g) x = (∑ a₁ in f.support, ∑ a₂ in g.support, F (a₁, a₂))
... = ∑ p in (f.support.product g.support).filter (λ p : G × G, p.1 * p.2 = x), f p.1 * g p.2 :
(finset.sum_filter _ _).symm
... = ∑ p in s.filter (λ p : G × G, p.1 ∈ f.support ∧ p.2 ∈ g.support), f p.1 * g p.2 :
sum_congr (by { ext, simp [hs, and_comm] }) (λ _ _, rfl)
sum_congr (by { ext, simp only [mem_filter, mem_product, hs, and_comm] }) (λ _ _, rfl)
... = ∑ p in s, f p.1 * g p.2 : sum_subset (filter_subset _) $ λ p hps hp,
begin
simp only [mem_filter, mem_support_iff, not_and, not_not] at hp ⊢,
Expand Down

0 comments on commit f77d5d6

Please sign in to comment.