mathlib documentation

order.filter.countable_Inter

Filters with countable intersection property #

In this file we define countable_Inter_filter to be the class of filters with the following property: for any countable collection of sets s ∈ l their intersection belongs to l as well.

Two main examples are the residual filter defined in topology.metric_space.baire and the measure.ae filter defined in measure_theory.measure_space.

We reformulate the definition in terms of indexed intersection and in terms of filter.eventually and provide instances for some basic constructions (, , filter.principal, filter.map, filter.comap, has_inf.inf). We also provide a custom constructor filter.of_countable_Inter that deduces two axioms of a filter from the countable intersection property.

Tags #

filter, countable

@[class]
structure countable_Inter_filter {α : Type u_2} (l : filter α) :
Prop

A filter l has the countable intersection property if for any countable collection of sets s ∈ l their intersection belongs to l as well.

Instances of this typeclass
theorem countable_sInter_mem {α : Type u_2} {l : filter α} [countable_Inter_filter l] {S : set (set α)} (hSc : S.countable) :
⋂₀ S l ∀ (s : set α), s Ss l
theorem countable_Inter_mem {ι : Sort u_1} {α : Type u_2} {l : filter α} [countable_Inter_filter l] [countable ι] {s : ι → set α} :
(⋂ (i : ι), s i) l ∀ (i : ι), s i l
theorem countable_bInter_mem {α : Type u_2} {l : filter α} [countable_Inter_filter l] {ι : Type u_1} {S : set ι} (hS : S.countable) {s : Π (i : ι), i Sset α} :
(⋂ (i : ι) (H : i S), s i H) l ∀ (i : ι) (H : i S), s i H l
theorem eventually_countable_forall {ι : Sort u_1} {α : Type u_2} {l : filter α} [countable_Inter_filter l] [countable ι] {p : α → ι → Prop} :
(∀ᶠ (x : α) in l, ∀ (i : ι), p x i) ∀ (i : ι), ∀ᶠ (x : α) in l, p x i
theorem eventually_countable_ball {α : Type u_2} {l : filter α} [countable_Inter_filter l] {ι : Type u_1} {S : set ι} (hS : S.countable) {p : α → Π (i : ι), i S → Prop} :
(∀ᶠ (x : α) in l, ∀ (i : ι) (H : i S), p x i H) ∀ (i : ι) (H : i S), ∀ᶠ (x : α) in l, p x i H
theorem eventually_le.countable_Union {ι : Sort u_1} {α : Type u_2} {l : filter α} [countable_Inter_filter l] [countable ι] {s t : ι → set α} (h : ∀ (i : ι), s i ≤ᶠ[l] t i) :
(⋃ (i : ι), s i) ≤ᶠ[l] ⋃ (i : ι), t i
theorem eventually_eq.countable_Union {ι : Sort u_1} {α : Type u_2} {l : filter α} [countable_Inter_filter l] [countable ι] {s t : ι → set α} (h : ∀ (i : ι), s i =ᶠ[l] t i) :
(⋃ (i : ι), s i) =ᶠ[l] ⋃ (i : ι), t i
theorem eventually_le.countable_bUnion {α : Type u_2} {l : filter α} [countable_Inter_filter l] {ι : Type u_1} {S : set ι} (hS : S.countable) {s t : Π (i : ι), i Sset α} (h : ∀ (i : ι) (H : i S), s i H ≤ᶠ[l] t i H) :
(⋃ (i : ι) (H : i S), s i H) ≤ᶠ[l] ⋃ (i : ι) (H : i S), t i H
theorem eventually_eq.countable_bUnion {α : Type u_2} {l : filter α} [countable_Inter_filter l] {ι : Type u_1} {S : set ι} (hS : S.countable) {s t : Π (i : ι), i Sset α} (h : ∀ (i : ι) (H : i S), s i H =ᶠ[l] t i H) :
(⋃ (i : ι) (H : i S), s i H) =ᶠ[l] ⋃ (i : ι) (H : i S), t i H
theorem eventually_le.countable_Inter {ι : Sort u_1} {α : Type u_2} {l : filter α} [countable_Inter_filter l] [countable ι] {s t : ι → set α} (h : ∀ (i : ι), s i ≤ᶠ[l] t i) :
(⋂ (i : ι), s i) ≤ᶠ[l] ⋂ (i : ι), t i
theorem eventually_eq.countable_Inter {ι : Sort u_1} {α : Type u_2} {l : filter α} [countable_Inter_filter l] [countable ι] {s t : ι → set α} (h : ∀ (i : ι), s i =ᶠ[l] t i) :
(⋂ (i : ι), s i) =ᶠ[l] ⋂ (i : ι), t i
theorem eventually_le.countable_bInter {α : Type u_2} {l : filter α} [countable_Inter_filter l] {ι : Type u_1} {S : set ι} (hS : S.countable) {s t : Π (i : ι), i Sset α} (h : ∀ (i : ι) (H : i S), s i H ≤ᶠ[l] t i H) :
(⋂ (i : ι) (H : i S), s i H) ≤ᶠ[l] ⋂ (i : ι) (H : i S), t i H
theorem eventually_eq.countable_bInter {α : Type u_2} {l : filter α} [countable_Inter_filter l] {ι : Type u_1} {S : set ι} (hS : S.countable) {s t : Π (i : ι), i Sset α} (h : ∀ (i : ι) (H : i S), s i H =ᶠ[l] t i H) :
(⋂ (i : ι) (H : i S), s i H) =ᶠ[l] ⋂ (i : ι) (H : i S), t i H
def filter.of_countable_Inter {α : Type u_2} (l : set (set α)) (hp : ∀ (S : set (set α)), S.countableS l⋂₀ S l) (h_mono : ∀ (s t : set α), s ls tt l) :

Construct a filter with countable intersection property. This constructor deduces filter.univ_sets and filter.inter_sets from the countable intersection property.

Equations
Instances for filter.of_countable_Inter
@[protected, instance]
def filter.countable_Inter_of_countable_Inter {α : Type u_2} (l : set (set α)) (hp : ∀ (S : set (set α)), S.countableS l⋂₀ S l) (h_mono : ∀ (s t : set α), s ls tt l) :
@[simp]
theorem filter.mem_of_countable_Inter {α : Type u_2} {l : set (set α)} (hp : ∀ (S : set (set α)), S.countableS l⋂₀ S l) (h_mono : ∀ (s t : set α), s ls tt l) {s : set α} :
@[protected, instance]
@[protected, instance]
@[protected, instance]
@[protected, instance]
def filter.comap.countable_Inter_filter {α : Type u_2} {β : Type u_3} (l : filter β) [countable_Inter_filter l] (f : α → β) :
@[protected, instance]
def filter.map.countable_Inter_filter {α : Type u_2} {β : Type u_3} (l : filter α) [countable_Inter_filter l] (f : α → β) :
@[protected, instance]
def countable_Inter_filter_inf {α : Type u_2} (l₁ l₂ : filter α) [countable_Inter_filter l₁] [countable_Inter_filter l₂] :

Infimum of two countable_Inter_filters is a countable_Inter_filter. This is useful, e.g., to automatically get an instance for residual α ⊓ 𝓟 s.

@[protected, instance]
def countable_Inter_filter_sup {α : Type u_2} (l₁ l₂ : filter α) [countable_Inter_filter l₁] [countable_Inter_filter l₂] :

Supremum of two countable_Inter_filters is a countable_Inter_filter.