Documentation

Mathlib.Analysis.Normed.Field.UnitBall

Algebraic structures on unit balls and spheres #

In this file we define algebraic structures (Semigroup, CommSemigroup, Monoid, CommMonoid, Group, CommGroup) on Metric.ball (0 : ๐•œ) 1, Metric.closedBall (0 : ๐•œ) 1, and Metric.sphere (0 : ๐•œ) 1. In each case we use the weakest possible typeclass assumption on ๐•œ, from NonUnitalSeminormedRing to NormedField.

Algebraic structures on Metric.ball 0 1 #

def Subsemigroup.unitBall (๐•œ : Type u_2) [NonUnitalSeminormedRing ๐•œ] :
Subsemigroup ๐•œ

Unit ball in a non-unital seminormed ring as a bundled Subsemigroup.

Equations
instance Metric.unitBall.instContinuousMul {๐•œ : Type u_1} [NonUnitalSeminormedRing ๐•œ] :
ContinuousMul โ†‘(ball 0 1)
@[simp]
theorem Metric.unitBall.coe_mul {๐•œ : Type u_1} [NonUnitalSeminormedRing ๐•œ] (x y : โ†‘(ball 0 1)) :
โ†‘(x * y) = โ†‘x * โ†‘y
@[deprecated Metric.unitBall.coe_mul (since := "2025-04-18")]
theorem coe_mul_unitBall {๐•œ : Type u_1} [NonUnitalSeminormedRing ๐•œ] (x y : โ†‘(Metric.ball 0 1)) :
โ†‘(x * y) = โ†‘x * โ†‘y

Alias of Metric.unitBall.coe_mul.

instance Metric.unitBall.instZero {๐•œ : Type u_1} [Zero ๐•œ] [PseudoMetricSpace ๐•œ] :
Zero โ†‘(ball 0 1)
Equations
@[simp]
theorem Metric.unitBall.coe_zero {๐•œ : Type u_1} [Zero ๐•œ] [PseudoMetricSpace ๐•œ] :
โ†‘0 = 0
@[simp]
theorem Metric.unitBall.coe_eq_zero {๐•œ : Type u_1} [Zero ๐•œ] [PseudoMetricSpace ๐•œ] {a : โ†‘(ball 0 1)} :
โ†‘a = 0 โ†” a = 0
instance Metric.unitBall.instSemigroupWithZero {๐•œ : Type u_1} [NonUnitalSeminormedRing ๐•œ] :
SemigroupWithZero โ†‘(ball 0 1)
Equations
instance Metric.unitBall.instIsCancelMulZero {๐•œ : Type u_1} [NonUnitalSeminormedRing ๐•œ] [IsCancelMulZero ๐•œ] :
IsCancelMulZero โ†‘(ball 0 1)

Algebraic instances for Metric.closedBall 0 1 #

def Subsemigroup.unitClosedBall (๐•œ : Type u_2) [NonUnitalSeminormedRing ๐•œ] :
Subsemigroup ๐•œ

Closed unit ball in a non-unital seminormed ring as a bundled Subsemigroup.

Equations
@[simp]
theorem Metric.unitClosedBall.coe_mul {๐•œ : Type u_1} [NonUnitalSeminormedRing ๐•œ] (x y : โ†‘(closedBall 0 1)) :
โ†‘(x * y) = โ†‘x * โ†‘y
@[deprecated Metric.unitClosedBall.coe_mul (since := "2025-04-18")]
theorem coe_mul_unitClosedBall {๐•œ : Type u_1} [NonUnitalSeminormedRing ๐•œ] (x y : โ†‘(Metric.closedBall 0 1)) :
โ†‘(x * y) = โ†‘x * โ†‘y

Alias of Metric.unitClosedBall.coe_mul.

instance Metric.unitClosedBall.instZero {๐•œ : Type u_1} [Zero ๐•œ] [PseudoMetricSpace ๐•œ] :
Zero โ†‘(closedBall 0 1)
Equations
@[simp]
theorem Metric.unitClosedBall.coe_zero {๐•œ : Type u_1} [Zero ๐•œ] [PseudoMetricSpace ๐•œ] :
โ†‘0 = 0
@[simp]
theorem Metric.unitClosedBall.coe_eq_zero {๐•œ : Type u_1} [Zero ๐•œ] [PseudoMetricSpace ๐•œ] {a : โ†‘(closedBall 0 1)} :
โ†‘a = 0 โ†” a = 0
def Submonoid.unitClosedBall (๐•œ : Type u_2) [SeminormedRing ๐•œ] [NormOneClass ๐•œ] :
Submonoid ๐•œ

Closed unit ball in a seminormed ring as a bundled Submonoid.

Equations
@[simp]
theorem Metric.unitClosedBall.coe_one {๐•œ : Type u_1} [SeminormedRing ๐•œ] [NormOneClass ๐•œ] :
โ†‘1 = 1
@[deprecated Metric.unitClosedBall.coe_one (since := "2025-04-18")]
theorem coe_one_unitClosedBall {๐•œ : Type u_1} [SeminormedRing ๐•œ] [NormOneClass ๐•œ] :
โ†‘1 = 1

Alias of Metric.unitClosedBall.coe_one.

@[simp]
theorem Metric.unitClosedBall.coe_eq_one {๐•œ : Type u_1} [SeminormedRing ๐•œ] [NormOneClass ๐•œ] {a : โ†‘(closedBall 0 1)} :
โ†‘a = 1 โ†” a = 1
@[simp]
theorem Metric.unitClosedBall.coe_pow {๐•œ : Type u_1} [SeminormedRing ๐•œ] [NormOneClass ๐•œ] (x : โ†‘(closedBall 0 1)) (n : โ„•) :
โ†‘(x ^ n) = โ†‘x ^ n
@[deprecated Metric.unitClosedBall.coe_pow (since := "2025-04-18")]
theorem coe_pow_unitClosedBall {๐•œ : Type u_1} [SeminormedRing ๐•œ] [NormOneClass ๐•œ] (x : โ†‘(Metric.closedBall 0 1)) (n : โ„•) :
โ†‘(x ^ n) = โ†‘x ^ n

Alias of Metric.unitClosedBall.coe_pow.

instance Metric.unitClosedBall.instMonoidWithZero {๐•œ : Type u_1} [SeminormedRing ๐•œ] [NormOneClass ๐•œ] :
Equations
instance Metric.unitClosedBall.instCancelMonoidWithZero {๐•œ : Type u_1} [SeminormedRing ๐•œ] [IsCancelMulZero ๐•œ] [NormOneClass ๐•œ] :
Equations

Algebraic instances on the unit sphere #

def Submonoid.unitSphere (๐•œ : Type u_2) [SeminormedRing ๐•œ] [NormMulClass ๐•œ] [NormOneClass ๐•œ] :
Submonoid ๐•œ

Unit sphere in a seminormed ring (with strictly multiplicative norm) as a bundled Submonoid.

Equations
@[simp]
theorem Submonoid.coe_unitSphere (๐•œ : Type u_2) [SeminormedRing ๐•œ] [NormMulClass ๐•œ] [NormOneClass ๐•œ] :
โ†‘(unitSphere ๐•œ) = Metric.sphere 0 1
instance Metric.unitSphere.instInv {๐•œ : Type u_1} [NormedDivisionRing ๐•œ] :
Inv โ†‘(sphere 0 1)
Equations
@[simp]
theorem Metric.unitSphere.coe_inv {๐•œ : Type u_1} [NormedDivisionRing ๐•œ] (x : โ†‘(sphere 0 1)) :
โ†‘xโปยน = (โ†‘x)โปยน
@[deprecated Metric.unitSphere.coe_inv (since := "2025-04-18")]
theorem coe_inv_unitSphere {๐•œ : Type u_1} [NormedDivisionRing ๐•œ] (x : โ†‘(Metric.sphere 0 1)) :
โ†‘xโปยน = (โ†‘x)โปยน

Alias of Metric.unitSphere.coe_inv.

instance Metric.unitSphere.instDiv {๐•œ : Type u_1} [NormedDivisionRing ๐•œ] :
Div โ†‘(sphere 0 1)
Equations
@[simp]
theorem Metric.unitSphere.coe_div {๐•œ : Type u_1} [NormedDivisionRing ๐•œ] (x y : โ†‘(sphere 0 1)) :
โ†‘(x / y) = โ†‘x / โ†‘y
@[deprecated Metric.unitSphere.coe_div (since := "2025-04-18")]
theorem coe_div_unitSphere {๐•œ : Type u_1} [NormedDivisionRing ๐•œ] (x y : โ†‘(Metric.sphere 0 1)) :
โ†‘(x / y) = โ†‘x / โ†‘y

Alias of Metric.unitSphere.coe_div.

instance Metric.unitSphere.instZPow {๐•œ : Type u_1} [NormedDivisionRing ๐•œ] :
Pow โ†‘(sphere 0 1) โ„ค
Equations
@[simp]
theorem Metric.unitSphere.coe_zpow {๐•œ : Type u_1} [NormedDivisionRing ๐•œ] (x : โ†‘(sphere 0 1)) (n : โ„ค) :
โ†‘(x ^ n) = โ†‘x ^ n
@[deprecated Metric.unitSphere.coe_zpow (since := "2025-04-18")]
theorem coe_zpow_unitSphere {๐•œ : Type u_1} [NormedDivisionRing ๐•œ] (x : โ†‘(Metric.sphere 0 1)) (n : โ„ค) :
โ†‘(x ^ n) = โ†‘x ^ n

Alias of Metric.unitSphere.coe_zpow.

instance Metric.unitSphere.instMonoid {๐•œ : Type u_1} [SeminormedRing ๐•œ] [NormMulClass ๐•œ] [NormOneClass ๐•œ] :
Monoid โ†‘(sphere 0 1)
Equations
@[simp]
theorem Metric.unitSphere.coe_one {๐•œ : Type u_1} [SeminormedRing ๐•œ] [NormMulClass ๐•œ] [NormOneClass ๐•œ] :
โ†‘1 = 1
@[deprecated Metric.unitSphere.coe_one (since := "2025-04-18")]
theorem coe_one_unitSphere {๐•œ : Type u_1} [SeminormedRing ๐•œ] [NormMulClass ๐•œ] [NormOneClass ๐•œ] :
โ†‘1 = 1

Alias of Metric.unitSphere.coe_one.

@[simp]
theorem Metric.unitSphere.coe_mul {๐•œ : Type u_1} [SeminormedRing ๐•œ] [NormMulClass ๐•œ] [NormOneClass ๐•œ] (x y : โ†‘(sphere 0 1)) :
โ†‘(x * y) = โ†‘x * โ†‘y
@[deprecated Metric.unitSphere.coe_mul (since := "2025-04-18")]
theorem coe_mul_unitSphere {๐•œ : Type u_1} [SeminormedRing ๐•œ] [NormMulClass ๐•œ] [NormOneClass ๐•œ] (x y : โ†‘(Metric.sphere 0 1)) :
โ†‘(x * y) = โ†‘x * โ†‘y

Alias of Metric.unitSphere.coe_mul.

@[simp]
theorem Metric.unitSphere.coe_pow {๐•œ : Type u_1} [SeminormedRing ๐•œ] [NormMulClass ๐•œ] [NormOneClass ๐•œ] (x : โ†‘(sphere 0 1)) (n : โ„•) :
โ†‘(x ^ n) = โ†‘x ^ n
@[deprecated Metric.unitSphere.coe_pow (since := "2025-04-18")]
theorem coe_pow_unitSphere {๐•œ : Type u_1} [SeminormedRing ๐•œ] [NormMulClass ๐•œ] [NormOneClass ๐•œ] (x : โ†‘(Metric.sphere 0 1)) (n : โ„•) :
โ†‘(x ^ n) = โ†‘x ^ n

Alias of Metric.unitSphere.coe_pow.

def unitSphereToUnits (๐•œ : Type u_2) [NormedDivisionRing ๐•œ] :
โ†‘(Metric.sphere 0 1) โ†’* ๐•œหฃ

Monoid homomorphism from the unit sphere in a normed division ring to the group of units.

Equations
@[simp]
theorem unitSphereToUnits_apply_coe {๐•œ : Type u_1} [NormedDivisionRing ๐•œ] (x : โ†‘(Metric.sphere 0 1)) :
โ†‘((unitSphereToUnits ๐•œ) x) = โ†‘x
theorem unitSphereToUnits_injective {๐•œ : Type u_1} [NormedDivisionRing ๐•œ] :
instance Metric.unitSphere.instGroup {๐•œ : Type u_1} [NormedDivisionRing ๐•œ] :
Group โ†‘(sphere 0 1)
Equations
instance Metric.sphere.instHasDistribNeg {๐•œ : Type u_1} [SeminormedRing ๐•œ] [NormMulClass ๐•œ] [NormOneClass ๐•œ] :
HasDistribNeg โ†‘(sphere 0 1)
Equations
instance Metric.sphere.instContinuousMul {๐•œ : Type u_1} [SeminormedRing ๐•œ] [NormMulClass ๐•œ] [NormOneClass ๐•œ] :
ContinuousMul โ†‘(sphere 0 1)
instance Metric.sphere.instCommGroup {๐•œ : Type u_1} [NormedField ๐•œ] :
CommGroup โ†‘(sphere 0 1)
Equations