Skip to content

Commit

Permalink
feat(Finset): add gcongr attributes (#9520)
Browse files Browse the repository at this point in the history
Co-authored-by: Johan Commelin <johan@commelin.net>
  • Loading branch information
urkud and jcommelin committed Feb 29, 2024
1 parent d8924e0 commit efdfa67
Show file tree
Hide file tree
Showing 4 changed files with 28 additions and 22 deletions.
8 changes: 6 additions & 2 deletions Mathlib/Data/Finset/Basic.lean
Expand Up @@ -2866,15 +2866,17 @@ theorem filter_empty : filter p ∅ = ∅ :=
subset_empty.1 <| filter_subset _ _
#align finset.filter_empty Finset.filter_empty

@[gcongr]
theorem filter_subset_filter {s t : Finset α} (h : s ⊆ t) : s.filter p ⊆ t.filter p := fun _a ha =>
mem_filter.2 ⟨h (mem_filter.1 ha).1, (mem_filter.1 ha).2
#align finset.filter_subset_filter Finset.filter_subset_filter

theorem monotone_filter_left : Monotone (filter p) := fun _ _ => filter_subset_filter p
#align finset.monotone_filter_left Finset.monotone_filter_left

-- TODO: `@[gcongr]` doesn't accept this lemma because of the `DecidablePred` arguments
theorem monotone_filter_right (s : Finset α) ⦃p q : α → Prop⦄ [DecidablePred p] [DecidablePred q]
(h : p ≤ q) : s.filter p s.filter q :=
(h : p ≤ q) : s.filter p s.filter q :=
Multiset.subset_of_le (Multiset.monotone_filter_right s.val h)
#align finset.monotone_filter_right Finset.monotone_filter_right

Expand All @@ -2887,7 +2889,7 @@ theorem subset_coe_filter_of_subset_forall (s : Finset α) {t : Set α} (h₁ :
(h₂ : ∀ x ∈ t, p x) : t ⊆ s.filter p := fun x hx => (s.coe_filter p).symm ▸ ⟨h₁ hx, h₂ x hx⟩
#align finset.subset_coe_filter_of_subset_forall Finset.subset_coe_filter_of_subset_forall

theorem filter_singleton (a : α) : filter p (singleton a) = if p a then singleton a else ∅ := by
theorem filter_singleton (a : α) : filter p {a} = if p a then {a} else ∅ := by
classical
ext x
simp only [mem_singleton, forall_eq, mem_filter]
Expand Down Expand Up @@ -3171,6 +3173,8 @@ theorem range_subset {n m} : range n ⊆ range m ↔ n ≤ m :=
theorem range_mono : Monotone range := fun _ _ => range_subset.2
#align finset.range_mono Finset.range_mono

@[gcongr] alias ⟨_, _root_.GCongr.finset_range_subset_of_le⟩ := range_subset

theorem mem_range_succ_iff {a b : ℕ} : a ∈ Finset.range b.succ ↔ a ≤ b :=
Finset.mem_range.trans Nat.lt_succ_iff
#align finset.mem_range_succ_iff Finset.mem_range_succ_iff
Expand Down
1 change: 1 addition & 0 deletions Mathlib/Data/Finset/Card.lean
Expand Up @@ -57,6 +57,7 @@ theorem card_empty : card (∅ : Finset α) = 0 :=
rfl
#align finset.card_empty Finset.card_empty

@[gcongr]
theorem card_le_card : s ⊆ t → s.card ≤ t.card :=
Multiset.card_le_card ∘ val_le_iff.mpr
#align finset.card_le_of_subset Finset.card_le_card
Expand Down
16 changes: 16 additions & 0 deletions Mathlib/Data/Finset/Lattice.lean
Expand Up @@ -148,6 +148,7 @@ theorem sup_mono_fun {g : β → α} (h : ∀ b ∈ s, f b ≤ g b) : s.sup f
Finset.sup_le fun b hb => le_trans (h b hb) (le_sup hb)
#align finset.sup_mono_fun Finset.sup_mono_fun

@[gcongr]
theorem sup_mono (h : s₁ ⊆ s₂) : s₁.sup f ≤ s₂.sup f :=
Finset.sup_le (fun _ hb => le_sup (h hb))
#align finset.sup_mono Finset.sup_mono
Expand Down Expand Up @@ -418,6 +419,7 @@ theorem inf_mono_fun {g : β → α} (h : ∀ b ∈ s, f b ≤ g b) : s.inf f
Finset.le_inf fun b hb => le_trans (inf_le hb) (h b hb)
#align finset.inf_mono_fun Finset.inf_mono_fun

@[gcongr]
theorem inf_mono (h : s₁ ⊆ s₂) : s₂.inf f ≤ s₁.inf f :=
Finset.le_inf (fun _ hb => inf_le (h hb))
#align finset.inf_mono Finset.inf_mono
Expand Down Expand Up @@ -964,6 +966,12 @@ theorem sup'_mono {s₁ s₂ : Finset β} (h : s₁ ⊆ s₂) (h₁ : s₁.Nonem
s₁.sup' h₁ f ≤ s₂.sup' (h₁.mono h) f :=
Finset.sup'_le h₁ _ (fun _ hb => le_sup' _ (h hb))

/-- A version of `Finset.sup'_mono` acceptable for `@[gcongr]`. -/
@[gcongr]
lemma _root_.GCongr.finset_sup'_le {s₁ s₂ : Finset β} (h : s₁ ⊆ s₂)
{h₁ : s₁.Nonempty} {h₂ : s₂.Nonempty} : s₁.sup' h₁ f ≤ s₂.sup' h₂ f :=
sup'_mono f h h₁

end Sup'

section Inf'
Expand Down Expand Up @@ -1135,6 +1143,12 @@ theorem inf'_mono {s₁ s₂ : Finset β} (h : s₁ ⊆ s₂) (h₁ : s₁.Nonem
s₂.inf' (h₁.mono h) f ≤ s₁.inf' h₁ f :=
Finset.le_inf' h₁ _ (fun _ hb => inf'_le _ (h hb))

/-- A version of `Finset.inf'_mono` acceptable for `@[gcongr]`. -/
@[gcongr]
lemma _root_.GCongr.finset_inf'_mono {s₁ s₂ : Finset β} (h : s₁ ⊆ s₂)
{h₁ : s₁.Nonempty} {h₂ : s₂.Nonempty} : s₂.inf' h₂ f ≤ s₁.inf' h₁ f :=
inf'_mono f h h₁

end Inf'

section Sup
Expand Down Expand Up @@ -1400,6 +1414,7 @@ theorem not_mem_of_max_lt {s : Finset α} {a b : α} (h₁ : b < a) (h₂ : s.ma
Finset.not_mem_of_max_lt_coe <| h₂.trans_lt <| WithBot.coe_lt_coe.mpr h₁
#align finset.not_mem_of_max_lt Finset.not_mem_of_max_lt

@[gcongr]
theorem max_mono {s t : Finset α} (st : s ⊆ t) : s.max ≤ t.max :=
sup_mono st
#align finset.max_mono Finset.max_mono
Expand Down Expand Up @@ -1474,6 +1489,7 @@ theorem not_mem_of_lt_min {s : Finset α} {a b : α} (h₁ : a < b) (h₂ : s.mi
Finset.not_mem_of_coe_lt_min <| (WithTop.coe_lt_coe.mpr h₁).trans_eq h₂.symm
#align finset.not_mem_of_lt_min Finset.not_mem_of_lt_min

@[gcongr]
theorem min_mono {s t : Finset α} (st : s ⊆ t) : t.min ≤ s.min :=
inf_mono st
#align finset.min_mono Finset.min_mono
Expand Down
25 changes: 5 additions & 20 deletions Mathlib/Order/Filter/AtTopBot.lean
Expand Up @@ -1904,28 +1904,13 @@ theorem exists_seq_tendsto (f : Filter α) [IsCountablyGenerated f] [NeBot f] :
theorem exists_seq_monotone_tendsto_atTop_atTop (α : Type*) [SemilatticeSup α] [Nonempty α]
[(atTop : Filter α).IsCountablyGenerated] :
∃ xs : ℕ → α, Monotone xs ∧ Tendsto xs atTop atTop := by
haveI h_ne_bot : (atTop : Filter α).NeBot := atTop_neBot
obtain ⟨ys, h⟩ := exists_seq_tendsto (atTop : Filter α)
let xs : ℕ → α := fun n => Finset.sup' (Finset.range (n + 1)) Finset.nonempty_range_succ ys
have h_mono : Monotone xs := by
intro i j hij
rw [Finset.sup'_le_iff]
intro k hk
refine' Finset.le_sup'_of_le _ _ le_rfl
rw [Finset.mem_range] at hk ⊢
exact hk.trans_le (add_le_add_right hij _)
refine' ⟨xs, h_mono, _⟩
· refine' tendsto_atTop_atTop_of_monotone h_mono _
have : ∀ a : α, ∃ n : ℕ, a ≤ ys n := by
rw [tendsto_atTop_atTop] at h
intro a
obtain ⟨i, hi⟩ := h a
exact ⟨i, hi i le_rfl⟩
intro a
obtain ⟨i, hi⟩ := this a
refine' ⟨i, hi.trans _⟩
refine' Finset.le_sup'_of_le _ _ le_rfl
rw [Finset.mem_range_succ_iff]
have h_mono : Monotone xs := fun i j hij ↦ by
simp only -- Need to unfold `xs` and do alpha reduction, otherwise `gcongr` fails
gcongr
refine ⟨xs, h_mono, tendsto_atTop_mono (fun n ↦ Finset.le_sup' _ ?_) h⟩
simp
#align exists_seq_monotone_tendsto_at_top_at_top Filter.exists_seq_monotone_tendsto_atTop_atTop

theorem exists_seq_antitone_tendsto_atTop_atBot (α : Type*) [SemilatticeInf α] [Nonempty α]
Expand Down

0 comments on commit efdfa67

Please sign in to comment.