Skip to content

Commit

Permalink
feat: port Order.Filter.AtTopBot (#1795)
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 Jan 24, 2023
1 parent fec4dbb commit cea40eb
Show file tree
Hide file tree
Showing 3 changed files with 1,973 additions and 2 deletions.
1 change: 1 addition & 0 deletions Mathlib.lean
Expand Up @@ -655,6 +655,7 @@ import Mathlib.Order.Cover
import Mathlib.Order.Directed
import Mathlib.Order.Disjoint
import Mathlib.Order.Extension.Linear
import Mathlib.Order.Filter.AtTopBot
import Mathlib.Order.Filter.Bases
import Mathlib.Order.Filter.Basic
import Mathlib.Order.Filter.CountableInter
Expand Down
11 changes: 9 additions & 2 deletions Mathlib/Order/ConditionallyCompleteLattice/Basic.lean
Expand Up @@ -828,10 +828,17 @@ theorem cinfᵢ_unique [Unique ι] {s : ι → α} : (⨅ i, s i) = s default :=
@csupᵢ_unique αᵒᵈ _ _ _ _
#align infi_unique cinfᵢ_unique

-- porting note: new lemma
theorem csupᵢ_subsingleton [Subsingleton ι] (i : ι) (s : ι → α) : (⨆ i, s i) = s i :=
@csupᵢ_unique α ι _ ⟨⟨i⟩, fun j => Subsingleton.elim j i⟩ _

-- porting note: new lemma
theorem cinfᵢ_subsingleton [Subsingleton ι] (i : ι) (s : ι → α) : (⨅ i, s i) = s i :=
@cinfᵢ_unique α ι _ ⟨⟨i⟩, fun j => Subsingleton.elim j i⟩ _

@[simp]
theorem csupᵢ_pos {p : Prop} {f : p → α} (hp : p) : (⨆ h : p, f h) = f hp :=
haveI := uniqueProp hp
csupᵢ_unique
csupᵢ_subsingleton hp f
#align csupr_pos csupᵢ_pos

@[simp]
Expand Down

0 comments on commit cea40eb

Please sign in to comment.