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

Commit 740b41b

Browse files
committed
feat(data/fintype/basic): add finset.(sup|inf)_univ_eq_(supr|infi) (#8397)
1 parent 0c024a6 commit 740b41b

File tree

1 file changed

+8
-0
lines changed

1 file changed

+8
-0
lines changed

src/data/fintype/basic.lean

Lines changed: 8 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -170,6 +170,14 @@ lemma univ_filter_mem_range (f : α → β) [fintype β]
170170
finset.univ.filter (λ y, y ∈ set.range f) = finset.univ.image f :=
171171
univ_filter_exists f
172172

173+
/-- A special case of `finset.sup_eq_supr` that omits the useless `x ∈ univ` binder. -/
174+
lemma sup_univ_eq_supr [complete_lattice β] (f : α → β) : finset.univ.sup f = supr f :=
175+
(sup_eq_supr _ f).trans $ congr_arg _ $ funext $ λ a, supr_pos (mem_univ _)
176+
177+
/-- A special case of `finset.inf_eq_infi` that omits the useless `x ∈ univ` binder. -/
178+
lemma inf_univ_eq_infi [complete_lattice β] (f : α → β) : finset.univ.inf f = infi f :=
179+
sup_univ_eq_supr (by exact f : α → order_dual β)
180+
173181
end finset
174182

175183
open finset function

0 commit comments

Comments
 (0)