Product of pseudometric spaces #
This file constructs the infinity distance on finite products of pseudometric spaces.
instance
pseudoMetricSpacePi
{β : Type u_2}
{π : β → Type u_3}
[Fintype β]
[(b : β) → PseudoMetricSpace (π b)]
:
PseudoMetricSpace ((b : β) → π b)
A finite product of pseudometric spaces is a pseudometric space, with the sup distance.
Equations
- pseudoMetricSpacePi = (PseudoEMetricSpace.toPseudoMetricSpaceOfDist (fun (f g : (b : β) → π b) => ↑(Finset.univ.sup fun (b : β) => nndist (f b) (g b))) ⋯ ⋯).replaceBornology ⋯
theorem
nndist_pi_def
{β : Type u_2}
{π : β → Type u_3}
[Fintype β]
[(b : β) → PseudoMetricSpace (π b)]
(f g : (b : β) → π b)
:
theorem
dist_pi_def
{β : Type u_2}
{π : β → Type u_3}
[Fintype β]
[(b : β) → PseudoMetricSpace (π b)]
(f g : (b : β) → π b)
:
theorem
dist_pi_const_le
{α : Type u_1}
{β : Type u_2}
[PseudoMetricSpace α]
[Fintype β]
(a b : α)
:
theorem
nndist_pi_const_le
{α : Type u_1}
{β : Type u_2}
[PseudoMetricSpace α]
[Fintype β]
(a b : α)
:
@[simp]
theorem
dist_pi_const
{α : Type u_1}
{β : Type u_2}
[PseudoMetricSpace α]
[Fintype β]
[Nonempty β]
(a b : α)
:
@[simp]
theorem
nndist_pi_const
{α : Type u_1}
{β : Type u_2}
[PseudoMetricSpace α]
[Fintype β]
[Nonempty β]
(a b : α)
:
theorem
nndist_le_pi_nndist
{β : Type u_2}
{π : β → Type u_3}
[Fintype β]
[(b : β) → PseudoMetricSpace (π b)]
(f g : (b : β) → π b)
(b : β)
:
theorem
dist_le_pi_dist
{β : Type u_2}
{π : β → Type u_3}
[Fintype β]
[(b : β) → PseudoMetricSpace (π b)]
(f g : (b : β) → π b)
(b : β)
:
theorem
closedBall_pi
{β : Type u_2}
{π : β → Type u_3}
[Fintype β]
[(b : β) → PseudoMetricSpace (π b)]
(x : (b : β) → π b)
{r : ℝ}
(hr : 0 ≤ r)
:
A closed ball in a product space is a product of closed balls. See also closedBall_pi'
for a version assuming Nonempty β
instead of 0 ≤ r
.
theorem
closedBall_pi'
{β : Type u_2}
{π : β → Type u_3}
[Fintype β]
[(b : β) → PseudoMetricSpace (π b)]
[Nonempty β]
(x : (b : β) → π b)
(r : ℝ)
:
A closed ball in a product space is a product of closed balls. See also closedBall_pi
for a version assuming 0 ≤ r
instead of Nonempty β
.