Ideal quotients #
This file defines ideal quotients as a special case of submodule quotients and proves some basic results about these quotients.
See algebra.ring_quot
for quotients of non-commutative rings.
Main definitions #
ideal.quotient
: the quotient of a commutative ringR
by an idealI : ideal R
Main results #
ideal.quotient_inf_ring_equiv_pi_quotient
: the Chinese Remainder Theorem
The quotient R/I
of a ring R
by an ideal I
.
The ideal quotient of I
is defined to equal the quotient of I
as an R
-submodule of R
.
This definition is marked reducible
so that typeclass instances can be shared between
ideal.quotient I
and submodule.quotient I
.
Equations
Equations
Equations
- ideal.quotient.has_mul I = {mul := λ (a b : R ⧸ I), quotient.lift_on₂' a b (λ (a b : R), submodule.quotient.mk (a * b)) _}
Equations
- ideal.quotient.comm_ring I = {add := add_comm_group.add (submodule.quotient.add_comm_group I), add_assoc := _, zero := add_comm_group.zero (submodule.quotient.add_comm_group I), zero_add := _, add_zero := _, nsmul := add_comm_group.nsmul (submodule.quotient.add_comm_group I), nsmul_zero' := _, nsmul_succ' := _, neg := add_comm_group.neg (submodule.quotient.add_comm_group I), sub := add_comm_group.sub (submodule.quotient.add_comm_group I), sub_eq_add_neg := _, zsmul := add_comm_group.zsmul (submodule.quotient.add_comm_group I), zsmul_zero' := _, zsmul_succ' := _, zsmul_neg' := _, add_left_neg := _, add_comm := _, mul := has_mul.mul (ideal.quotient.has_mul I), mul_assoc := _, one := 1, one_mul := _, mul_one := _, npow := ring.npow._default 1 has_mul.mul _ _, npow_zero' := _, npow_succ' := _, left_distrib := _, right_distrib := _, mul_comm := _}
The ring homomorphism from a ring R
to a quotient ring R/I
.
Equations
- ideal.quotient.mk I = {to_fun := λ (a : R), submodule.quotient.mk a, map_one' := _, map_mul' := _, map_zero' := _, map_add' := _}
Equations
- ideal.quotient.inhabited = {default := ⇑(ideal.quotient.mk I) 37}
Equations
- ideal.quotient.has_quotient.quotient.unique = {to_inhabited := {default := 0}, uniq := _}
If I
is an ideal of a commutative ring R
, if q : R → R/I
is the quotient map, and if
s ⊆ R
is a subset, then q⁻¹(q(s)) = ⋃ᵢ(i + s)
, the union running over all i ∈ I
.
quotient by maximal ideal is a field. def rather than instance, since users will have computable inverses in some applications. See note [reducible non-instances].
Equations
- ideal.quotient.field I = {add := comm_ring.add (ideal.quotient.comm_ring I), add_assoc := _, zero := comm_ring.zero (ideal.quotient.comm_ring I), zero_add := _, add_zero := _, nsmul := comm_ring.nsmul (ideal.quotient.comm_ring I), nsmul_zero' := _, nsmul_succ' := _, neg := comm_ring.neg (ideal.quotient.comm_ring I), sub := comm_ring.sub (ideal.quotient.comm_ring I), sub_eq_add_neg := _, zsmul := comm_ring.zsmul (ideal.quotient.comm_ring I), zsmul_zero' := _, zsmul_succ' := _, zsmul_neg' := _, add_left_neg := _, add_comm := _, mul := comm_ring.mul (ideal.quotient.comm_ring I), mul_assoc := _, one := comm_ring.one (ideal.quotient.comm_ring I), one_mul := _, mul_one := _, npow := comm_ring.npow (ideal.quotient.comm_ring I), npow_zero' := _, npow_succ' := _, left_distrib := _, right_distrib := _, mul_comm := _, inv := λ (a : R ⧸ I), dite (a = 0) (λ (ha : a = 0), 0) (λ (ha : ¬a = 0), classical.some _), div := div_inv_monoid.div._default comm_ring.mul _ comm_ring.one _ _ comm_ring.npow _ _ (λ (a : R ⧸ I), dite (a = 0) (λ (ha : a = 0), 0) (λ (ha : ¬a = 0), classical.some _)), div_eq_mul_inv := _, zpow := div_inv_monoid.zpow._default comm_ring.mul _ comm_ring.one _ _ comm_ring.npow _ _ (λ (a : R ⧸ I), dite (a = 0) (λ (ha : a = 0), 0) (λ (ha : ¬a = 0), classical.some _)), zpow_zero' := _, zpow_succ' := _, zpow_neg' := _, exists_pair_ne := _, mul_inv_cancel := _, inv_zero := _}
If the quotient by an ideal is a field, then the ideal is maximal.
The quotient of a ring by an ideal is a field iff the ideal is maximal.
The ring homomorphism from the quotient by a smaller ideal to the quotient by a larger ideal.
This is the ideal.quotient
version of quot.factor
Equations
- ideal.quotient.factor S T H = ideal.quotient.lift S (ideal.quotient.mk T) _
Quotienting by equal ideals gives equivalent rings.
See also submodule.quot_equiv_of_eq
.
Equations
- ideal.quot_equiv_of_eq h = {to_fun := (submodule.quot_equiv_of_eq I J h).to_fun, inv_fun := (submodule.quot_equiv_of_eq I J h).inv_fun, left_inv := _, right_inv := _, map_mul' := _, map_add' := _}
R^n/I^n
is a R/I
-module.
Equations
- I.module_pi ι = {to_distrib_mul_action := {to_mul_action := {to_has_scalar := {smul := λ (c : R ⧸ I) (m : (ι → R) ⧸ I.pi ι), quotient.lift_on₂' c m (λ (r : R) (m : ι → R), submodule.quotient.mk (r • m)) _}, one_smul := _, mul_smul := _}, smul_add := _, smul_zero := _}, add_smul := _, zero_smul := _}
R^n/I^n
is isomorphic to (R/I)^n
as an R/I
-module.
Equations
- I.pi_quot_equiv ι = {to_fun := λ (x : (ι → R) ⧸ I.pi ι), quotient.lift_on' x (λ (f : ι → R) (i : ι), ⇑(ideal.quotient.mk I) (f i)) _, map_add' := _, map_smul' := _, inv_fun := λ (x : ι → R ⧸ I), ⇑(ideal.quotient.mk (I.pi ι)) (λ (i : ι), quotient.out' (x i)), left_inv := _, right_inv := _}
The homomorphism from R/(⋂ i, f i)
to ∏ i, (R / f i)
featured in the Chinese
Remainder Theorem. It is bijective if the ideals f i
are comaximal.
Equations
- ideal.quotient_inf_to_pi_quotient f = ideal.quotient.lift (⨅ (i : ι), f i) (pi.ring_hom (λ (i : ι), ideal.quotient.mk (f i))) _
Chinese Remainder Theorem. Eisenbud Ex.2.6. Similar to Atiyah-Macdonald 1.10 and Stacks 00DT
Equations
- ideal.quotient_inf_ring_equiv_pi_quotient f hf = {to_fun := (equiv.of_bijective ⇑(ideal.quotient_inf_to_pi_quotient (λ (i : ι), f i)) _).to_fun, inv_fun := (equiv.of_bijective ⇑(ideal.quotient_inf_to_pi_quotient (λ (i : ι), f i)) _).inv_fun, left_inv := _, right_inv := _, map_mul' := _, map_add' := _}