Propositional typeclasses on several ring homs #
This file contains three typeclasses used in the definition of (semi)linear maps:
- RingHomId σ, which expresses the fact that- σ₂₃ = id
- RingHomCompTriple σ₁₂ σ₂₃ σ₁₃, which expresses the fact that- σ₂₃.comp σ₁₂ = σ₁₃
- RingHomInvPair σ₁₂ σ₂₁, which states that- σ₁₂and- σ₂₁are inverses of each other
- RingHomSurjective σ, which states that- σis surjective These typeclasses ensure that objects such as- σ₂₃.comp σ₁₂never end up in the type of a semilinear map; instead, the typeclass system directly finds the appropriate- RingHomto use. A typical use-case is conjugate-linear maps, i.e. when- σ = Complex.conj; this system ensures that composing two conjugate-linear maps is a linear map, and not a- conj.comp conj-linear map.
Instances of these typeclasses mostly involving RingHom.id are also provided:
- RingHomInvPair (RingHom.id R) (RingHom.id R)
- [RingHomInvPair σ₁₂ σ₂₁] : RingHomCompTriple σ₁₂ σ₂₁ (RingHom.id R₁)
- RingHomCompTriple (RingHom.id R₁) σ₁₂ σ₁₂
- RingHomCompTriple σ₁₂ (RingHom.id R₂) σ₁₂
- RingHomSurjective (RingHom.id R)
- [RingHomInvPair σ₁ σ₂] : RingHomSurjective σ₁
Implementation notes #
- For the typeclass - RingHomInvPair σ₁₂ σ₂₁,- σ₂₁is marked as an- outParam, as it must typically be found via the typeclass inference system.
- Likewise, for - RingHomCompTriple σ₁₂ σ₂₃ σ₁₃,- σ₁₃is marked as an- outParam, for the same reason.
Tags #
Class that expresses the fact that three ring homomorphisms form a composition triple. This is used to handle composition of semilinear maps.
- comp_eq : σ₂₃.comp σ₁₂ = σ₁₃The morphisms form a commutative triangle 
Instances
Class that expresses the fact that two ring homomorphisms are inverses of each other. This is
used to handle symm for semilinear equivalences.
- comp_eq : RingHom.comp σ' σ = RingHom.id R₁σ'is a left inverse ofσ
- comp_eq₂ : σ.comp σ' = RingHom.id R₂σ'is a left inverse ofσ'
Instances
σ' is a left inverse of σ
σ' is a left inverse of σ'
Equations
- ⋯ = ⋯
Equations
- ⋯ = ⋯
Equations
- ⋯ = ⋯
Construct a RingHomInvPair from both directions of a ring equiv.
This is not an instance, as for equivalences that are involutions, a better instance
would be RingHomInvPair e e. Indeed, this declaration is not currently used in mathlib.
Swap the direction of a RingHomInvPair. This is not an instance as it would loop, and better
instances are often available and may often be preferrable to using this one. Indeed, this
declaration is not currently used in mathlib.
Equations
- ⋯ = ⋯
Equations
- ⋯ = ⋯
Class expressing the fact that a RingHom is surjective. This is needed in the context
of semilinear maps, where some lemmas require this.
- is_surjective : Function.Surjective ⇑σThe ring homomorphism is surjective 
Instances
The ring homomorphism is surjective
Equations
- ⋯ = ⋯
Equations
- ⋯ = ⋯
This cannot be an instance as there is no way to infer σ₁₂ and σ₂₃.