Skip to content

Commit

Permalink
chore(UniformSpace.Basic): make UniformSpace.comap reducible (#10010)
Browse files Browse the repository at this point in the history
`UniformSpace.comap` is used in instance construction so needs to be reducible for unification purposes.
  • Loading branch information
mattrobball committed Jan 26, 2024
1 parent 79a9e0e commit 66439f5
Showing 1 changed file with 6 additions and 4 deletions.
10 changes: 6 additions & 4 deletions Mathlib/Topology/UniformSpace/Basic.lean
Expand Up @@ -1258,7 +1258,9 @@ instance inhabitedUniformSpaceCore : Inhabited (UniformSpace.Core α) :=
#align inhabited_uniform_space_core inhabitedUniformSpaceCore

/-- Given `f : α → β` and a uniformity `u` on `β`, the inverse image of `u` under `f`
is the inverse image in the filter sense of the induced function `α × α → β × β`. -/
is the inverse image in the filter sense of the induced function `α × α → β × β`.
See note [reducible non-instances]. -/
@[reducible]
def UniformSpace.comap (f : α → β) (u : UniformSpace β) : UniformSpace α :=
.ofNhdsEqComap
{ uniformity := 𝓤[u].comap fun p : α × α => (f p.1, f p.2)
Expand Down Expand Up @@ -1570,10 +1572,10 @@ section Prod
instance instUniformSpaceProd [u₁ : UniformSpace α] [u₂ : UniformSpace β] : UniformSpace (α × β) :=
u₁.comap Prod.fst ⊓ u₂.comap Prod.snd

-- check the above produces no diamond
-- check the above produces no diamond for `simp` and typeclass search
example [UniformSpace α] [UniformSpace β] :
(instTopologicalSpaceProd : TopologicalSpace (α × β)) = UniformSpace.toTopologicalSpace :=
rfl
(instTopologicalSpaceProd : TopologicalSpace (α × β)) = UniformSpace.toTopologicalSpace := by
with_reducible_and_instances rfl

theorem uniformity_prod [UniformSpace α] [UniformSpace β] :
𝓤 (α × β) =
Expand Down

0 comments on commit 66439f5

Please sign in to comment.