Continuous linear equivalences #
Continuous semilinear / linear / star-linear equivalences between topological modules are denoted
by M ≃SL[σ] M₂
, M ≃L[R] M₂
and M ≃L⋆[R] M₂
.
Continuous linear equivalences between modules. We only put the type classes that are necessary
for the definition, although in applications M
and M₂
will be topological modules over the
topological semiring R
.
- toFun : M → M₂
- invFun : M₂ → M
- left_inv : Function.LeftInverse self.invFun (↑self.toLinearEquiv).toFun
- right_inv : Function.RightInverse self.invFun (↑self.toLinearEquiv).toFun
- continuous_toFun : Continuous (↑self.toLinearEquiv).toFun
Continuous linear equivalences between modules. We only put the type classes that are necessary for the definition, although in applications
M
andM₂
will be topological modules over the topological semiringR
. - continuous_invFun : Continuous self.invFun
Continuous linear equivalences between modules. We only put the type classes that are necessary for the definition, although in applications
M
andM₂
will be topological modules over the topological semiringR
.
Instances For
Continuous linear equivalences between modules. We only put the type classes that are necessary
for the definition, although in applications M
and M₂
will be topological modules over the
topological semiring R
.
Equations
- One or more equations did not get rendered due to their size.
Instances For
Continuous linear equivalences between modules. We only put the type classes that are necessary
for the definition, although in applications M
and M₂
will be topological modules over the
topological semiring R
.
Equations
- One or more equations did not get rendered due to their size.
Instances For
ContinuousSemilinearEquivClass F σ M M₂
asserts F
is a type of bundled continuous
σ
-semilinear equivs M → M₂
. See also ContinuousLinearEquivClass F R M M₂
for the case
where σ
is the identity map on R
. A map f
between an R
-module and an S
-module over a ring
homomorphism σ : R →+* S
is semilinear if it satisfies the two properties f (x + y) = f x + f y
and f (c • x) = (σ c) • f x
.
- map_continuous (f : F) : Continuous ⇑f
ContinuousSemilinearEquivClass F σ M M₂
assertsF
is a type of bundled continuousσ
-semilinear equivsM → M₂
. See alsoContinuousLinearEquivClass F R M M₂
for the case whereσ
is the identity map onR
. A mapf
between anR
-module and anS
-module over a ring homomorphismσ : R →+* S
is semilinear if it satisfies the two propertiesf (x + y) = f x + f y
andf (c • x) = (σ c) • f x
. - inv_continuous (f : F) : Continuous (EquivLike.inv f)
ContinuousSemilinearEquivClass F σ M M₂
assertsF
is a type of bundled continuousσ
-semilinear equivsM → M₂
. See alsoContinuousLinearEquivClass F R M M₂
for the case whereσ
is the identity map onR
. A mapf
between anR
-module and anS
-module over a ring homomorphismσ : R →+* S
is semilinear if it satisfies the two propertiesf (x + y) = f x + f y
andf (c • x) = (σ c) • f x
.
Instances
ContinuousLinearEquivClass F σ M M₂
asserts F
is a type of bundled continuous
R
-linear equivs M → M₂
. This is an abbreviation for
ContinuousSemilinearEquivClass F (RingHom.id R) M M₂
.
Equations
- ContinuousLinearEquivClass F R M M₂ = ContinuousSemilinearEquivClass F (RingHom.id R) M M₂
Instances For
If I
and J
are complementary index sets, the product of the kernels of the J
th projections
of φ
is linearly equivalent to the product over I
.
Equations
- ContinuousLinearMap.iInfKerProjEquiv R φ hd hu = { toLinearEquiv := LinearMap.iInfKerProjEquiv R φ hd hu, continuous_toFun := ⋯, continuous_invFun := ⋯ }
Instances For
A continuous linear equivalence induces a continuous linear map.
Equations
- ↑e = { toLinearMap := ↑e.toLinearEquiv, cont := ⋯ }
Instances For
Coerce continuous linear equivs to continuous linear maps.
Equations
- ContinuousLinearEquiv.ContinuousLinearMap.coe = { coe := ContinuousLinearEquiv.toContinuousLinearMap }
Equations
- One or more equations did not get rendered due to their size.
A continuous linear equivalence induces a homeomorphism.
Equations
- e.toHomeomorph = { toEquiv := e.toEquiv, continuous_toFun := ⋯, continuous_invFun := ⋯ }
Instances For
An extensionality lemma for R ≃L[R] M
.
The identity map as a continuous linear equivalence.
Equations
- ContinuousLinearEquiv.refl R₁ M₁ = { toLinearEquiv := LinearEquiv.refl R₁ M₁, continuous_toFun := ⋯, continuous_invFun := ⋯ }
Instances For
The inverse of a continuous linear equivalence as a continuous linear equivalence
Equations
- e.symm = { toLinearEquiv := e.symm, continuous_toFun := ⋯, continuous_invFun := ⋯ }
Instances For
See Note [custom simps projection]. We need to specify this projection explicitly in this case, because it is a composition of multiple projections.
Equations
Instances For
See Note [custom simps projection]
Equations
- ContinuousLinearEquiv.Simps.symm_apply h = ⇑h.symm
Instances For
The composition of two continuous linear equivalences as a continuous linear equivalence.
Equations
- e₁.trans e₂ = { toLinearEquiv := e₁.trans e₂.toLinearEquiv, continuous_toFun := ⋯, continuous_invFun := ⋯ }
Instances For
Product of two continuous linear equivalences. The map comes from Equiv.prodCongr
.
Equations
- e.prod e' = { toLinearEquiv := e.prod e'.toLinearEquiv, continuous_toFun := ⋯, continuous_invFun := ⋯ }
Instances For
Product of modules is commutative up to continuous linear isomorphism.
Equations
- ContinuousLinearEquiv.prodComm R₁ M₁ M₂ = { toLinearEquiv := LinearEquiv.prodComm R₁ M₁ M₂, continuous_toFun := ⋯, continuous_invFun := ⋯ }
Instances For
Alias of ContinuousLinearEquiv.isUniformEmbedding
.
Alias of LinearEquiv.isUniformEmbedding
.
Create a ContinuousLinearEquiv
from two ContinuousLinearMap
s that are
inverse of each other. See also equivOfInverse'
.
Equations
- ContinuousLinearEquiv.equivOfInverse f₁ f₂ h₁ h₂ = { toLinearMap := ↑f₁, invFun := ⇑f₂, left_inv := h₁, right_inv := h₂, continuous_toFun := ⋯, continuous_invFun := ⋯ }
Instances For
Create a ContinuousLinearEquiv
from two ContinuousLinearMap
s that are
inverse of each other, in the ContinuousLinearMap.comp
sense. See also equivOfInverse
.
Equations
- ContinuousLinearEquiv.equivOfInverse' f₁ f₂ h₁ h₂ = ContinuousLinearEquiv.equivOfInverse f₁ f₂ ⋯ ⋯
Instances For
The inverse of equivOfInverse'
is obtained by swapping the order of its parameters.
The continuous linear equivalences from M
to itself form a group under composition.
Equations
The continuous linear equivalence between ULift M₁
and M₁
.
This is a continuous version of ULift.moduleEquiv
.
Equations
- ContinuousLinearEquiv.ulift = { toLinearEquiv := ULift.moduleEquiv, continuous_toFun := ⋯, continuous_invFun := ⋯ }
Instances For
A pair of continuous (semi)linear equivalences generates an equivalence between the spaces of
continuous linear maps. See also ContinuousLinearEquiv.arrowCongr
.
Equations
- One or more equations did not get rendered due to their size.
Instances For
Combine a family of continuous linear equivalences into a continuous linear equivalence of pi-types.
Equations
- ContinuousLinearEquiv.piCongrRight f = { toLinearEquiv := LinearEquiv.piCongrRight fun (i : ι) => ↑(f i), continuous_toFun := ⋯, continuous_invFun := ⋯ }
Instances For
Equivalence given by a block lower diagonal matrix. e
and e'
are diagonal square blocks,
and f
is a rectangular block below the diagonal.
Equations
- e.skewProd e' f = { toLinearEquiv := e.skewProd e'.toLinearEquiv ↑f, continuous_toFun := ⋯, continuous_invFun := ⋯ }
Instances For
The negation map as a continuous linear equivalence.
Equations
- ContinuousLinearEquiv.neg R = { toLinearEquiv := LinearEquiv.neg R, continuous_toFun := ⋯, continuous_invFun := ⋯ }
Instances For
The next theorems cover the identification between M ≃L[R] M
and the group of units of the ring
M →L[R] M
.
An invertible continuous linear map f
determines a continuous equivalence from M
to itself.
Equations
- ContinuousLinearEquiv.ofUnit f = { toFun := ⇑↑f, map_add' := ⋯, map_smul' := ⋯, invFun := ⇑f.inv, left_inv := ⋯, right_inv := ⋯, continuous_toFun := ⋯, continuous_invFun := ⋯ }
Instances For
A continuous equivalence from M
to itself determines an invertible continuous linear map.
Equations
- f.toUnit = { val := ↑f, inv := ↑f.symm, val_inv := ⋯, inv_val := ⋯ }
Instances For
The units of the algebra of continuous R
-linear endomorphisms of M
is multiplicatively
equivalent to the type of continuous linear equivalences between M
and itself.
Equations
- ContinuousLinearEquiv.unitsEquiv R M = { toFun := ContinuousLinearEquiv.ofUnit, invFun := ContinuousLinearEquiv.toUnit, left_inv := ⋯, right_inv := ⋯, map_mul' := ⋯ }
Instances For
Continuous linear equivalences R ≃L[R] R
are enumerated by Rˣ
.
Equations
- One or more equations did not get rendered due to their size.
Instances For
A pair of continuous linear maps such that f₁ ∘ f₂ = id
generates a continuous
linear equivalence e
between M
and M₂ × f₁.ker
such that (e x).2 = x
for x ∈ f₁.ker
,
(e x).1 = f₁ x
, and (e (f₂ y)).2 = 0
. The map is given by e x = (f₁ x, x - f₂ (f₁ x))
.
Equations
- ContinuousLinearEquiv.equivOfRightInverse f₁ f₂ h = ContinuousLinearEquiv.equivOfInverse (f₁.prod (f₁.projKerOfRightInverse f₂ h)) (f₂.coprod (LinearMap.ker f₁).subtypeL) ⋯ ⋯
Instances For
If ι
has a unique element, then ι → M
is continuously linear equivalent to M
.
Equations
- ContinuousLinearEquiv.funUnique ι R M = { toLinearEquiv := LinearEquiv.funUnique ι R M, continuous_toFun := ⋯, continuous_invFun := ⋯ }
Instances For
Continuous linear equivalence between dependent functions (i : Fin 2) → M i
and M 0 × M 1
.
Equations
- ContinuousLinearEquiv.piFinTwo R M = { toLinearEquiv := LinearEquiv.piFinTwo R M, continuous_toFun := ⋯, continuous_invFun := ⋯ }
Instances For
Continuous linear equivalence between vectors in M² = Fin 2 → M
and M × M
.
Equations
- ContinuousLinearEquiv.finTwoArrow R M = { toLinearEquiv := LinearEquiv.finTwoArrow R M, continuous_toFun := ⋯, continuous_invFun := ⋯ }
Instances For
A continuous linear map is invertible if it is the forward direction of a continuous linear equivalence.
Instances For
Introduce a function inverse
from M →L[R] M₂
to M₂ →L[R] M
, which sends f
to f.symm
if
f
is a continuous linear equivalence and to 0
otherwise. This definition is somewhat ad hoc,
but one needs a fully (rather than partially) defined inverse function for some purposes, including
for calculus.
Equations
- f.inverse = if h : f.IsInvertible then ↑(Classical.choose h).symm else 0
Instances For
By definition, if f
is not invertible then inverse f = 0
.
Alias of ContinuousLinearMap.inverse_of_not_isInvertible
.
By definition, if f
is not invertible then inverse f = 0
.
The function ContinuousLinearEquiv.inverse
can be written in terms of Ring.inverse
for the
ring of self-maps of the domain.
If p
is a closed complemented submodule,
then there exists a submodule q
and a continuous linear equivalence M ≃L[R] (p × q)
such that
e (x : p) = (x, 0)
, e (y : q) = (0, y)
, and e.symm x = x.1 + x.2
.
In fact, the properties of e
imply the properties of e.symm
and vice versa,
but we provide both for convenience.
The function op
is a continuous linear equivalence.
Equations
- MulOpposite.opContinuousLinearEquiv R = { toLinearEquiv := MulOpposite.opLinearEquiv R, continuous_toFun := ⋯, continuous_invFun := ⋯ }