Documentation

Mathlib.Algebra.Ring.Equiv

(Semi)ring equivs #

In this file we define an extension of Equiv called RingEquiv, which is a datatype representing an isomorphism of Semirings, Rings, DivisionRings, or Fields.

Notations #

The extended equiv have coercions to functions, and the coercion is the canonical notation when treating the isomorphism as maps.

Implementation notes #

The fields for RingEquiv now avoid the unbundled isMulHom and isAddHom, as these are deprecated.

Definition of multiplication in the groups of automorphisms agrees with function composition, multiplication in Equiv.Perm, and multiplication in CategoryTheory.End, not with CategoryTheory.CategoryStruct.comp.

Tags #

Equiv, MulEquiv, AddEquiv, RingEquiv, MulAut, AddAut, RingAut

def NonUnitalRingHom.inverse {R : Type u_4} {S : Type u_5} [NonUnitalNonAssocSemiring R] [NonUnitalNonAssocSemiring S] (f : R →ₙ+* S) (g : SR) (h₁ : Function.LeftInverse g f) (h₂ : Function.RightInverse g f) :

makes a NonUnitalRingHom from the bijective inverse of a NonUnitalRingHom

Equations
  • f.inverse g h₁ h₂ = { toFun := g, map_mul' := , map_zero' := , map_add' := }
@[simp]
theorem NonUnitalRingHom.inverse_apply {R : Type u_4} {S : Type u_5} [NonUnitalNonAssocSemiring R] [NonUnitalNonAssocSemiring S] (f : R →ₙ+* S) (g : SR) (h₁ : Function.LeftInverse g f) (h₂ : Function.RightInverse g f) (a✝ : S) :
(f.inverse g h₁ h₂) a✝ = g a✝
def RingHom.inverse {R : Type u_4} {S : Type u_5} [NonAssocSemiring R] [NonAssocSemiring S] (f : R →+* S) (g : SR) (h₁ : Function.LeftInverse g f) (h₂ : Function.RightInverse g f) :
S →+* R

makes a RingHom from the bijective inverse of a RingHom

Equations
  • f.inverse g h₁ h₂ = { toFun := g, map_one' := , map_mul' := , map_zero' := , map_add' := }
@[simp]
theorem RingHom.inverse_apply {R : Type u_4} {S : Type u_5} [NonAssocSemiring R] [NonAssocSemiring S] (f : R →+* S) (g : SR) (h₁ : Function.LeftInverse g f) (h₂ : Function.RightInverse g f) (a✝ : S) :
(f.inverse g h₁ h₂) a✝ = g a✝
structure RingEquiv (R : Type u_7) (S : Type u_8) [Mul R] [Mul S] [Add R] [Add S] extends R S, R ≃* S, R ≃+ S :
Type (max u_7 u_8)

An equivalence between two (non-unital non-associative semi)rings that preserves the algebraic structure.

Instances For
class RingEquivClass (F : Type u_7) (R : Type u_8) (S : Type u_9) [Mul R] [Add R] [Mul S] [Add S] [EquivLike F R S] extends MulEquivClass F R S :

RingEquivClass F R S states that F is a type of ring structure preserving equivalences. You should extend this class when you extend RingEquiv.

  • map_mul (f : F) (a b : R) : f (a * b) = f a * f b
  • map_add (f : F) (a b : R) : f (a + b) = f a + f b

    By definition, a ring isomorphism preserves the additive structure.

Instances
@[instance 100]
instance RingEquivClass.toAddEquivClass {F : Type u_1} {R : Type u_4} {S : Type u_5} [EquivLike F R S] [Mul R] [Add R] [Mul S] [Add S] [h : RingEquivClass F R S] :
@[instance 100]
instance RingEquivClass.toRingHomClass {F : Type u_1} {R : Type u_4} {S : Type u_5} [EquivLike F R S] [NonAssocSemiring R] [NonAssocSemiring S] [h : RingEquivClass F R S] :
def RingEquivClass.toRingEquiv {F : Type u_1} {α : Type u_2} {β : Type u_3} [Mul α] [Add α] [Mul β] [Add β] [EquivLike F α β] [RingEquivClass F α β] (f : F) :
α ≃+* β

Turn an element of a type F satisfying RingEquivClass F α β into an actual RingEquiv. This is declared as the default coercion from F to α ≃+* β.

Equations
  • f = { toEquiv := (↑f).toEquiv, map_mul' := , map_add' := }
instance instCoeTCRingEquivOfRingEquivClass {F : Type u_1} {α : Type u_2} {β : Type u_3} [Mul α] [Add α] [Mul β] [Add β] [EquivLike F α β] [RingEquivClass F α β] :
CoeTC F (α ≃+* β)

Any type satisfying RingEquivClass can be cast into RingEquiv via RingEquivClass.toRingEquiv.

Equations
instance RingEquiv.instEquivLike {R : Type u_4} {S : Type u_5} [Mul R] [Mul S] [Add R] [Add S] :
EquivLike (R ≃+* S) R S
Equations
instance RingEquiv.instRingEquivClass {R : Type u_4} {S : Type u_5} [Mul R] [Mul S] [Add R] [Add S] :
theorem RingEquiv.ext {R : Type u_4} {S : Type u_5} [Mul R] [Mul S] [Add R] [Add S] {f g : R ≃+* S} (h : ∀ (x : R), f x = g x) :
f = g

Two ring isomorphisms agree if they are defined by the same underlying function.

theorem RingEquiv.ext_iff {R : Type u_4} {S : Type u_5} [Mul R] [Mul S] [Add R] [Add S] {f g : R ≃+* S} :
f = g ∀ (x : R), f x = g x
theorem RingEquiv.congr_arg {R : Type u_4} {S : Type u_5} [Mul R] [Mul S] [Add R] [Add S] {f : R ≃+* S} {x x' : R} :
x = x'f x = f x'
theorem RingEquiv.congr_fun {R : Type u_4} {S : Type u_5} [Mul R] [Mul S] [Add R] [Add S] {f g : R ≃+* S} (h : f = g) (x : R) :
f x = g x
@[simp]
theorem RingEquiv.coe_mk {R : Type u_4} {S : Type u_5} [Mul R] [Mul S] [Add R] [Add S] (e : R S) (h₃ : ∀ (x y : R), e.toFun (x * y) = e.toFun x * e.toFun y) (h₄ : ∀ (x y : R), e.toFun (x + y) = e.toFun x + e.toFun y) :
{ toEquiv := e, map_mul' := h₃, map_add' := h₄ } = e
@[simp]
theorem RingEquiv.mk_coe {R : Type u_4} {S : Type u_5} [Mul R] [Mul S] [Add R] [Add S] (e : R ≃+* S) (e' : SR) (h₁ : Function.LeftInverse e' e) (h₂ : Function.RightInverse e' e) (h₃ : ∀ (x y : R), { toFun := e, invFun := e', left_inv := h₁, right_inv := h₂ }.toFun (x * y) = { toFun := e, invFun := e', left_inv := h₁, right_inv := h₂ }.toFun x * { toFun := e, invFun := e', left_inv := h₁, right_inv := h₂ }.toFun y) (h₄ : ∀ (x y : R), { toFun := e, invFun := e', left_inv := h₁, right_inv := h₂ }.toFun (x + y) = { toFun := e, invFun := e', left_inv := h₁, right_inv := h₂ }.toFun x + { toFun := e, invFun := e', left_inv := h₁, right_inv := h₂ }.toFun y) :
{ toFun := e, invFun := e', left_inv := h₁, right_inv := h₂, map_mul' := h₃, map_add' := h₄ } = e
@[simp]
theorem RingEquiv.toEquiv_eq_coe {R : Type u_4} {S : Type u_5} [Mul R] [Mul S] [Add R] [Add S] (f : R ≃+* S) :
f.toEquiv = f
@[simp]
theorem RingEquiv.coe_toEquiv {R : Type u_4} {S : Type u_5} [Mul R] [Mul S] [Add R] [Add S] (f : R ≃+* S) :
f = f
@[simp]
theorem RingEquiv.toAddEquiv_eq_coe {R : Type u_4} {S : Type u_5} [Mul R] [Mul S] [Add R] [Add S] (f : R ≃+* S) :
f.toAddEquiv = f
@[simp]
theorem RingEquiv.toMulEquiv_eq_coe {R : Type u_4} {S : Type u_5} [Mul R] [Mul S] [Add R] [Add S] (f : R ≃+* S) :
f.toMulEquiv = f
@[simp]
theorem RingEquiv.coe_toMulEquiv {R : Type u_4} {S : Type u_5} [Mul R] [Mul S] [Add R] [Add S] (f : R ≃+* S) :
f = f
@[simp]
theorem RingEquiv.coe_toAddEquiv {R : Type u_4} {S : Type u_5} [Mul R] [Mul S] [Add R] [Add S] (f : R ≃+* S) :
f = f
theorem RingEquiv.map_mul {R : Type u_4} {S : Type u_5} [Mul R] [Mul S] [Add R] [Add S] (e : R ≃+* S) (x y : R) :
e (x * y) = e x * e y

A ring isomorphism preserves multiplication.

theorem RingEquiv.map_add {R : Type u_4} {S : Type u_5} [Mul R] [Mul S] [Add R] [Add S] (e : R ≃+* S) (x y : R) :
e (x + y) = e x + e y

A ring isomorphism preserves addition.

theorem RingEquiv.bijective {R : Type u_4} {S : Type u_5} [Mul R] [Mul S] [Add R] [Add S] (e : R ≃+* S) :
theorem RingEquiv.injective {R : Type u_4} {S : Type u_5} [Mul R] [Mul S] [Add R] [Add S] (e : R ≃+* S) :
theorem RingEquiv.surjective {R : Type u_4} {S : Type u_5} [Mul R] [Mul S] [Add R] [Add S] (e : R ≃+* S) :
def RingEquiv.refl (R : Type u_4) [Mul R] [Add R] :
R ≃+* R

The identity map is a ring isomorphism.

Equations
instance RingEquiv.instInhabited (R : Type u_4) [Mul R] [Add R] :
Equations
@[simp]
theorem RingEquiv.refl_apply (R : Type u_4) [Mul R] [Add R] (x : R) :
(refl R) x = x
@[simp]
theorem RingEquiv.coe_refl (R : Type u_7) [Mul R] [Add R] :
(refl R) = id
@[deprecated RingEquiv.coe_refl (since := "2025-02-10")]
theorem RingEquiv.coe_refl_id (R : Type u_7) [Mul R] [Add R] :
(refl R) = id

Alias of RingEquiv.coe_refl.

@[simp]
theorem RingEquiv.coe_addEquiv_refl (R : Type u_4) [Mul R] [Add R] :
@[simp]
theorem RingEquiv.coe_mulEquiv_refl (R : Type u_4) [Mul R] [Add R] :
def RingEquiv.symm {R : Type u_4} {S : Type u_5} [Mul R] [Mul S] [Add R] [Add S] (e : R ≃+* S) :
S ≃+* R

The inverse of a ring isomorphism is a ring isomorphism.

Equations
@[simp]
theorem RingEquiv.invFun_eq_symm {R : Type u_4} {S : Type u_5} [Mul R] [Mul S] [Add R] [Add S] (f : R ≃+* S) :
@[simp]
theorem RingEquiv.symm_symm {R : Type u_4} {S : Type u_5} [Mul R] [Mul S] [Add R] [Add S] (e : R ≃+* S) :
e.symm.symm = e
@[simp]
theorem RingEquiv.mk_coe' {R : Type u_4} {S : Type u_5} [Mul R] [Mul S] [Add R] [Add S] (e : R ≃+* S) (f : SR) (h₁ : Function.LeftInverse (⇑e) f) (h₂ : Function.RightInverse (⇑e) f) (h₃ : ∀ (x y : S), { toFun := f, invFun := e, left_inv := h₁, right_inv := h₂ }.toFun (x * y) = { toFun := f, invFun := e, left_inv := h₁, right_inv := h₂ }.toFun x * { toFun := f, invFun := e, left_inv := h₁, right_inv := h₂ }.toFun y) (h₄ : ∀ (x y : S), { toFun := f, invFun := e, left_inv := h₁, right_inv := h₂ }.toFun (x + y) = { toFun := f, invFun := e, left_inv := h₁, right_inv := h₂ }.toFun x + { toFun := f, invFun := e, left_inv := h₁, right_inv := h₂ }.toFun y) :
{ toFun := f, invFun := e, left_inv := h₁, right_inv := h₂, map_mul' := h₃, map_add' := h₄ } = e.symm
def RingEquiv.symm_mk.aux {R : Type u_4} {S : Type u_5} [Mul R] [Mul S] [Add R] [Add S] (f : RS) (g : SR) (h₁ : Function.LeftInverse g f) (h₂ : Function.RightInverse g f) (h₃ : ∀ (x y : R), { toFun := f, invFun := g, left_inv := h₁, right_inv := h₂ }.toFun (x * y) = { toFun := f, invFun := g, left_inv := h₁, right_inv := h₂ }.toFun x * { toFun := f, invFun := g, left_inv := h₁, right_inv := h₂ }.toFun y) (h₄ : ∀ (x y : R), { toFun := f, invFun := g, left_inv := h₁, right_inv := h₂ }.toFun (x + y) = { toFun := f, invFun := g, left_inv := h₁, right_inv := h₂ }.toFun x + { toFun := f, invFun := g, left_inv := h₁, right_inv := h₂ }.toFun y) :
S ≃+* R

Auxiliary definition to avoid looping in dsimp with RingEquiv.symm_mk.

Equations
  • RingEquiv.symm_mk.aux f g h₁ h₂ h₃ h₄ = { toFun := f, invFun := g, left_inv := h₁, right_inv := h₂, map_mul' := h₃, map_add' := h₄ }.symm
@[simp]
theorem RingEquiv.symm_mk {R : Type u_4} {S : Type u_5} [Mul R] [Mul S] [Add R] [Add S] (f : RS) (g : SR) (h₁ : Function.LeftInverse g f) (h₂ : Function.RightInverse g f) (h₃ : ∀ (x y : R), { toFun := f, invFun := g, left_inv := h₁, right_inv := h₂ }.toFun (x * y) = { toFun := f, invFun := g, left_inv := h₁, right_inv := h₂ }.toFun x * { toFun := f, invFun := g, left_inv := h₁, right_inv := h₂ }.toFun y) (h₄ : ∀ (x y : R), { toFun := f, invFun := g, left_inv := h₁, right_inv := h₂ }.toFun (x + y) = { toFun := f, invFun := g, left_inv := h₁, right_inv := h₂ }.toFun x + { toFun := f, invFun := g, left_inv := h₁, right_inv := h₂ }.toFun y) :
{ toFun := f, invFun := g, left_inv := h₁, right_inv := h₂, map_mul' := h₃, map_add' := h₄ }.symm = let __src := symm_mk.aux f g h₁ h₂ h₃ h₄; { toFun := g, invFun := f, left_inv := , right_inv := , map_mul' := , map_add' := }
@[simp]
theorem RingEquiv.symm_refl {R : Type u_4} [Mul R] [Add R] :
(refl R).symm = refl R
@[simp]
theorem RingEquiv.coe_toEquiv_symm {R : Type u_4} {S : Type u_5} [Mul R] [Mul S] [Add R] [Add S] (e : R ≃+* S) :
e.symm = (↑e).symm
@[simp]
theorem RingEquiv.coe_toMulEquiv_symm {R : Type u_4} {S : Type u_5} [Mul R] [Mul S] [Add R] [Add S] (e : R ≃+* S) :
e.symm = (↑e).symm
@[simp]
theorem RingEquiv.coe_toAddEquiv_symm {R : Type u_4} {S : Type u_5} [Mul R] [Mul S] [Add R] [Add S] (e : R ≃+* S) :
e.symm = (↑e).symm
@[simp]
theorem RingEquiv.apply_symm_apply {R : Type u_4} {S : Type u_5} [Mul R] [Mul S] [Add R] [Add S] (e : R ≃+* S) (x : S) :
e (e.symm x) = x
@[simp]
theorem RingEquiv.symm_apply_apply {R : Type u_4} {S : Type u_5} [Mul R] [Mul S] [Add R] [Add S] (e : R ≃+* S) (x : R) :
e.symm (e x) = x
theorem RingEquiv.image_eq_preimage {R : Type u_4} {S : Type u_5} [Mul R] [Mul S] [Add R] [Add S] (e : R ≃+* S) (s : Set R) :
e '' s = e.symm ⁻¹' s
def RingEquiv.Simps.symm_apply {R : Type u_4} {S : Type u_5} [Mul R] [Mul S] [Add R] [Add S] (e : R ≃+* S) :
SR

See Note [custom simps projection]

Equations
def RingEquiv.trans {R : Type u_4} {S : Type u_5} {S' : Type u_6} [Mul R] [Mul S] [Add R] [Add S] [Mul S'] [Add S'] (e₁ : R ≃+* S) (e₂ : S ≃+* S') :
R ≃+* S'

Transitivity of RingEquiv.

Equations
@[simp]
theorem RingEquiv.coe_trans {R : Type u_4} {S : Type u_5} {S' : Type u_6} [Mul R] [Mul S] [Add R] [Add S] [Mul S'] [Add S'] (e₁ : R ≃+* S) (e₂ : S ≃+* S') :
(e₁.trans e₂) = e₂ e₁
theorem RingEquiv.trans_apply {R : Type u_4} {S : Type u_5} {S' : Type u_6} [Mul R] [Mul S] [Add R] [Add S] [Mul S'] [Add S'] (e₁ : R ≃+* S) (e₂ : S ≃+* S') (a : R) :
(e₁.trans e₂) a = e₂ (e₁ a)
@[simp]
theorem RingEquiv.symm_trans_apply {R : Type u_4} {S : Type u_5} {S' : Type u_6} [Mul R] [Mul S] [Add R] [Add S] [Mul S'] [Add S'] (e₁ : R ≃+* S) (e₂ : S ≃+* S') (a : S') :
(e₁.trans e₂).symm a = e₁.symm (e₂.symm a)
theorem RingEquiv.symm_trans {R : Type u_4} {S : Type u_5} {S' : Type u_6} [Mul R] [Mul S] [Add R] [Add S] [Mul S'] [Add S'] (e₁ : R ≃+* S) (e₂ : S ≃+* S') :
(e₁.trans e₂).symm = e₂.symm.trans e₁.symm
@[simp]
theorem RingEquiv.coe_mulEquiv_trans {R : Type u_4} {S : Type u_5} {S' : Type u_6} [Mul R] [Mul S] [Add R] [Add S] [Mul S'] [Add S'] (e₁ : R ≃+* S) (e₂ : S ≃+* S') :
(e₁.trans e₂) = (↑e₁).trans e₂
@[simp]
theorem RingEquiv.coe_addEquiv_trans {R : Type u_4} {S : Type u_5} {S' : Type u_6} [Mul R] [Mul S] [Add R] [Add S] [Mul S'] [Add S'] (e₁ : R ≃+* S) (e₂ : S ≃+* S') :
(e₁.trans e₂) = (↑e₁).trans e₂
def RingEquiv.ofUnique {M : Type u_7} {N : Type u_8} [Unique M] [Unique N] [Add M] [Mul M] [Add N] [Mul N] :
M ≃+* N

The RingEquiv between two semirings with a unique element.

Equations
@[deprecated RingEquiv.ofUnique (since := "2024-12-26")]
def RingEquiv.ringEquivOfUnique {M : Type u_7} {N : Type u_8} [Unique M] [Unique N] [Add M] [Mul M] [Add N] [Mul N] :
M ≃+* N

Alias of RingEquiv.ofUnique.


The RingEquiv between two semirings with a unique element.

Equations
instance RingEquiv.instUnique {M : Type u_7} {N : Type u_8} [Unique M] [Unique N] [Add M] [Mul M] [Add N] [Mul N] :
Equations
def RingEquiv.op {α : Type u_7} {β : Type u_8} [Add α] [Mul α] [Add β] [Mul β] :

A ring iso α ≃+* β can equivalently be viewed as a ring iso αᵐᵒᵖ ≃+* βᵐᵒᵖ.

Equations
  • One or more equations did not get rendered due to their size.
@[simp]
theorem RingEquiv.op_symm_apply_symm_apply {α : Type u_7} {β : Type u_8} [Add α] [Mul α] [Add β] [Mul β] (f : αᵐᵒᵖ ≃+* βᵐᵒᵖ) (a✝ : β) :
@[simp]
theorem RingEquiv.op_symm_apply_apply {α : Type u_7} {β : Type u_8} [Add α] [Mul α] [Add β] [Mul β] (f : αᵐᵒᵖ ≃+* βᵐᵒᵖ) (a✝ : α) :
@[simp]
theorem RingEquiv.op_apply_apply {α : Type u_7} {β : Type u_8} [Add α] [Mul α] [Add β] [Mul β] (f : α ≃+* β) (a✝ : αᵐᵒᵖ) :
@[simp]
theorem RingEquiv.op_apply_symm_apply {α : Type u_7} {β : Type u_8} [Add α] [Mul α] [Add β] [Mul β] (f : α ≃+* β) (a✝ : βᵐᵒᵖ) :
def RingEquiv.unop {α : Type u_7} {β : Type u_8} [Add α] [Mul α] [Add β] [Mul β] :

The 'unopposite' of a ring iso αᵐᵒᵖ ≃+* βᵐᵒᵖ. Inverse to RingEquiv.op.

Equations

A ring is isomorphic to the opposite of its opposite.

Equations
@[simp]
@[simp]
theorem RingEquiv.opOp_apply (R : Type u_7) [Add R] [Mul R] (a✝ : R) :

A non-unital commutative ring is isomorphic to its opposite.

Equations
theorem RingEquiv.map_zero {R : Type u_4} {S : Type u_5} [NonUnitalNonAssocSemiring R] [NonUnitalNonAssocSemiring S] (f : R ≃+* S) :
f 0 = 0

A ring isomorphism sends zero to zero.

theorem RingEquiv.map_eq_zero_iff {R : Type u_4} {S : Type u_5} [NonUnitalNonAssocSemiring R] [NonUnitalNonAssocSemiring S] (f : R ≃+* S) {x : R} :
f x = 0 x = 0
theorem RingEquiv.map_ne_zero_iff {R : Type u_4} {S : Type u_5} [NonUnitalNonAssocSemiring R] [NonUnitalNonAssocSemiring S] (f : R ≃+* S) {x : R} :
f x 0 x 0
noncomputable def RingEquiv.ofBijective {F : Type u_1} {R : Type u_4} {S : Type u_5} [NonUnitalNonAssocSemiring R] [NonUnitalNonAssocSemiring S] [FunLike F R S] [NonUnitalRingHomClass F R S] (f : F) (hf : Function.Bijective f) :
R ≃+* S

Produce a ring isomorphism from a bijective ring homomorphism.

Equations
@[simp]
theorem RingEquiv.coe_ofBijective {F : Type u_1} {R : Type u_4} {S : Type u_5} [NonUnitalNonAssocSemiring R] [NonUnitalNonAssocSemiring S] [FunLike F R S] [NonUnitalRingHomClass F R S] (f : F) (hf : Function.Bijective f) :
(ofBijective f hf) = f
theorem RingEquiv.ofBijective_apply {F : Type u_1} {R : Type u_4} {S : Type u_5} [NonUnitalNonAssocSemiring R] [NonUnitalNonAssocSemiring S] [FunLike F R S] [NonUnitalRingHomClass F R S] (f : F) (hf : Function.Bijective f) (x : R) :
(ofBijective f hf) x = f x
def RingEquiv.piUnique {ι : Type u_7} (R : ιType u_8) [Unique ι] [(i : ι) → NonUnitalNonAssocSemiring (R i)] :
((i : ι) → R i) ≃+* R default

Product of a singleton family of (non-unital non-associative semi)rings is isomorphic to the only member of this family.

Equations
@[simp]
theorem RingEquiv.piUnique_symm_apply {ι : Type u_7} (R : ιType u_8) [Unique ι] [(i : ι) → NonUnitalNonAssocSemiring (R i)] :
@[simp]
theorem RingEquiv.piUnique_apply {ι : Type u_7} (R : ιType u_8) [Unique ι] [(i : ι) → NonUnitalNonAssocSemiring (R i)] :
(piUnique R) = fun (f : (i : ι) → R i) => f default
def RingEquiv.piCongrRight {ι : Type u_7} {R : ιType u_8} {S : ιType u_9} [(i : ι) → NonUnitalNonAssocSemiring (R i)] [(i : ι) → NonUnitalNonAssocSemiring (S i)] (e : (i : ι) → R i ≃+* S i) :
((i : ι) → R i) ≃+* ((i : ι) → S i)

A family of ring isomorphisms ∀ j, (R j ≃+* S j) generates a ring isomorphisms between ∀ j, R j and ∀ j, S j.

This is the RingEquiv version of Equiv.piCongrRight, and the dependent version of RingEquiv.arrowCongr.

Equations
  • One or more equations did not get rendered due to their size.
@[simp]
theorem RingEquiv.piCongrRight_apply {ι : Type u_7} {R : ιType u_8} {S : ιType u_9} [(i : ι) → NonUnitalNonAssocSemiring (R i)] [(i : ι) → NonUnitalNonAssocSemiring (S i)] (e : (i : ι) → R i ≃+* S i) (x : (i : ι) → R i) (j : ι) :
(piCongrRight e) x j = (e j) (x j)
@[simp]
theorem RingEquiv.piCongrRight_refl {ι : Type u_7} {R : ιType u_8} [(i : ι) → NonUnitalNonAssocSemiring (R i)] :
(piCongrRight fun (i : ι) => refl (R i)) = refl ((i : ι) → R i)
@[simp]
theorem RingEquiv.piCongrRight_symm {ι : Type u_7} {R : ιType u_8} {S : ιType u_9} [(i : ι) → NonUnitalNonAssocSemiring (R i)] [(i : ι) → NonUnitalNonAssocSemiring (S i)] (e : (i : ι) → R i ≃+* S i) :
(piCongrRight e).symm = piCongrRight fun (i : ι) => (e i).symm
@[simp]
theorem RingEquiv.piCongrRight_trans {ι : Type u_7} {R : ιType u_8} {S : ιType u_9} {T : ιType u_10} [(i : ι) → NonUnitalNonAssocSemiring (R i)] [(i : ι) → NonUnitalNonAssocSemiring (S i)] [(i : ι) → NonUnitalNonAssocSemiring (T i)] (e : (i : ι) → R i ≃+* S i) (f : (i : ι) → S i ≃+* T i) :
(piCongrRight e).trans (piCongrRight f) = piCongrRight fun (i : ι) => (e i).trans (f i)
def RingEquiv.piCongrLeft' {ι : Type u_7} {ι' : Type u_8} (R : ιType u_9) (e : ι ι') [(i : ι) → NonUnitalNonAssocSemiring (R i)] :
((i : ι) → R i) ≃+* ((i : ι') → R (e.symm i))

Transport dependent functions through an equivalence of the base space.

This is Equiv.piCongrLeft' as a RingEquiv.

Equations
@[simp]
theorem RingEquiv.piCongrLeft'_apply {ι : Type u_7} {ι' : Type u_8} (R : ιType u_9) (e : ι ι') [(i : ι) → NonUnitalNonAssocSemiring (R i)] (f : (a : ι) → R a) (x : ι') :
(piCongrLeft' R e) f x = f (e.symm x)
@[simp]
theorem RingEquiv.piCongrLeft'_symm_apply {ι : Type u_7} {ι' : Type u_8} (R : ιType u_9) (e : ι ι') [(i : ι) → NonUnitalNonAssocSemiring (R i)] (f : (b : ι') → R (e.symm b)) (x : ι) :
(piCongrLeft' R e).symm f x = f (e x)
@[simp]
theorem RingEquiv.piCongrLeft'_symm {α : Type u_2} {β : Type u_3} {R : Type u_7} [NonUnitalNonAssocSemiring R] (e : α β) :
(piCongrLeft' (fun (x : α) => R) e).symm = piCongrLeft' (fun (i : β) => R) e.symm
def RingEquiv.piCongrLeft {ι : Type u_7} {ι' : Type u_8} (S : ι'Type u_9) (e : ι ι') [(i : ι') → NonUnitalNonAssocSemiring (S i)] :
((i : ι) → S (e i)) ≃+* ((i : ι') → S i)

Transport dependent functions through an equivalence of the base space.

This is Equiv.piCongrLeft as a RingEquiv.

Equations
@[simp]
theorem RingEquiv.piCongrLeft_symm_apply {ι : Type u_7} {ι' : Type u_8} (S : ι'Type u_9) (e : ι ι') [(i : ι') → NonUnitalNonAssocSemiring (S i)] (a✝ : (i : ι') → S i) (i : ι) :
(piCongrLeft S e).symm a✝ i = (piCongrLeft' S e.symm) a✝ i
@[simp]
theorem RingEquiv.piCongrLeft_apply {ι : Type u_7} {ι' : Type u_8} (S : ι'Type u_9) (e : ι ι') [(i : ι') → NonUnitalNonAssocSemiring (S i)] (a✝ : (i : ι) → S (e.symm.symm i)) (i : ι') :
(piCongrLeft S e) a✝ i = (↑(piCongrLeft' S e.symm)).symm a✝ i
def RingEquiv.piEquivPiSubtypeProd {ι : Type u_7} (p : ιProp) [DecidablePred p] (Y : ιType u_8) [(i : ι) → NonUnitalNonAssocSemiring (Y i)] :
((i : ι) → Y i) ≃+* ((i : { x : ι // p x }) → Y i) × ((i : { x : ι // ¬p x }) → Y i)

Splits the indices of ring ∀ (i : ι), Y i along the predicate p. This is Equiv.piEquivPiSubtypeProd as a RingEquiv.

Equations
@[simp]
theorem RingEquiv.piEquivPiSubtypeProd_symm_apply {ι : Type u_7} (p : ιProp) [DecidablePred p] (Y : ιType u_8) [(i : ι) → NonUnitalNonAssocSemiring (Y i)] (f : ((i : { x : ι // p x }) → Y i) × ((i : { x : ι // ¬p x }) → Y i)) (x : ι) :
(piEquivPiSubtypeProd p Y).symm f x = if h : p x then f.1 x, h else f.2 x, h
@[simp]
theorem RingEquiv.piEquivPiSubtypeProd_apply {ι : Type u_7} (p : ιProp) [DecidablePred p] (Y : ιType u_8) [(i : ι) → NonUnitalNonAssocSemiring (Y i)] (f : (i : ι) → Y i) :
(piEquivPiSubtypeProd p Y) f = (fun (x : { x : ι // p x }) => f x, fun (x : { x : ι // ¬p x }) => f x)
def RingEquiv.prodCongr {R : Type u_7} {R' : Type u_8} {S : Type u_9} {S' : Type u_10} [NonUnitalNonAssocSemiring R] [NonUnitalNonAssocSemiring R'] [NonUnitalNonAssocSemiring S] [NonUnitalNonAssocSemiring S'] (f : R ≃+* R') (g : S ≃+* S') :
R × S ≃+* R' × S'

Product of ring equivalences. This is Equiv.prodCongr as a RingEquiv.

Equations
@[simp]
theorem RingEquiv.prodCongr_apply {R : Type u_7} {R' : Type u_8} {S : Type u_9} {S' : Type u_10} [NonUnitalNonAssocSemiring R] [NonUnitalNonAssocSemiring R'] [NonUnitalNonAssocSemiring S] [NonUnitalNonAssocSemiring S'] (f : R ≃+* R') (g : S ≃+* S') (a✝ : R × S) :
(f.prodCongr g) a✝ = Prod.map (⇑f) (⇑g) a✝
@[simp]
theorem RingEquiv.prodCongr_symm_apply {R : Type u_7} {R' : Type u_8} {S : Type u_9} {S' : Type u_10} [NonUnitalNonAssocSemiring R] [NonUnitalNonAssocSemiring R'] [NonUnitalNonAssocSemiring S] [NonUnitalNonAssocSemiring S'] (f : R ≃+* R') (g : S ≃+* S') (a✝ : R' × S') :
(f.prodCongr g).symm a✝ = Prod.map (⇑(↑f).symm) (⇑(↑g).symm) a✝
@[simp]
theorem RingEquiv.coe_prodCongr {R : Type u_7} {R' : Type u_8} {S : Type u_9} {S' : Type u_10} [NonUnitalNonAssocSemiring R] [NonUnitalNonAssocSemiring R'] [NonUnitalNonAssocSemiring S] [NonUnitalNonAssocSemiring S'] (f : R ≃+* R') (g : S ≃+* S') :
(f.prodCongr g) = Prod.map f g
def RingEquiv.piOptionEquivProd {ι : Type u_7} {R : Option ιType u_8} [(i : Option ι) → NonUnitalNonAssocSemiring (R i)] :
((i : Option ι) → R i) ≃+* R none × ((i : ι) → R (some i))

This is Equiv.piOptionEquivProd as a RingEquiv.

Equations
@[simp]
theorem RingEquiv.piOptionEquivProd_apply {ι : Type u_7} {R : Option ιType u_8} [(i : Option ι) → NonUnitalNonAssocSemiring (R i)] (f : (a : Option ι) → R a) :
piOptionEquivProd f = (f none, fun (a : ι) => f (some a))
@[simp]
theorem RingEquiv.piOptionEquivProd_symm_apply {ι : Type u_7} {R : Option ιType u_8} [(i : Option ι) → NonUnitalNonAssocSemiring (R i)] (x : R none × ((a : ι) → R (some a))) (a : Option ι) :
piOptionEquivProd.symm x a = Option.rec x.1 (fun (val : ι) => x.2 val) a
theorem RingEquiv.map_one {R : Type u_4} {S : Type u_5} [NonAssocSemiring R] [NonAssocSemiring S] (f : R ≃+* S) :
f 1 = 1

A ring isomorphism sends one to one.

theorem RingEquiv.map_eq_one_iff {R : Type u_4} {S : Type u_5} [NonAssocSemiring R] [NonAssocSemiring S] (f : R ≃+* S) {x : R} :
f x = 1 x = 1
theorem RingEquiv.map_ne_one_iff {R : Type u_4} {S : Type u_5} [NonAssocSemiring R] [NonAssocSemiring S] (f : R ≃+* S) {x : R} :
f x 1 x 1

RingEquiv.coe_mulEquiv_refl and RingEquiv.coe_addEquiv_refl are proved above in higher generality

@[simp]
@[simp]
theorem RingEquiv.coe_monoidHom_trans {R : Type u_4} {S : Type u_5} {S' : Type u_6} [NonAssocSemiring R] [NonAssocSemiring S] [NonAssocSemiring S'] (e₁ : R ≃+* S) (e₂ : S ≃+* S') :
(e₁.trans e₂) = (↑e₂).comp e₁
@[simp]
theorem RingEquiv.coe_addMonoidHom_trans {R : Type u_4} {S : Type u_5} {S' : Type u_6} [NonAssocSemiring R] [NonAssocSemiring S] [NonUnitalNonAssocSemiring S'] (e₁ : R ≃+* S) (e₂ : S ≃+* S') :
(e₁.trans e₂) = (↑e₂).comp e₁

RingEquiv.coe_mulEquiv_trans and RingEquiv.coe_addEquiv_trans are proved above in higher generality

@[simp]
theorem RingEquiv.coe_ringHom_trans {R : Type u_4} {S : Type u_5} {S' : Type u_6} [NonAssocSemiring R] [NonAssocSemiring S] [NonAssocSemiring S'] (e₁ : R ≃+* S) (e₂ : S ≃+* S') :
(e₁.trans e₂) = (↑e₂).comp e₁
@[simp]
theorem RingEquiv.comp_symm {R : Type u_4} {S : Type u_5} [NonAssocSemiring R] [NonAssocSemiring S] (e : R ≃+* S) :
(↑e).comp e.symm = RingHom.id S
@[simp]
theorem RingEquiv.symm_comp {R : Type u_4} {S : Type u_5} [NonAssocSemiring R] [NonAssocSemiring S] (e : R ≃+* S) :
(↑e.symm).comp e = RingHom.id R
theorem RingEquiv.map_neg {R : Type u_4} {S : Type u_5} [NonUnitalNonAssocRing R] [NonUnitalNonAssocRing S] (f : R ≃+* S) (x : R) :
f (-x) = -f x
theorem RingEquiv.map_sub {R : Type u_4} {S : Type u_5} [NonUnitalNonAssocRing R] [NonUnitalNonAssocRing S] (f : R ≃+* S) (x y : R) :
f (x - y) = f x - f y
@[simp]
theorem RingEquiv.map_neg_one {R : Type u_4} {S : Type u_5} [NonAssocRing R] [NonAssocRing S] (f : R ≃+* S) :
f (-1) = -1
theorem RingEquiv.map_eq_neg_one_iff {R : Type u_4} {S : Type u_5} [NonAssocRing R] [NonAssocRing S] (f : R ≃+* S) {x : R} :
f x = -1 x = -1

Reinterpret a ring equivalence as a non-unital ring homomorphism.

Equations
@[simp]
theorem RingEquiv.coe_toNonUnitalRingHom {R : Type u_4} {S : Type u_5} [NonUnitalNonAssocSemiring R] [NonUnitalNonAssocSemiring S] (f : R ≃+* S) :
f = f
def RingEquiv.toRingHom {R : Type u_4} {S : Type u_5} [NonAssocSemiring R] [NonAssocSemiring S] (e : R ≃+* S) :
R →+* S

Reinterpret a ring equivalence as a ring homomorphism.

Equations
@[simp]
theorem RingEquiv.toRingHom_eq_coe {R : Type u_4} {S : Type u_5} [NonAssocSemiring R] [NonAssocSemiring S] (f : R ≃+* S) :
f.toRingHom = f
@[simp]
theorem RingEquiv.coe_toRingHom {R : Type u_4} {S : Type u_5} [NonAssocSemiring R] [NonAssocSemiring S] (f : R ≃+* S) :
f = f
theorem RingEquiv.coe_ringHom_inj_iff {R : Type u_7} {S : Type u_8} [NonAssocSemiring R] [NonAssocSemiring S] (f g : R ≃+* S) :
f = g f = g
@[simp]
theorem RingEquiv.toNonUnitalRingHom_commutes {R : Type u_4} {S : Type u_5} [NonAssocSemiring R] [NonAssocSemiring S] (f : R ≃+* S) :
f = f

The two paths coercion can take to a NonUnitalRingEquiv are equivalent

@[reducible, inline]
abbrev RingEquiv.toMonoidHom {R : Type u_4} {S : Type u_5} [NonAssocSemiring R] [NonAssocSemiring S] (e : R ≃+* S) :
R →* S

Reinterpret a ring equivalence as a monoid homomorphism.

Equations
@[reducible, inline]
abbrev RingEquiv.toAddMonoidHom {R : Type u_4} {S : Type u_5} [NonAssocSemiring R] [NonAssocSemiring S] (e : R ≃+* S) :
R →+ S

Reinterpret a ring equivalence as an AddMonoid homomorphism.

Equations

The two paths coercion can take to an AddMonoidHom are equivalent

theorem RingEquiv.toMonoidHom_commutes {R : Type u_4} {S : Type u_5} [NonAssocSemiring R] [NonAssocSemiring S] (f : R ≃+* S) :
f = (↑f).toMonoidHom

The two paths coercion can take to a MonoidHom are equivalent

theorem RingEquiv.toEquiv_commutes {R : Type u_4} {S : Type u_5} [NonAssocSemiring R] [NonAssocSemiring S] (f : R ≃+* S) :
(↑f).toEquiv = (↑f).toEquiv

The two paths coercion can take to an Equiv are equivalent

@[simp]
theorem RingEquiv.toRingHom_trans {R : Type u_4} {S : Type u_5} {S' : Type u_6} [NonAssocSemiring R] [NonAssocSemiring S] [NonAssocSemiring S'] (e₁ : R ≃+* S) (e₂ : S ≃+* S') :
(e₁.trans e₂).toRingHom = e₂.toRingHom.comp e₁.toRingHom
def RingEquiv.ofHomInv' {R : Type u_7} {S : Type u_8} {F : Type u_9} {G : Type u_10} [NonUnitalNonAssocSemiring R] [NonUnitalNonAssocSemiring S] [FunLike F R S] [FunLike G S R] [NonUnitalRingHomClass F R S] [NonUnitalRingHomClass G S R] (hom : F) (inv : G) (hom_inv_id : (↑inv).comp hom = NonUnitalRingHom.id R) (inv_hom_id : (↑hom).comp inv = NonUnitalRingHom.id S) :
R ≃+* S

Construct an equivalence of rings from homomorphisms in both directions, which are inverses.

Equations
  • RingEquiv.ofHomInv' hom inv hom_inv_id inv_hom_id = { toFun := hom, invFun := inv, left_inv := , right_inv := , map_mul' := , map_add' := }
@[simp]
theorem RingEquiv.ofHomInv'_apply {R : Type u_7} {S : Type u_8} {F : Type u_9} {G : Type u_10} [NonUnitalNonAssocSemiring R] [NonUnitalNonAssocSemiring S] [FunLike F R S] [FunLike G S R] [NonUnitalRingHomClass F R S] [NonUnitalRingHomClass G S R] (hom : F) (inv : G) (hom_inv_id : (↑inv).comp hom = NonUnitalRingHom.id R) (inv_hom_id : (↑hom).comp inv = NonUnitalRingHom.id S) (a : R) :
(ofHomInv' hom inv hom_inv_id inv_hom_id) a = hom a
@[simp]
theorem RingEquiv.ofHomInv'_symm_apply {R : Type u_7} {S : Type u_8} {F : Type u_9} {G : Type u_10} [NonUnitalNonAssocSemiring R] [NonUnitalNonAssocSemiring S] [FunLike F R S] [FunLike G S R] [NonUnitalRingHomClass F R S] [NonUnitalRingHomClass G S R] (hom : F) (inv : G) (hom_inv_id : (↑inv).comp hom = NonUnitalRingHom.id R) (inv_hom_id : (↑hom).comp inv = NonUnitalRingHom.id S) (a : S) :
(ofHomInv' hom inv hom_inv_id inv_hom_id).symm a = inv a
def RingEquiv.ofHomInv {R : Type u_7} {S : Type u_8} {F : Type u_9} {G : Type u_10} [NonAssocSemiring R] [NonAssocSemiring S] [FunLike F R S] [FunLike G S R] [RingHomClass F R S] [RingHomClass G S R] (hom : F) (inv : G) (hom_inv_id : (↑inv).comp hom = RingHom.id R) (inv_hom_id : (↑hom).comp inv = RingHom.id S) :
R ≃+* S

Construct an equivalence of rings from unital homomorphisms in both directions, which are inverses.

Equations
  • RingEquiv.ofHomInv hom inv hom_inv_id inv_hom_id = { toFun := hom, invFun := inv, left_inv := , right_inv := , map_mul' := , map_add' := }
@[simp]
theorem RingEquiv.ofHomInv_symm_apply {R : Type u_7} {S : Type u_8} {F : Type u_9} {G : Type u_10} [NonAssocSemiring R] [NonAssocSemiring S] [FunLike F R S] [FunLike G S R] [RingHomClass F R S] [RingHomClass G S R] (hom : F) (inv : G) (hom_inv_id : (↑inv).comp hom = RingHom.id R) (inv_hom_id : (↑hom).comp inv = RingHom.id S) (a : S) :
(ofHomInv hom inv hom_inv_id inv_hom_id).symm a = inv a
@[simp]
theorem RingEquiv.ofHomInv_apply {R : Type u_7} {S : Type u_8} {F : Type u_9} {G : Type u_10} [NonAssocSemiring R] [NonAssocSemiring S] [FunLike F R S] [FunLike G S R] [RingHomClass F R S] [RingHomClass G S R] (hom : F) (inv : G) (hom_inv_id : (↑inv).comp hom = RingHom.id R) (inv_hom_id : (↑hom).comp inv = RingHom.id S) (a : R) :
(ofHomInv hom inv hom_inv_id inv_hom_id) a = hom a
theorem RingEquiv.map_pow {R : Type u_4} {S : Type u_5} [Semiring R] [Semiring S] (f : R ≃+* S) (a : R) (n : ) :
f (a ^ n) = f a ^ n
def MulEquiv.toRingEquiv {R : Type u_7} {S : Type u_8} {F : Type u_9} [Add R] [Add S] [Mul R] [Mul S] [EquivLike F R S] [MulEquivClass F R S] (f : F) (H : ∀ (x y : R), f (x + y) = f x + f y) :
R ≃+* S

Gives a RingEquiv from an element of a MulEquivClass preserving addition.

Equations
def AddEquiv.toRingEquiv {R : Type u_7} {S : Type u_8} {F : Type u_9} [Add R] [Add S] [Mul R] [Mul S] [EquivLike F R S] [AddEquivClass F R S] (f : F) (H : ∀ (x y : R), f (x * y) = f x * f y) :
R ≃+* S

Gives a RingEquiv from an element of an AddEquivClass preserving addition.

Equations
@[simp]
theorem RingEquiv.self_trans_symm {R : Type u_4} {S : Type u_5} [Add R] [Add S] [Mul R] [Mul S] (e : R ≃+* S) :
@[simp]
theorem RingEquiv.symm_trans_self {R : Type u_4} {S : Type u_5} [Add R] [Add S] [Mul R] [Mul S] (e : R ≃+* S) :
def RingEquiv.ofRingHom {R : Type u_4} {S : Type u_5} [NonAssocSemiring R] [NonAssocSemiring S] (f : R →+* S) (g : S →+* R) (h₁ : f.comp g = RingHom.id S) (h₂ : g.comp f = RingHom.id R) :
R ≃+* S

If a ring homomorphism has an inverse, it is a ring isomorphism.

Equations
  • RingEquiv.ofRingHom f g h₁ h₂ = { toFun := f, invFun := g, left_inv := , right_inv := , map_mul' := , map_add' := }
@[simp]
theorem RingEquiv.ofRingHom_symm_apply {R : Type u_4} {S : Type u_5} [NonAssocSemiring R] [NonAssocSemiring S] (f : R →+* S) (g : S →+* R) (h₁ : f.comp g = RingHom.id S) (h₂ : g.comp f = RingHom.id R) (a : S) :
(ofRingHom f g h₁ h₂).symm a = g a
@[simp]
theorem RingEquiv.ofRingHom_apply {R : Type u_4} {S : Type u_5} [NonAssocSemiring R] [NonAssocSemiring S] (f : R →+* S) (g : S →+* R) (h₁ : f.comp g = RingHom.id S) (h₂ : g.comp f = RingHom.id R) (a : R) :
(ofRingHom f g h₁ h₂) a = f a
theorem RingEquiv.coe_ringHom_ofRingHom {R : Type u_4} {S : Type u_5} [NonAssocSemiring R] [NonAssocSemiring S] (f : R →+* S) (g : S →+* R) (h₁ : f.comp g = RingHom.id S) (h₂ : g.comp f = RingHom.id R) :
(ofRingHom f g h₁ h₂) = f
@[simp]
theorem RingEquiv.ofRingHom_coe_ringHom {R : Type u_4} {S : Type u_5} [NonAssocSemiring R] [NonAssocSemiring S] (f : R ≃+* S) (g : S →+* R) (h₁ : (↑f).comp g = RingHom.id S) (h₂ : g.comp f = RingHom.id R) :
ofRingHom (↑f) g h₁ h₂ = f
theorem RingEquiv.ofRingHom_symm {R : Type u_4} {S : Type u_5} [NonAssocSemiring R] [NonAssocSemiring S] (f : R →+* S) (g : S →+* R) (h₁ : f.comp g = RingHom.id S) (h₂ : g.comp f = RingHom.id R) :
(ofRingHom f g h₁ h₂).symm = ofRingHom g f h₂ h₁
theorem MulEquiv.noZeroDivisors {A : Type u_7} (B : Type u_8) [MulZeroClass A] [MulZeroClass B] [NoZeroDivisors B] (e : A ≃* B) :

If two rings are isomorphic, and the second doesn't have zero divisors, then so does the first.

theorem MulEquiv.isDomain {A : Type u_7} (B : Type u_8) [Semiring A] [Semiring B] [IsDomain B] (e : A ≃* B) :

If two rings are isomorphic, and the second is a domain, then so is the first.