Skip to content

Commit

Permalink
ToMathlib/LocallyFinite: use Type* and comment lemmas
Browse files Browse the repository at this point in the history
Three will stay for now (explain why); fourth has been PRed.
  • Loading branch information
grunweg committed Jan 20, 2024
1 parent 29442ec commit 7229a2c
Showing 1 changed file with 13 additions and 7 deletions.
20 changes: 13 additions & 7 deletions SphereEversion/ToMathlib/Topology/LocallyFinite.lean
Original file line number Diff line number Diff line change
Expand Up @@ -3,24 +3,29 @@ import Mathlib.Topology.LocallyFinite

open Function Set

theorem LocallyFinite.smul_left {ι : Type _} {α : Type _} [TopologicalSpace α] {M : Type _}
{R : Type _} [Zero R] [Zero M] [SMulWithZero R M]
-- these two lemmas require additional imports compared to what's in mathlib
-- `Algebra.SMulWithZero` is not required in `Topology.LocallyFinite` so far;
-- so it remains to find a nice home for these lemmas...
theorem LocallyFinite.smul_left {ι : Type*} {α : Type*} [TopologicalSpace α] {M : Type*}
{R : Type*} [Zero R] [Zero M] [SMulWithZero R M]
{s : ι → α → R} (h : LocallyFinite fun i => support <| s i) (f : ι → α → M) :
LocallyFinite fun i => support (s i • f i) :=
h.subset fun i x ↦ mt <| fun h ↦ by rw [Pi.smul_apply', h, zero_smul]

theorem LocallyFinite.smul_right {ι : Type _} {α : Type _} [TopologicalSpace α] {M : Type _}
{R : Type _} [Zero R] [Zero M] [SMulZeroClass R M]
theorem LocallyFinite.smul_right {ι : Type*} {α : Type*} [TopologicalSpace α] {M : Type*}
{R : Type*} [Zero R] [Zero M] [SMulZeroClass R M]
{f : ι → α → M} (h : LocallyFinite fun i => support <| f i) (s : ι → α → R) :
LocallyFinite fun i => support <| s i • f i :=
h.subset fun i x ↦ mt <| fun h ↦ by rw [Pi.smul_apply', h, smul_zero]

section

variable {ι X : Type _} [TopologicalSpace X]
variable {ι X : Type*} [TopologicalSpace X]

-- TODO: remove this; we don't want to have this lemma in mathlib.
-- See https://github.com/leanprover-community/mathlib4/pull/9813#discussion_r1455617707.
@[to_additive]
theorem LocallyFinite.exists_finset_mulSupport_eq {M : Type _} [CommMonoid M] {ρ : ι → X → M}
theorem LocallyFinite.exists_finset_mulSupport_eq {M : Type*} [CommMonoid M] {ρ : ι → X → M}
(hρ : LocallyFinite fun i => mulSupport <| ρ i) (x₀ : X) :
∃ I : Finset ι, (mulSupport fun i => ρ i x₀) = I := by
use (hρ.point_finite x₀).toFinset
Expand All @@ -33,7 +38,8 @@ open scoped Topology

open Filter

theorem LocallyFinite.eventually_subset {ι X : Type _} [TopologicalSpace X] {s : ι → Set X}
-- submitted as PR #9813
theorem LocallyFinite.eventually_subset {ι X : Type*} [TopologicalSpace X] {s : ι → Set X}
(hs : LocallyFinite s) (hs' : ∀ i, IsClosed (s i)) (x : X) :
∀ᶠ y in 𝓝 x, {i | y ∈ s i} ⊆ {i | x ∈ s i} := by
filter_upwards [hs.iInter_compl_mem_nhds hs' x] with y hy i hi
Expand Down

0 comments on commit 7229a2c

Please sign in to comment.