(Semi)linear equivalences #
In this file we define
LinearEquiv σ M M₂
,M ≃ₛₗ[σ] M₂
: an invertible semilinear map. Here,σ
is aRingHom
fromR
toR₂
and ane : M ≃ₛₗ[σ] M₂
satisfiese (c • x) = (σ c) • (e x)
. The plain linear version, withσ
beingRingHom.id R
, is denoted byM ≃ₗ[R] M₂
, and the star-linear version (withσ
beingstarRingEnd
) is denoted byM ≃ₗ⋆[R] M₂
.
Implementation notes #
To ensure that composition works smoothly for semilinear equivalences, we use the typeclasses
RingHomCompTriple
, RingHomInvPair
and RingHomSurjective
from
Algebra/Ring/CompTypeclasses
.
The group structure on automorphisms, LinearEquiv.automorphismGroup
, is provided elsewhere.
TODO #
- Parts of this file have not yet been generalized to semilinear maps
Tags #
linear equiv, linear equivalences, linear isomorphism, linear isomorphic
A linear equivalence is an invertible linear map.
- toFun : M → M₂
- invFun : M₂ → M
- left_inv : Function.LeftInverse self.invFun (↑self).toFun
- right_inv : Function.RightInverse self.invFun (↑self).toFun
Instances For
M ≃ₛₗ[σ] M₂
denotes the type of linear equivalences between M
and M₂
over a
ring homomorphism σ
.
Equations
- One or more equations did not get rendered due to their size.
Instances For
M ≃ₗ [R] M₂
denotes the type of linear equivalences between M
and M₂
over
a plain linear map M →ₗ M₂
.
Equations
- One or more equations did not get rendered due to their size.
Instances For
SemilinearEquivClass F σ M M₂
asserts F
is a type of bundled σ
-semilinear equivs
M → M₂
.
See also LinearEquivClass 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
.
Applying a semilinear equivalence
f
overσ
tor • x
equalsσ r • f x
.
Instances
LinearEquivClass F R M M₂
asserts F
is a type of bundled R
-linear equivs M → M₂
.
This is an abbreviation for SemilinearEquivClass F (RingHom.id R) M M₂
.
Equations
- LinearEquivClass F R M M₂ = SemilinearEquivClass F (RingHom.id R) M M₂
Instances For
Reinterpret an element of a type of semilinear equivalences as a semilinear equivalence.
Equations
Instances For
Reinterpret an element of a type of semilinear equivalences as a semilinear equivalence.
Equations
- SemilinearEquivClass.instCoeToSemilinearEquiv = { coe := fun (f : F) => ↑f }
Equations
- LinearEquiv.instCoeLinearMap = { coe := LinearEquiv.toLinearMap }
The equivalence of types underlying a linear equivalence.
Equations
- f.toEquiv = f.toAddEquiv.toEquiv
Instances For
Equations
- LinearEquiv.instEquivLike = { coe := fun (e : M ≃ₛₗ[σ] M₂) => (↑e).toFun, inv := LinearEquiv.invFun, left_inv := ⋯, right_inv := ⋯, coe_injective' := ⋯ }
The identity map is a linear equivalence.
Equations
- LinearEquiv.refl R M = { toLinearMap := LinearMap.id, invFun := (Equiv.refl M).invFun, left_inv := ⋯, right_inv := ⋯ }
Instances For
Linear equivalences are symmetric.
Equations
Instances For
See Note [custom simps projection]
Equations
- LinearEquiv.Simps.apply e = ⇑e
Instances For
See Note [custom simps projection]
Equations
Instances For
Linear equivalences are transitive.
Equations
Instances For
e₁ ≪≫ₗ e₂
denotes the composition of the linear equivalences e₁
and e₂
.
Equations
- LinearEquiv.transNotation = Lean.ParserDescr.trailingNode `LinearEquiv.transNotation 80 80 (Lean.ParserDescr.binary `andthen (Lean.ParserDescr.symbol " ≪≫ₗ ") (Lean.ParserDescr.cat `term 81))
Instances For
Pretty printer defined by notation3
command.
Equations
- One or more equations did not get rendered due to their size.
Instances For
The two paths coercion can take to an AddMonoidHom
are equivalent
Auxiliary definition to avoid looping in dsimp
with LinearEquiv.symm_mk
.
Equations
- LinearEquiv.symm_mk.aux e f h₁ h₂ h₃ h₄ = { toFun := ⇑e, map_add' := h₁, map_smul' := h₂, invFun := f, left_inv := h₃, right_inv := h₄ }.symm
Instances For
Interpret a RingEquiv
f
as an f
-semilinear equiv.
Equations
- f.toSemilinearEquiv = { toFun := ⇑f, map_add' := ⋯, map_smul' := ⋯, invFun := f.invFun, left_inv := ⋯, right_inv := ⋯ }
Instances For
An involutive linear map is a linear equivalence.
Equations
- LinearEquiv.ofInvolutive f hf = { toLinearMap := f, invFun := (Function.Involutive.toPerm (⇑f) hf).invFun, left_inv := ⋯, right_inv := ⋯ }