@@ -1903,7 +1903,7 @@ section Comap
1903
1903
equivalent conditions hold.
1904
1904
1905
1905
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'`.
1907
1907
3. The set `(m '' sᶜ)ᶜ` belongs to `f`, see `Filter.mem_comap_iff_compl` and
1908
1908
`Filter.compl_mem_comap`. -/
1909
1909
def comap (m : α → β) (f : Filter β) : Filter α
@@ -1922,6 +1922,11 @@ theorem mem_comap' : s ∈ comap f l ↔ { y | ∀ ⦃x⦄, f x = y → x ∈ s
1922
1922
fun h => ⟨_, h, fun x hx => hx rfl⟩⟩
1923
1923
#align filter.mem_comap' Filter.mem_comap'
1924
1924
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
+
1925
1930
/-- RHS form is used, e.g., in the definition of `UniformSpace`. -/
1926
1931
lemma mem_comap_prod_mk {x : α} {s : Set β} {F : Filter (α × β)} :
1927
1932
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
1939
1944
#align filter.frequently_comap Filter.frequently_comap
1940
1945
1941
1946
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 ]
1943
1948
#align filter.mem_comap_iff_compl Filter.mem_comap_iff_compl
1944
1949
1945
1950
theorem compl_mem_comap : sᶜ ∈ comap f l ↔ (f '' s)ᶜ ∈ l := by rw [mem_comap_iff_compl, compl_compl]
1946
1951
#align filter.compl_mem_comap Filter.compl_mem_comap
1947
1952
1948
1953
end Comap
1949
1954
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
+
1950
1994
/-- The monadic bind operation on filter is defined the usual way in terms of `map` and `join`.
1951
1995
1952
1996
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) :=
2169
2213
fun _ _ => map_le_iff_le_comap
2170
2214
#align filter.gc_map_comap Filter.gc_map_comap
2171
2215
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
+
2172
2226
@[mono]
2173
2227
theorem map_mono : Monotone (map m) :=
2174
2228
(gc_map_comap m).monotone_l
@@ -2242,15 +2296,7 @@ theorem disjoint_comap (h : Disjoint g₁ g₂) : Disjoint (comap m g₁) (comap
2242
2296
#align filter.disjoint_comap Filter.disjoint_comap
2243
2297
2244
2298
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
2254
2300
#align filter.comap_supr Filter.comap_iSup
2255
2301
2256
2302
theorem comap_sSup {s : Set (Filter β)} {m : α → β} : comap m (sSup s) = ⨆ f ∈ s, comap m f := by
0 commit comments