Ring structures on the multiplicative opposite #
Equations
Equations
Equations
Equations
- MulOpposite.instSemiring = Semiring.mk ⋯ ⋯ ⋯ ⋯ Monoid.npow ⋯ ⋯
Equations
Equations
Equations
Equations
- MulOpposite.instNonAssocRing = NonAssocRing.mk ⋯ ⋯ ⋯ ⋯ ⋯ ⋯
Equations
- MulOpposite.instRing = Ring.mk ⋯ SubNegMonoid.zsmul ⋯ ⋯ ⋯ ⋯ ⋯ ⋯
Equations
Equations
Equations
Equations
Equations
Equations
- AddOpposite.instSemiring = Semiring.mk ⋯ ⋯ ⋯ ⋯ Monoid.npow ⋯ ⋯
Equations
Equations
Equations
Equations
- AddOpposite.instNonAssocRing = NonAssocRing.mk ⋯ ⋯ ⋯ ⋯ ⋯ ⋯
Equations
- AddOpposite.instRing = Ring.mk ⋯ SubNegMonoid.zsmul ⋯ ⋯ ⋯ ⋯ ⋯ ⋯
Equations
A non-unital ring homomorphism f : R →ₙ+* S
such that f x
commutes with f y
for all x, y
defines a non-unital ring homomorphism to Sᵐᵒᵖ
.
Equations
- f.toOpposite hf = { toFun := MulOpposite.op ∘ ⇑f, map_mul' := ⋯, map_zero' := ⋯, map_add' := ⋯ }
Instances For
A non-unital ring homomorphism f : R →ₙ* S
such that f x
commutes with f y
for all x, y
defines a non-unital ring homomorphism from Rᵐᵒᵖ
.
Equations
- f.fromOpposite hf = { toFun := ⇑f ∘ MulOpposite.unop, map_mul' := ⋯, map_zero' := ⋯, map_add' := ⋯ }
Instances For
A non-unital ring hom α →ₙ+* β
can equivalently be viewed as a non-unital ring hom
αᵐᵒᵖ →+* βᵐᵒᵖ
. This is the action of the (fully faithful) ᵐᵒᵖ
-functor on morphisms.
Equations
- One or more equations did not get rendered due to their size.
Instances For
The 'unopposite' of a non-unital ring hom αᵐᵒᵖ →ₙ+* βᵐᵒᵖ
. Inverse to
NonUnitalRingHom.op
.
Equations
Instances For
A ring homomorphism f : R →+* S
such that f x
commutes with f y
for all x, y
defines
a ring homomorphism to Sᵐᵒᵖ
.
Equations
- f.toOpposite hf = { toFun := MulOpposite.op ∘ ⇑f, map_one' := ⋯, map_mul' := ⋯, map_zero' := ⋯, map_add' := ⋯ }
Instances For
A ring homomorphism f : R →+* S
such that f x
commutes with f y
for all x, y
defines
a ring homomorphism from Rᵐᵒᵖ
.
Equations
- f.fromOpposite hf = { toFun := ⇑f ∘ MulOpposite.unop, map_one' := ⋯, map_mul' := ⋯, map_zero' := ⋯, map_add' := ⋯ }
Instances For
A ring hom α →+* β
can equivalently be viewed as a ring hom αᵐᵒᵖ →+* βᵐᵒᵖ
. This is the
action of the (fully faithful) ᵐᵒᵖ
-functor on morphisms.
Equations
- One or more equations did not get rendered due to their size.
Instances For
The 'unopposite' of a ring hom αᵐᵒᵖ →+* βᵐᵒᵖ
. Inverse to RingHom.op
.
Equations
- RingHom.unop = RingHom.op.symm