Skip to content

Commit

Permalink
feat: sups and limsups are measurable in conditionally complete linea…
Browse files Browse the repository at this point in the history
…r orders (#6979)

Currently, we only have that sups and limsups are measurable in complete linear orders, which excludes the main case of the real line. With more complicated proofs, these measurability results can be extended to all conditionally complete linear orders, without any further assumption in the statements.
  • Loading branch information
sgouezel committed Sep 7, 2023
1 parent f0c7f6d commit c719310
Show file tree
Hide file tree
Showing 5 changed files with 385 additions and 142 deletions.
11 changes: 10 additions & 1 deletion Mathlib/Data/Set/Lattice.lean
Expand Up @@ -2168,9 +2168,18 @@ end Disjoint

/-! ### Intervals -/


namespace Set

lemma nonempty_iInter_Iic_iff [Preorder α] {f : ι → α} :
(⋂ i, Iic (f i)).Nonempty ↔ BddBelow (range f) := by
have : (⋂ (i : ι), Iic (f i)) = lowerBounds (range f) := by
ext c; simp [lowerBounds]
simp [this, BddBelow]

lemma nonempty_iInter_Ici_iff [Preorder α] {f : ι → α} :
(⋂ i, Ici (f i)).Nonempty ↔ BddAbove (range f) :=
nonempty_iInter_Iic_iff (α := αᵒᵈ)

variable [CompleteLattice α]

theorem Ici_iSup (f : ι → α) : Ici (⨆ i, f i) = ⋂ i, Ici (f i) :=
Expand Down

0 comments on commit c719310

Please sign in to comment.