Documentation

Mathlib.Algebra.Group.Opposite

Group structures on the multiplicative and additive opposites #

Additive structures on αᵐᵒᵖ #

Equations
Equations
Equations
Equations

Multiplicative structures on αᵐᵒᵖ #

We also generate additive structures on αᵃᵒᵖ using to_additive

instance MulOpposite.instMonoid {α : Type u_1} [Monoid α] :
Equations
@[simp]
theorem MulOpposite.op_pow {α : Type u_1} [Monoid α] (x : α) (n : ) :
op (x ^ n) = op x ^ n
@[simp]
theorem MulOpposite.unop_pow {α : Type u_1} [Monoid α] (x : αᵐᵒᵖ) (n : ) :
unop (x ^ n) = unop x ^ n
@[simp]
theorem MulOpposite.op_zpow {α : Type u_1} [DivInvMonoid α] (x : α) (z : ) :
op (x ^ z) = op x ^ z
@[simp]
theorem MulOpposite.unop_zpow {α : Type u_1} [DivInvMonoid α] (x : αᵐᵒᵖ) (z : ) :
unop (x ^ z) = unop x ^ z
@[simp]
theorem MulOpposite.op_natCast {α : Type u_1} [NatCast α] (n : ) :
op n = n
@[simp]
theorem AddOpposite.op_natCast {α : Type u_1} [NatCast α] (n : ) :
op n = n
@[simp]
theorem MulOpposite.op_ofNat {α : Type u_1} [NatCast α] (n : ) [n.AtLeastTwo] :
@[simp]
theorem AddOpposite.op_ofNat {α : Type u_1} [NatCast α] (n : ) [n.AtLeastTwo] :
@[simp]
theorem MulOpposite.op_intCast {α : Type u_1} [IntCast α] (n : ) :
op n = n
@[simp]
theorem AddOpposite.op_intCast {α : Type u_1} [IntCast α] (n : ) :
op n = n
@[simp]
theorem MulOpposite.unop_natCast {α : Type u_1} [NatCast α] (n : ) :
unop n = n
@[simp]
theorem AddOpposite.unop_natCast {α : Type u_1} [NatCast α] (n : ) :
unop n = n
@[simp]
theorem MulOpposite.unop_ofNat {α : Type u_1} [NatCast α] (n : ) [n.AtLeastTwo] :
@[simp]
theorem AddOpposite.unop_ofNat {α : Type u_1} [NatCast α] (n : ) [n.AtLeastTwo] :
@[simp]
theorem MulOpposite.unop_intCast {α : Type u_1} [IntCast α] (n : ) :
unop n = n
@[simp]
theorem AddOpposite.unop_intCast {α : Type u_1} [IntCast α] (n : ) :
unop n = n
@[simp]
theorem MulOpposite.unop_div {α : Type u_1} [DivInvMonoid α] (x y : αᵐᵒᵖ) :
unop (x / y) = (unop y)⁻¹ * unop x
@[simp]
theorem AddOpposite.unop_sub {α : Type u_1} [SubNegMonoid α] (x y : αᵃᵒᵖ) :
unop (x - y) = -unop y + unop x
@[simp]
theorem MulOpposite.op_div {α : Type u_1} [DivInvMonoid α] (x y : α) :
op (x / y) = (op y)⁻¹ * op x
@[simp]
theorem AddOpposite.op_sub {α : Type u_1} [SubNegMonoid α] (x y : α) :
op (x - y) = -op y + op x
@[simp]
theorem MulOpposite.semiconjBy_op {α : Type u_1} [Mul α] {a x y : α} :
SemiconjBy (op a) (op y) (op x) SemiconjBy a x y
@[simp]
theorem AddOpposite.addSemiconjBy_op {α : Type u_1} [Add α] {a x y : α} :
@[simp]
theorem MulOpposite.semiconjBy_unop {α : Type u_1} [Mul α] {a x y : αᵐᵒᵖ} :
@[simp]
theorem AddOpposite.addSemiconjBy_unop {α : Type u_1} [Add α] {a x y : αᵃᵒᵖ} :
theorem SemiconjBy.op {α : Type u_1} [Mul α] {a x y : α} (h : SemiconjBy a x y) :
theorem AddSemiconjBy.op {α : Type u_1} [Add α] {a x y : α} (h : AddSemiconjBy a x y) :
theorem Commute.op {α : Type u_1} [Mul α] {x y : α} (h : Commute x y) :
theorem AddCommute.op {α : Type u_1} [Add α] {x y : α} (h : AddCommute x y) :
theorem Commute.unop {α : Type u_1} [Mul α] {x y : αᵐᵒᵖ} (h : Commute x y) :
theorem AddCommute.unop {α : Type u_1} [Add α] {x y : αᵃᵒᵖ} (h : AddCommute x y) :
@[simp]
theorem MulOpposite.commute_op {α : Type u_1} [Mul α] {x y : α} :
Commute (op x) (op y) Commute x y
@[simp]
theorem AddOpposite.addCommute_op {α : Type u_1} [Add α] {x y : α} :
@[simp]
theorem MulOpposite.commute_unop {α : Type u_1} [Mul α] {x y : αᵐᵒᵖ} :
@[simp]
theorem AddOpposite.addCommute_unop {α : Type u_1} [Add α] {x y : αᵃᵒᵖ} :
def MulOpposite.opAddEquiv {α : Type u_1} [Add α] :

The function MulOpposite.op is an additive equivalence.

Equations
@[simp]
theorem MulOpposite.opAddEquiv_apply {α : Type u_1} [Add α] :
@[simp]
theorem MulOpposite.opAddEquiv_symm_apply {α : Type u_1} [Add α] :
opAddEquiv.symm = unop
@[simp]

Multiplicative structures on αᵃᵒᵖ #

instance AddOpposite.pow {α : Type u_1} {β : Type u_2} [Pow α β] :
Equations
@[simp]
theorem AddOpposite.op_pow {α : Type u_1} {β : Type u_2} [Pow α β] (a : α) (b : β) :
op (a ^ b) = op a ^ b
@[simp]
theorem AddOpposite.unop_pow {α : Type u_1} {β : Type u_2} [Pow α β] (a : αᵃᵒᵖ) (b : β) :
unop (a ^ b) = unop a ^ b
def AddOpposite.opMulEquiv {α : Type u_1} [Mul α] :

The function AddOpposite.op is a multiplicative equivalence.

Equations
@[simp]
theorem AddOpposite.opMulEquiv_apply {α : Type u_1} [Mul α] :
@[simp]
theorem AddOpposite.opMulEquiv_symm_apply {α : Type u_1} [Mul α] :
opMulEquiv.symm = unop
@[simp]

Inversion on a group is a MulEquiv to the opposite group. When G is commutative, there is MulEquiv.inv.

Equations

Negation on an additive group is an AddEquiv to the opposite group. When G is commutative, there is AddEquiv.inv.

Equations
def MulHom.toOpposite {M : Type u_2} {N : Type u_3} [Mul M] [Mul N] (f : M →ₙ* N) (hf : ∀ (x y : M), Commute (f x) (f y)) :

A semigroup homomorphism f : M →ₙ* N such that f x commutes with f y for all x, y defines a semigroup homomorphism to Nᵐᵒᵖ.

Equations
def AddHom.toOpposite {M : Type u_2} {N : Type u_3} [Add M] [Add N] (f : M →ₙ+ N) (hf : ∀ (x y : M), AddCommute (f x) (f y)) :

An additive semigroup homomorphism f : AddHom M N such that f x additively commutes with f y for all x, y defines an additive semigroup homomorphism to Sᵃᵒᵖ.

Equations
@[simp]
theorem MulHom.toOpposite_apply {M : Type u_2} {N : Type u_3} [Mul M] [Mul N] (f : M →ₙ* N) (hf : ∀ (x y : M), Commute (f x) (f y)) :
(f.toOpposite hf) = MulOpposite.op f
@[simp]
theorem AddHom.toOpposite_apply {M : Type u_2} {N : Type u_3} [Add M] [Add N] (f : M →ₙ+ N) (hf : ∀ (x y : M), AddCommute (f x) (f y)) :
(f.toOpposite hf) = AddOpposite.op f
def MulHom.fromOpposite {M : Type u_2} {N : Type u_3} [Mul M] [Mul N] (f : M →ₙ* N) (hf : ∀ (x y : M), Commute (f x) (f y)) :

A semigroup homomorphism f : M →ₙ* N such that f x commutes with f y for all x, y defines a semigroup homomorphism from Mᵐᵒᵖ.

Equations
def AddHom.fromOpposite {M : Type u_2} {N : Type u_3} [Add M] [Add N] (f : M →ₙ+ N) (hf : ∀ (x y : M), AddCommute (f x) (f y)) :

An additive semigroup homomorphism f : AddHom M N such that f x additively commutes with f y for all x, y defines an additive semigroup homomorphism from Mᵃᵒᵖ.

Equations
@[simp]
theorem AddHom.fromOpposite_apply {M : Type u_2} {N : Type u_3} [Add M] [Add N] (f : M →ₙ+ N) (hf : ∀ (x y : M), AddCommute (f x) (f y)) :
(f.fromOpposite hf) = f AddOpposite.unop
@[simp]
theorem MulHom.fromOpposite_apply {M : Type u_2} {N : Type u_3} [Mul M] [Mul N] (f : M →ₙ* N) (hf : ∀ (x y : M), Commute (f x) (f y)) :
(f.fromOpposite hf) = f MulOpposite.unop
def MonoidHom.toOpposite {M : Type u_2} {N : Type u_3} [MulOneClass M] [MulOneClass N] (f : M →* N) (hf : ∀ (x y : M), Commute (f x) (f y)) :

A monoid homomorphism f : M →* N such that f x commutes with f y for all x, y defines a monoid homomorphism to Nᵐᵒᵖ.

Equations
def AddMonoidHom.toOpposite {M : Type u_2} {N : Type u_3} [AddZeroClass M] [AddZeroClass N] (f : M →+ N) (hf : ∀ (x y : M), AddCommute (f x) (f y)) :

An additive monoid homomorphism f : M →+ N such that f x additively commutes with f y for all x, y defines an additive monoid homomorphism to Sᵃᵒᵖ.

Equations
@[simp]
theorem AddMonoidHom.toOpposite_apply {M : Type u_2} {N : Type u_3} [AddZeroClass M] [AddZeroClass N] (f : M →+ N) (hf : ∀ (x y : M), AddCommute (f x) (f y)) :
(f.toOpposite hf) = AddOpposite.op f
@[simp]
theorem MonoidHom.toOpposite_apply {M : Type u_2} {N : Type u_3} [MulOneClass M] [MulOneClass N] (f : M →* N) (hf : ∀ (x y : M), Commute (f x) (f y)) :
(f.toOpposite hf) = MulOpposite.op f
def MonoidHom.fromOpposite {M : Type u_2} {N : Type u_3} [MulOneClass M] [MulOneClass N] (f : M →* N) (hf : ∀ (x y : M), Commute (f x) (f y)) :

A monoid homomorphism f : M →* N such that f x commutes with f y for all x, y defines a monoid homomorphism from Mᵐᵒᵖ.

Equations
def AddMonoidHom.fromOpposite {M : Type u_2} {N : Type u_3} [AddZeroClass M] [AddZeroClass N] (f : M →+ N) (hf : ∀ (x y : M), AddCommute (f x) (f y)) :

An additive monoid homomorphism f : M →+ N such that f x additively commutes with f y for all x, y defines an additive monoid homomorphism from Mᵃᵒᵖ.

Equations
@[simp]
theorem MonoidHom.fromOpposite_apply {M : Type u_2} {N : Type u_3} [MulOneClass M] [MulOneClass N] (f : M →* N) (hf : ∀ (x y : M), Commute (f x) (f y)) :
(f.fromOpposite hf) = f MulOpposite.unop
@[simp]
theorem AddMonoidHom.fromOpposite_apply {M : Type u_2} {N : Type u_3} [AddZeroClass M] [AddZeroClass N] (f : M →+ N) (hf : ∀ (x y : M), AddCommute (f x) (f y)) :
(f.fromOpposite hf) = f AddOpposite.unop
def MulHom.op {M : Type u_2} {N : Type u_3} [Mul M] [Mul N] :

A semigroup homomorphism M →ₙ* N can equivalently be viewed as a semigroup homomorphism Mᵐᵒᵖ →ₙ* Nᵐᵒᵖ. This is the action of the (fully faithful) ᵐᵒᵖ-functor on morphisms.

Equations
  • One or more equations did not get rendered due to their size.
def AddHom.op {M : Type u_2} {N : Type u_3} [Add M] [Add N] :

An additive semigroup homomorphism AddHom M N can equivalently be viewed as an additive semigroup homomorphism AddHom Mᵃᵒᵖ Nᵃᵒᵖ. This is the action of the (fully faithful)ᵃᵒᵖ-functor on morphisms.

Equations
  • One or more equations did not get rendered due to their size.
@[simp]
theorem MulHom.op_symm_apply_apply {M : Type u_2} {N : Type u_3} [Mul M] [Mul N] (f : Mᵐᵒᵖ →ₙ* Nᵐᵒᵖ) (a✝ : M) :
(op.symm f) a✝ = (MulOpposite.unop f MulOpposite.op) a✝
@[simp]
theorem AddHom.op_apply_apply {M : Type u_2} {N : Type u_3} [Add M] [Add N] (f : M →ₙ+ N) (a✝ : Mᵃᵒᵖ) :
(op f) a✝ = (AddOpposite.op f AddOpposite.unop) a✝
@[simp]
theorem AddHom.op_symm_apply_apply {M : Type u_2} {N : Type u_3} [Add M] [Add N] (f : Mᵃᵒᵖ →ₙ+ Nᵃᵒᵖ) (a✝ : M) :
(op.symm f) a✝ = (AddOpposite.unop f AddOpposite.op) a✝
@[simp]
theorem MulHom.op_apply_apply {M : Type u_2} {N : Type u_3} [Mul M] [Mul N] (f : M →ₙ* N) (a✝ : Mᵐᵒᵖ) :
(op f) a✝ = (MulOpposite.op f MulOpposite.unop) a✝
def MulHom.unop {M : Type u_2} {N : Type u_3} [Mul M] [Mul N] :

The 'unopposite' of a semigroup homomorphism Mᵐᵒᵖ →ₙ* Nᵐᵒᵖ. Inverse to MulHom.op.

Equations
def AddHom.unop {M : Type u_2} {N : Type u_3} [Add M] [Add N] :

The 'unopposite' of an additive semigroup homomorphism Mᵃᵒᵖ →ₙ+ Nᵃᵒᵖ. Inverse to AddHom.op.

Equations
def AddHom.mulOp {M : Type u_2} {N : Type u_3} [Add M] [Add N] :

An additive semigroup homomorphism AddHom M N can equivalently be viewed as an additive homomorphism AddHom Mᵐᵒᵖ Nᵐᵒᵖ. This is the action of the (fully faithful) ᵐᵒᵖ-functor on morphisms.

Equations
  • One or more equations did not get rendered due to their size.
@[simp]
theorem AddHom.mulOp_apply_apply {M : Type u_2} {N : Type u_3} [Add M] [Add N] (f : M →ₙ+ N) (a✝ : Mᵐᵒᵖ) :
(mulOp f) a✝ = (MulOpposite.op f MulOpposite.unop) a✝
@[simp]
theorem AddHom.mulOp_symm_apply_apply {M : Type u_2} {N : Type u_3} [Add M] [Add N] (f : Mᵐᵒᵖ →ₙ+ Nᵐᵒᵖ) (a✝ : M) :
(mulOp.symm f) a✝ = (MulOpposite.unop f MulOpposite.op) a✝
def AddHom.mulUnop {α : Type u_2} {β : Type u_3} [Add α] [Add β] :

The 'unopposite' of an additive semigroup hom αᵐᵒᵖ →+ βᵐᵒᵖ. Inverse to AddHom.mul_op.

Equations
def MonoidHom.op {M : Type u_2} {N : Type u_3} [MulOneClass M] [MulOneClass N] :

A monoid homomorphism M →* N can equivalently be viewed as a monoid homomorphism Mᵐᵒᵖ →* Nᵐᵒᵖ. This is the action of the (fully faithful) ᵐᵒᵖ-functor on morphisms.

Equations
  • One or more equations did not get rendered due to their size.
def AddMonoidHom.op {M : Type u_2} {N : Type u_3} [AddZeroClass M] [AddZeroClass N] :

An additive monoid homomorphism M →+ N can equivalently be viewed as an additive monoid homomorphism Mᵃᵒᵖ →+ Nᵃᵒᵖ. This is the action of the (fully faithful) ᵃᵒᵖ-functor on morphisms.

Equations
  • One or more equations did not get rendered due to their size.
@[simp]
theorem MonoidHom.op_symm_apply_apply {M : Type u_2} {N : Type u_3} [MulOneClass M] [MulOneClass N] (f : Mᵐᵒᵖ →* Nᵐᵒᵖ) (a✝ : M) :
(op.symm f) a✝ = (MulOpposite.unop f MulOpposite.op) a✝
@[simp]
theorem MonoidHom.op_apply_apply {M : Type u_2} {N : Type u_3} [MulOneClass M] [MulOneClass N] (f : M →* N) (a✝ : Mᵐᵒᵖ) :
(op f) a✝ = (MulOpposite.op f MulOpposite.unop) a✝
@[simp]
theorem AddMonoidHom.op_symm_apply_apply {M : Type u_2} {N : Type u_3} [AddZeroClass M] [AddZeroClass N] (f : Mᵃᵒᵖ →+ Nᵃᵒᵖ) (a✝ : M) :
(op.symm f) a✝ = (AddOpposite.unop f AddOpposite.op) a✝
@[simp]
theorem AddMonoidHom.op_apply_apply {M : Type u_2} {N : Type u_3} [AddZeroClass M] [AddZeroClass N] (f : M →+ N) (a✝ : Mᵃᵒᵖ) :
(op f) a✝ = (AddOpposite.op f AddOpposite.unop) a✝
def MonoidHom.unop {M : Type u_2} {N : Type u_3} [MulOneClass M] [MulOneClass N] :

The 'unopposite' of a monoid homomorphism Mᵐᵒᵖ →* Nᵐᵒᵖ. Inverse to MonoidHom.op.

Equations
def AddMonoidHom.unop {M : Type u_2} {N : Type u_3} [AddZeroClass M] [AddZeroClass N] :

The 'unopposite' of an additive monoid homomorphism Mᵃᵒᵖ →+ Nᵃᵒᵖ. Inverse to AddMonoidHom.op.

Equations

A monoid is isomorphic to the opposite of its opposite.

Equations

A additive monoid is isomorphic to the opposite of its opposite.

Equations
@[simp]
theorem AddEquiv.opOp_symm_apply (M : Type u_2) [Add M] (a✝ : Mᵃᵒᵖᵃᵒᵖ) :
@[simp]
theorem MulEquiv.opOp_symm_apply (M : Type u_2) [Mul M] (a✝ : Mᵐᵒᵖᵐᵒᵖ) :
@[simp]
theorem AddEquiv.opOp_apply (M : Type u_2) [Add M] (a✝ : M) :
@[simp]
theorem MulEquiv.opOp_apply (M : Type u_2) [Mul M] (a✝ : M) :

An additive homomorphism M →+ N can equivalently be viewed as an additive homomorphism Mᵐᵒᵖ →+ Nᵐᵒᵖ. This is the action of the (fully faithful) ᵐᵒᵖ-functor on morphisms.

Equations
  • One or more equations did not get rendered due to their size.
@[simp]
theorem AddMonoidHom.mulOp_apply_apply {M : Type u_2} {N : Type u_3} [AddZeroClass M] [AddZeroClass N] (f : M →+ N) (a✝ : Mᵐᵒᵖ) :
(mulOp f) a✝ = (MulOpposite.op f MulOpposite.unop) a✝
@[simp]
theorem AddMonoidHom.mulOp_symm_apply_apply {M : Type u_2} {N : Type u_3} [AddZeroClass M] [AddZeroClass N] (f : Mᵐᵒᵖ →+ Nᵐᵒᵖ) (a✝ : M) :
(mulOp.symm f) a✝ = (MulOpposite.unop f MulOpposite.op) a✝
def AddMonoidHom.mulUnop {α : Type u_2} {β : Type u_3} [AddZeroClass α] [AddZeroClass β] :

The 'unopposite' of an additive monoid hom αᵐᵒᵖ →+ βᵐᵒᵖ. Inverse to AddMonoidHom.mul_op.

Equations
def AddEquiv.mulOp {α : Type u_2} {β : Type u_3} [Add α] [Add β] :

An iso α ≃+ β can equivalently be viewed as an iso αᵐᵒᵖ ≃+ βᵐᵒᵖ.

Equations
  • One or more equations did not get rendered due to their size.
@[simp]
theorem AddEquiv.mulOp_apply {α : Type u_2} {β : Type u_3} [Add α] [Add β] (f : α ≃+ β) :
mulOp f = MulOpposite.opAddEquiv.symm.trans (f.trans MulOpposite.opAddEquiv)
@[simp]
theorem AddEquiv.mulOp_symm_apply {α : Type u_2} {β : Type u_3} [Add α] [Add β] (f : αᵐᵒᵖ ≃+ βᵐᵒᵖ) :
def AddEquiv.mulUnop {α : Type u_2} {β : Type u_3} [Add α] [Add β] :

The 'unopposite' of an iso αᵐᵒᵖ ≃+ βᵐᵒᵖ. Inverse to AddEquiv.mul_op.

Equations
def MulEquiv.op {α : Type u_2} {β : Type u_3} [Mul α] [Mul β] :

An iso α ≃* β can equivalently be viewed as an iso αᵐᵒᵖ ≃* βᵐᵒᵖ.

Equations
  • One or more equations did not get rendered due to their size.
def AddEquiv.op {α : Type u_2} {β : Type u_3} [Add α] [Add β] :

An iso α ≃+ β can equivalently be viewed as an iso αᵃᵒᵖ ≃+ βᵃᵒᵖ.

Equations
  • One or more equations did not get rendered due to their size.
@[simp]
theorem MulEquiv.op_apply_symm_apply {α : Type u_2} {β : Type u_3} [Mul α] [Mul β] (f : α ≃* β) (a✝ : βᵐᵒᵖ) :
(op f).symm a✝ = (MulOpposite.op f.symm MulOpposite.unop) a✝
@[simp]
theorem MulEquiv.op_symm_apply_apply {α : Type u_2} {β : Type u_3} [Mul α] [Mul β] (f : αᵐᵒᵖ ≃* βᵐᵒᵖ) (a✝ : α) :
(op.symm f) a✝ = (MulOpposite.unop f MulOpposite.op) a✝
@[simp]
theorem MulEquiv.op_symm_apply_symm_apply {α : Type u_2} {β : Type u_3} [Mul α] [Mul β] (f : αᵐᵒᵖ ≃* βᵐᵒᵖ) (a✝ : β) :
(op.symm f).symm a✝ = (MulOpposite.unop f.symm MulOpposite.op) a✝
@[simp]
theorem MulEquiv.op_apply_apply {α : Type u_2} {β : Type u_3} [Mul α] [Mul β] (f : α ≃* β) (a✝ : αᵐᵒᵖ) :
(op f) a✝ = (MulOpposite.op f MulOpposite.unop) a✝
@[simp]
theorem AddEquiv.op_apply_apply {α : Type u_2} {β : Type u_3} [Add α] [Add β] (f : α ≃+ β) (a✝ : αᵃᵒᵖ) :
(op f) a✝ = (AddOpposite.op f AddOpposite.unop) a✝
@[simp]
theorem AddEquiv.op_symm_apply_symm_apply {α : Type u_2} {β : Type u_3} [Add α] [Add β] (f : αᵃᵒᵖ ≃+ βᵃᵒᵖ) (a✝ : β) :
(op.symm f).symm a✝ = (AddOpposite.unop f.symm AddOpposite.op) a✝
@[simp]
theorem AddEquiv.op_symm_apply_apply {α : Type u_2} {β : Type u_3} [Add α] [Add β] (f : αᵃᵒᵖ ≃+ βᵃᵒᵖ) (a✝ : α) :
(op.symm f) a✝ = (AddOpposite.unop f AddOpposite.op) a✝
@[simp]
theorem AddEquiv.op_apply_symm_apply {α : Type u_2} {β : Type u_3} [Add α] [Add β] (f : α ≃+ β) (a✝ : βᵃᵒᵖ) :
(op f).symm a✝ = (AddOpposite.op f.symm AddOpposite.unop) a✝
def MulEquiv.unop {α : Type u_2} {β : Type u_3} [Mul α] [Mul β] :

The 'unopposite' of an iso αᵐᵒᵖ ≃* βᵐᵒᵖ. Inverse to MulEquiv.op.

Equations
def AddEquiv.unop {α : Type u_2} {β : Type u_3} [Add α] [Add β] :

The 'unopposite' of an iso αᵃᵒᵖ ≃+ βᵃᵒᵖ. Inverse to AddEquiv.op.

Equations
theorem AddMonoidHom.mul_op_ext {α : Type u_2} {β : Type u_3} [AddZeroClass α] [AddZeroClass β] (f g : αᵐᵒᵖ →+ β) (h : f.comp MulOpposite.opAddEquiv.toAddMonoidHom = g.comp MulOpposite.opAddEquiv.toAddMonoidHom) :
f = g

This ext lemma changes equalities on αᵐᵒᵖ →+ β to equalities on α →+ β. This is useful because there are often ext lemmas for specific αs that will apply to an equality of α →+ β such as Finsupp.addHom_ext'.