Skip to content

Commit adddab8

Browse files
committed
feat(Data/Set/Lattice, Order/Filter/Basic): more lemmas about kernImage and filter analog (#5744)
Co-authored-by: Junyan Xu <junyanxumath@gmail.com> @alreadydone This was originally discussed on Zulip, and Junyan made most of the work in [this message](https://leanprover.zulipchat.com/#narrow/stream/217875-Is-there-code-for-X.3F/topic/congruence.20of.20filters.20for.20tendsto/near/294976522). I just changed some proofs to use a bit more Galois connections. This is a bit of a gadget but it does simplify the proof of `comap_iSup`, and it will also be convenient to define the space of functions with support a compact subset of a fixed set. See [this message](https://leanprover.zulipchat.com/#narrow/stream/217875-Is-there-code-for-X.3F/topic/congruence.20of.20filters.20for.20tendsto/near/302238889) for more details.
1 parent 8bbef90 commit adddab8

File tree

3 files changed

+111
-22
lines changed

3 files changed

+111
-22
lines changed

Mathlib/Data/Set/Lattice.lean

Lines changed: 47 additions & 11 deletions
Original file line numberDiff line numberDiff line change
@@ -26,7 +26,7 @@ for `Set α`, and some more set constructions.
2626
* `Set.completeAtomicBooleanAlgebra`: `Set α` is a `CompleteAtomicBooleanAlgebra` with `≤ = ⊆`,
2727
`< = ⊂`, `⊓ = ∩`, `⊔ = ∪`, `⨅ = ⋂`, `⨆ = ⋃` and `\` as the set difference.
2828
See `Set.BooleanAlgebra`.
29-
* `Set.kern_image`: For a function `f : α → β`, `s.kern_image f` is the set of `y` such that
29+
* `Set.kernImage`: For a function `f : α → β`, `s.kernImage f` is the set of `y` such that
3030
`f ⁻¹ y ⊆ s`.
3131
* `Set.seq`: Union of the image of a set under a **seq**uence of functions. `seq s t` is the union
3232
of `f '' t` over all `f ∈ s`, where `t : Set α` and `s : Set (α → β)`.
@@ -182,6 +182,14 @@ instance Set.completeAtomicBooleanAlgebra : CompleteAtomicBooleanAlgebra (Set α
182182
sInf_le := fun s t t_in a h => h _ t_in
183183
iInf_iSup_eq := by intros; ext; simp [Classical.skolem] }
184184

185+
/-- `kernImage f s` is the set of `y` such that `f ⁻¹ y ⊆ s`. -/
186+
def kernImage (f : α → β) (s : Set α) : Set β :=
187+
{ y | ∀ ⦃x⦄, f x = y → x ∈ s }
188+
#align set.kern_image Set.kernImage
189+
190+
lemma subset_kernImage_iff {f : α → β} : s ⊆ kernImage f t ↔ f ⁻¹' s ⊆ t :=
191+
fun h _ hx ↦ h hx rfl,
192+
fun h _ hx y hy ↦ h (show f y ∈ s from hy.symm ▸ hx)⟩
185193
section GaloisConnection
186194

187195
variable {f : α → β}
@@ -190,20 +198,48 @@ protected theorem image_preimage : GaloisConnection (image f) (preimage f) := fu
190198
image_subset_iff
191199
#align set.image_preimage Set.image_preimage
192200

193-
/-- `kernImage f s` is the set of `y` such that `f ⁻¹ y ⊆ s`. -/
194-
def kernImage (f : α → β) (s : Set α) : Set β :=
195-
{ y | ∀ ⦃x⦄, f x = y → x ∈ s }
196-
#align set.kern_image Set.kernImage
197-
198-
protected theorem preimage_kernImage : GaloisConnection (preimage f) (kernImage f) := fun a _ =>
199-
fun h _ hx y hy =>
200-
have : f y ∈ a := hy.symm ▸ hx
201-
h this,
202-
fun h x (hx : f x ∈ a) => h hx rfl⟩
201+
protected theorem preimage_kernImage : GaloisConnection (preimage f) (kernImage f) := fun _ _ =>
202+
subset_kernImage_iff.symm
203203
#align set.preimage_kern_image Set.preimage_kernImage
204204

205205
end GaloisConnection
206206

207+
section kernImage
208+
209+
variable {f : α → β}
210+
211+
lemma kernImage_mono : Monotone (kernImage f) :=
212+
Set.preimage_kernImage.monotone_u
213+
214+
lemma kernImage_eq_compl {s : Set α} : kernImage f s = (f '' sᶜ)ᶜ :=
215+
Set.preimage_kernImage.u_unique (Set.image_preimage.compl)
216+
(fun t ↦ compl_compl (f ⁻¹' t) ▸ Set.preimage_compl)
217+
218+
lemma kernImage_compl {s : Set α} : kernImage f (sᶜ) = (f '' s)ᶜ := by
219+
rw [kernImage_eq_compl, compl_compl]
220+
221+
lemma kernImage_empty : kernImage f ∅ = (range f)ᶜ := by
222+
rw [kernImage_eq_compl, compl_empty, image_univ]
223+
224+
lemma kernImage_preimage_eq_iff {s : Set β} : kernImage f (f ⁻¹' s) = s ↔ (range f)ᶜ ⊆ s := by
225+
rw [kernImage_eq_compl, ← preimage_compl, compl_eq_comm, eq_comm, image_preimage_eq_iff,
226+
compl_subset_comm]
227+
228+
lemma compl_range_subset_kernImage {s : Set α} : (range f)ᶜ ⊆ kernImage f s := by
229+
rw [← kernImage_empty]
230+
exact kernImage_mono (empty_subset _)
231+
232+
lemma kernImage_union_preimage {s : Set α} {t : Set β} :
233+
kernImage f (s ∪ f ⁻¹' t) = kernImage f s ∪ t := by
234+
rw [kernImage_eq_compl, kernImage_eq_compl, compl_union, ← preimage_compl, image_inter_preimage,
235+
compl_inter, compl_compl]
236+
237+
lemma kernImage_preimage_union {s : Set α} {t : Set β} :
238+
kernImage f (f ⁻¹' t ∪ s) = t ∪ kernImage f s := by
239+
rw [union_comm, kernImage_union_preimage, union_comm]
240+
241+
end kernImage
242+
207243
/-! ### Union and intersection over an indexed family of sets -/
208244

209245

Mathlib/Order/Filter/Basic.lean

Lines changed: 57 additions & 11 deletions
Original file line numberDiff line numberDiff line change
@@ -1903,7 +1903,7 @@ section Comap
19031903
equivalent conditions hold.
19041904
19051905
1. There exists a set `t ∈ f` such that `m ⁻¹' t ⊆ s`. This is used as a definition.
1906-
2. The set `{y | ∀ x, m x = y → x ∈ s}` belongs to `f`, see `Filter.mem_comap'`.
1906+
2. The set `kernImage m s = {y | ∀ x, m x = y → x ∈ s}` belongs to `f`, see `Filter.mem_comap'`.
19071907
3. The set `(m '' sᶜ)ᶜ` belongs to `f`, see `Filter.mem_comap_iff_compl` and
19081908
`Filter.compl_mem_comap`. -/
19091909
def comap (m : α → β) (f : Filter β) : Filter α
@@ -1922,6 +1922,11 @@ theorem mem_comap' : s ∈ comap f l ↔ { y | ∀ ⦃x⦄, f x = y → x ∈ s
19221922
fun h => ⟨_, h, fun x hx => hx rfl⟩⟩
19231923
#align filter.mem_comap' Filter.mem_comap'
19241924

1925+
-- TODO: it would be nice to use `kernImage` much more to take advantage of common name and API,
1926+
-- and then this would become `mem_comap'`
1927+
theorem mem_comap'' : s ∈ comap f l ↔ kernImage f s ∈ l :=
1928+
mem_comap'
1929+
19251930
/-- RHS form is used, e.g., in the definition of `UniformSpace`. -/
19261931
lemma mem_comap_prod_mk {x : α} {s : Set β} {F : Filter (α × β)} :
19271932
s ∈ comap (Prod.mk x) F ↔ {p : α × β | p.fst = x → p.snd ∈ s} ∈ F :=
@@ -1939,14 +1944,53 @@ theorem frequently_comap : (∃ᶠ a in comap f l, p a) ↔ ∃ᶠ b in l, ∃ a
19391944
#align filter.frequently_comap Filter.frequently_comap
19401945

19411946
theorem mem_comap_iff_compl : s ∈ comap f l ↔ (f '' sᶜ)ᶜ ∈ l := by
1942-
simp only [mem_comap', compl_def, mem_image, mem_setOf_eq, not_exists, not_and', not_not]
1947+
simp only [mem_comap'', kernImage_eq_compl]
19431948
#align filter.mem_comap_iff_compl Filter.mem_comap_iff_compl
19441949

19451950
theorem compl_mem_comap : sᶜ ∈ comap f l ↔ (f '' s)ᶜ ∈ l := by rw [mem_comap_iff_compl, compl_compl]
19461951
#align filter.compl_mem_comap Filter.compl_mem_comap
19471952

19481953
end Comap
19491954

1955+
section KernMap
1956+
1957+
/-- The analog of `kernImage` for filters. A set `s` belongs to `Filter.kernMap m f` if either of
1958+
the following equivalent conditions hold.
1959+
1960+
1. There exists a set `t ∈ f` such that `s = kernImage m t`. This is used as a definition.
1961+
2. There exists a set `t` such that `tᶜ ∈ f` and `sᶜ = m '' t`, see `Filter.mem_kernMap_iff_compl`
1962+
and `Filter.compl_mem_kernMap`.
1963+
1964+
This definition because it gives a right adjoint to `Filter.comap`, and because it has a nice
1965+
interpretation when working with `co-` filters (`Filter.cocompact`, `Filter.cofinite`, ...).
1966+
For example, `kernMap m (cocompact α)` is the filter generated by the complements of the sets
1967+
`m '' K` where `K` is a compact subset of `α`. -/
1968+
def kernMap (m : α → β) (f : Filter α) : Filter β where
1969+
sets := (kernImage m) '' f.sets
1970+
univ_sets := ⟨univ, f.univ_sets, by simp [kernImage_eq_compl]⟩
1971+
sets_of_superset := by
1972+
rintro _ t ⟨s, hs, rfl⟩ hst
1973+
refine ⟨s ∪ m ⁻¹' t, mem_of_superset hs (subset_union_left s _), ?_⟩
1974+
rw [kernImage_union_preimage, union_eq_right_iff_subset.mpr hst]
1975+
inter_sets := by
1976+
rintro _ _ ⟨s₁, h₁, rfl⟩ ⟨s₂, h₂, rfl⟩
1977+
exact ⟨s₁ ∩ s₂, f.inter_sets h₁ h₂, Set.preimage_kernImage.u_inf⟩
1978+
1979+
variable {m : α → β} {f : Filter α}
1980+
1981+
theorem mem_kernMap {s : Set β} : s ∈ kernMap m f ↔ ∃ t ∈ f, kernImage m t = s :=
1982+
Iff.rfl
1983+
1984+
theorem mem_kernMap_iff_compl {s : Set β} : s ∈ kernMap m f ↔ ∃ t, tᶜ ∈ f ∧ m '' t = sᶜ := by
1985+
rw [mem_kernMap, compl_surjective.exists]
1986+
refine exists_congr (fun x ↦ and_congr_right fun _ ↦ ?_)
1987+
rw [kernImage_compl, compl_eq_comm, eq_comm]
1988+
1989+
theorem compl_mem_kernMap {s : Set β} : sᶜ ∈ kernMap m f ↔ ∃ t, tᶜ ∈ f ∧ m '' t = s := by
1990+
simp_rw [mem_kernMap_iff_compl, compl_compl]
1991+
1992+
end KernMap
1993+
19501994
/-- The monadic bind operation on filter is defined the usual way in terms of `map` and `join`.
19511995
19521996
Unfortunately, this `bind` does not result in the expected applicative. See `Filter.seq` for the
@@ -2169,6 +2213,16 @@ theorem gc_map_comap (m : α → β) : GaloisConnection (map m) (comap m) :=
21692213
fun _ _ => map_le_iff_le_comap
21702214
#align filter.gc_map_comap Filter.gc_map_comap
21712215

2216+
theorem comap_le_iff_le_kernMap : comap m g ≤ f ↔ g ≤ kernMap m f := by
2217+
simp [Filter.le_def, mem_comap'', mem_kernMap, -mem_comap]
2218+
2219+
theorem gc_comap_kernMap (m : α → β) : GaloisConnection (comap m) (kernMap m) :=
2220+
fun _ _ ↦ comap_le_iff_le_kernMap
2221+
2222+
theorem kernMap_principal {s : Set α} : kernMap m (𝓟 s) = 𝓟 (kernImage m s) := by
2223+
refine eq_of_forall_le_iff (fun g ↦ ?_)
2224+
rw [← comap_le_iff_le_kernMap, le_principal_iff, le_principal_iff, mem_comap'']
2225+
21722226
@[mono]
21732227
theorem map_mono : Monotone (map m) :=
21742228
(gc_map_comap m).monotone_l
@@ -2242,15 +2296,7 @@ theorem disjoint_comap (h : Disjoint g₁ g₂) : Disjoint (comap m g₁) (comap
22422296
#align filter.disjoint_comap Filter.disjoint_comap
22432297

22442298
theorem comap_iSup {ι} {f : ι → Filter β} {m : α → β} : comap m (iSup f) = ⨆ i, comap m (f i) :=
2245-
le_antisymm
2246-
(fun s hs =>
2247-
have : ∀ i, ∃ t, t ∈ f i ∧ m ⁻¹' t ⊆ s := by
2248-
simpa only [mem_comap, exists_prop, mem_iSup] using mem_iSup.1 hs
2249-
let ⟨t, ht⟩ := Classical.axiom_of_choice this
2250-
⟨⋃ i, t i, mem_iSup.2 fun i => (f i).sets_of_superset (ht i).1 (subset_iUnion _ _), by
2251-
rw [preimage_iUnion, iUnion_subset_iff]
2252-
exact fun i => (ht i).2⟩)
2253-
(iSup_le fun i => comap_mono <| le_iSup _ _)
2299+
(gc_comap_kernMap m).l_iSup
22542300
#align filter.comap_supr Filter.comap_iSup
22552301

22562302
theorem comap_sSup {s : Set (Filter β)} {m : α → β} : comap m (sSup s) = ⨆ f ∈ s, comap m f := by

Mathlib/Order/GaloisConnection.lean

Lines changed: 7 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -339,6 +339,13 @@ protected theorem dfun {ι : Type u} {α : ι → Type v} {β : ι → Type w} [
339339
forall_congr' fun i => gc i (a i) (b i)
340340
#align galois_connection.dfun GaloisConnection.dfun
341341

342+
protected theorem compl [BooleanAlgebra α] [BooleanAlgebra β] {l : α → β} {u : β → α}
343+
(gc : GaloisConnection l u) :
344+
GaloisConnection (compl ∘ u ∘ compl) (compl ∘ l ∘ compl) := by
345+
intro a b
346+
dsimp
347+
rw [le_compl_iff_le_compl, gc, compl_le_iff_compl_le]
348+
342349
end Constructions
343350

344351
theorem l_comm_of_u_comm {X : Type _} [Preorder X] {Y : Type _} [Preorder Y] {Z : Type _}

0 commit comments

Comments
 (0)