@[instance]
def
submodule.has_scalar'
{R : Type u}
{M : Type v}
[comm_ring R]
[add_comm_group M]
[module R M] :
has_scalar (ideal R) (submodule R M)
Equations
- submodule.has_scalar' = {smul := λ (I : ideal R) (N : submodule R M), ⨆ (r : ↥I), submodule.map (r.val • linear_map.id) N}
def
submodule.annihilator
{R : Type u}
{M : Type v}
[comm_ring R]
[add_comm_group M]
[module R M] :
Equations
- N.annihilator = (linear_map.lsmul R ↥N).ker
Equations
- N.colon P = (submodule.map N.mkq P).annihilator
theorem
submodule.mem_annihilator
{R : Type u}
{M : Type v}
[comm_ring R]
[add_comm_group M]
[module R M]
{N : submodule R M}
{r : R} :
theorem
submodule.mem_annihilator'
{R : Type u}
{M : Type v}
[comm_ring R]
[add_comm_group M]
[module R M]
{N : submodule R M}
{r : R} :
r ∈ N.annihilator ↔ N ≤ submodule.comap (r • linear_map.id) ⊥
theorem
submodule.annihilator_bot
{R : Type u}
{M : Type v}
[comm_ring R]
[add_comm_group M]
[module R M] :
theorem
submodule.annihilator_eq_top_iff
{R : Type u}
{M : Type v}
[comm_ring R]
[add_comm_group M]
[module R M]
{N : submodule R M} :
theorem
submodule.annihilator_mono
{R : Type u}
{M : Type v}
[comm_ring R]
[add_comm_group M]
[module R M]
{N P : submodule R M} :
N ≤ P → P.annihilator ≤ N.annihilator
theorem
submodule.annihilator_supr
{R : Type u}
{M : Type v}
[comm_ring R]
[add_comm_group M]
[module R M]
(ι : Sort w)
(f : ι → submodule R M) :
(⨆ (i : ι), f i).annihilator = ⨅ (i : ι), (f i).annihilator
theorem
submodule.mem_colon'
{R : Type u}
{M : Type v}
[comm_ring R]
[add_comm_group M]
[module R M]
{N P : submodule R M}
{r : R} :
r ∈ N.colon P ↔ P ≤ submodule.comap (r • linear_map.id) N
theorem
submodule.infi_colon_supr
{R : Type u}
{M : Type v}
[comm_ring R]
[add_comm_group M]
[module R M]
(ι₁ : Sort w)
(f : ι₁ → submodule R M)
(ι₂ : Sort x)
(g : ι₂ → submodule R M) :
theorem
submodule.smul_induction_on
{R : Type u}
{M : Type v}
[comm_ring R]
[add_comm_group M]
[module R M]
{I : ideal R}
{N : submodule R M}
{p : M → Prop}
{x : M} :
theorem
submodule.mem_smul_span_singleton
{R : Type u}
{M : Type v}
[comm_ring R]
[add_comm_group M]
[module R M]
{I : ideal R}
{m x : M} :
theorem
submodule.smul_le_right
{R : Type u}
{M : Type v}
[comm_ring R]
[add_comm_group M]
[module R M]
{I : ideal R}
{N : submodule R M} :
@[simp]
theorem
submodule.smul_bot
{R : Type u}
{M : Type v}
[comm_ring R]
[add_comm_group M]
[module R M]
(I : ideal R) :
@[simp]
theorem
submodule.bot_smul
{R : Type u}
{M : Type v}
[comm_ring R]
[add_comm_group M]
[module R M]
(N : submodule R M) :
@[simp]
theorem
submodule.top_smul
{R : Type u}
{M : Type v}
[comm_ring R]
[add_comm_group M]
[module R M]
(N : submodule R M) :
theorem
submodule.span_smul_span
{R : Type u}
{M : Type v}
[comm_ring R]
[add_comm_group M]
[module R M]
(S : set R)
(T : set M) :
ideal.span S • submodule.span R T = submodule.span R (⋃ (s : R) (H : s ∈ S) (t : M) (H : t ∈ T), {s • t})
theorem
submodule.map_smul''
{R : Type u}
{M : Type v}
[comm_ring R]
[add_comm_group M]
[module R M]
(I : ideal R)
(N : submodule R M)
{M' : Type w}
[add_comm_group M']
[module R M']
(f : M →ₗ[R] M') :
submodule.map f (I • N) = I • submodule.map f N
Equations
- ideal.quotient_inf_to_pi_quotient f = ideal.quotient.lift (⨅ (i : ι), f i) (_.mpr (pi.ring_hom (λ (i : ι), ideal.quotient.mk (f i)))) _
theorem
ideal.quotient_inf_to_pi_quotient_bijective
{R : Type u}
[comm_ring R]
{ι : Type v}
[fintype ι]
{f : ι → ideal R} :
(∀ (i j : ι), i ≠ j → f i ⊔ f j = ⊤) → function.bijective ⇑(ideal.quotient_inf_to_pi_quotient f)
def
ideal.quotient_inf_ring_equiv_pi_quotient
{R : Type u}
[comm_ring R]
{ι : Type v}
[fintype ι]
(f : ι → ideal R) :
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' := _}
@[instance]
Equations
theorem
ideal.span_mul_span
{R : Type u}
[comm_ring R]
(S T : set R) :
ideal.span S * ideal.span T = ideal.span (⋃ (s : R) (H : s ∈ S) (t : R) (H : t ∈ T), {s * t})
@[instance]
Equations
theorem
ideal.comap_mono
{R : Type u}
{S : Type v}
[comm_ring R]
[comm_ring S]
{f : R →+* S}
{K L : ideal S} :
K ≤ L → ideal.comap f K ≤ ideal.comap f L
theorem
ideal.is_prime.comap
{R : Type u}
{S : Type v}
[comm_ring R]
[comm_ring S]
(f : R →+* S)
{K : ideal S}
[hK : K.is_prime] :
(ideal.comap f K).is_prime
theorem
ideal.gc_map_comap
{R : Type u}
{S : Type v}
[comm_ring R]
[comm_ring S]
(f : R →+* S) :
galois_connection (ideal.map f) (ideal.comap f)
@[simp]
theorem
ideal.comap_id
{R : Type u}
[comm_ring R]
(I : ideal R) :
ideal.comap (ring_hom.id R) I = I
@[simp]
theorem
ideal.comap_comap
{R : Type u}
{S : Type v}
[comm_ring R]
[comm_ring S]
{T : Type u_1}
[comm_ring T]
{I : ideal T}
(f : R →+* S)
(g : S →+* T) :
ideal.comap f (ideal.comap g I) = ideal.comap (g.comp f) I
theorem
ideal.le_comap_map
{R : Type u}
{S : Type v}
[comm_ring R]
[comm_ring S]
{f : R →+* S}
{I : ideal R} :
I ≤ ideal.comap f (ideal.map f I)
theorem
ideal.map_comap_le
{R : Type u}
{S : Type v}
[comm_ring R]
[comm_ring S]
{f : R →+* S}
{K : ideal S} :
ideal.map f (ideal.comap f K) ≤ K
@[simp]
theorem
ideal.comap_top
{R : Type u}
{S : Type v}
[comm_ring R]
[comm_ring S]
{f : R →+* S} :
ideal.comap f ⊤ = ⊤
@[simp]
theorem
ideal.comap_map_comap
{R : Type u}
{S : Type v}
[comm_ring R]
[comm_ring S]
(f : R →+* S)
(K : ideal S) :
ideal.comap f (ideal.map f (ideal.comap f K)) = ideal.comap f K
theorem
ideal.comap_inf
{R : Type u}
{S : Type v}
[comm_ring R]
[comm_ring S]
(f : R →+* S)
(K L : ideal S) :
ideal.comap f (K ⊓ L) = ideal.comap f K ⊓ ideal.comap f L
theorem
ideal.comap_infi
{R : Type u}
{S : Type v}
[comm_ring R]
[comm_ring S]
(f : R →+* S)
{ι : Sort u_1}
(K : ι → ideal S) :
ideal.comap f (infi K) = ⨅ (i : ι), ideal.comap f (K i)
theorem
ideal.comap_Inf
{R : Type u}
{S : Type v}
[comm_ring R]
[comm_ring S]
(f : R →+* S)
(s : set (ideal S)) :
ideal.comap f (has_Inf.Inf s) = ⨅ (I : ideal S) (H : I ∈ s), ideal.comap f I
theorem
ideal.comap_Inf'
{R : Type u}
{S : Type v}
[comm_ring R]
[comm_ring S]
(f : R →+* S)
(s : set (ideal S)) :
ideal.comap f (has_Inf.Inf s) = ⨅ (I : ideal R) (H : I ∈ ideal.comap f '' s), I
theorem
ideal.comap_radical
{R : Type u}
{S : Type v}
[comm_ring R]
[comm_ring S]
(f : R →+* S)
(K : ideal S) :
ideal.comap f K.radical = (ideal.comap f K).radical
theorem
ideal.comap_is_prime
{R : Type u}
{S : Type v}
[comm_ring R]
[comm_ring S]
(f : R →+* S)
(K : ideal S)
[H : K.is_prime] :
(ideal.comap f K).is_prime
@[simp]
theorem
ideal.map_quotient_self
{R : Type u}
[comm_ring R]
(I : ideal R) :
ideal.map (ideal.quotient.mk I) I = ⊥
theorem
ideal.le_comap_sup
{R : Type u}
{S : Type v}
[comm_ring R]
[comm_ring S]
(f : R →+* S)
{K L : ideal S} :
ideal.comap f K ⊔ ideal.comap f L ≤ ideal.comap f (K ⊔ L)
theorem
ideal.le_comap_mul
{R : Type u}
{S : Type v}
[comm_ring R]
[comm_ring S]
(f : R →+* S)
{K L : ideal S} :
ideal.comap f K * ideal.comap f L ≤ ideal.comap f (K * L)
theorem
ideal.map_comap_of_surjective
{R : Type u}
{S : Type v}
[comm_ring R]
[comm_ring S]
(f : R →+* S)
(hf : function.surjective ⇑f)
(I : ideal S) :
ideal.map f (ideal.comap f I) = I
def
ideal.gi_map_comap
{R : Type u}
{S : Type v}
[comm_ring R]
[comm_ring S]
(f : R →+* S) :
function.surjective ⇑f → galois_insertion (ideal.map f) (ideal.comap f)
map
and comap
are adjoint, and the composition map f ∘ comap f
is the
identity
Equations
- ideal.gi_map_comap f hf = galois_insertion.monotone_intro _ _ _ _
theorem
ideal.map_surjective_of_surjective
{R : Type u}
{S : Type v}
[comm_ring R]
[comm_ring S]
(f : R →+* S) :
theorem
ideal.comap_injective_of_surjective
{R : Type u}
{S : Type v}
[comm_ring R]
[comm_ring S]
(f : R →+* S) :
theorem
ideal.map_sup_comap_of_surjective
{R : Type u}
{S : Type v}
[comm_ring R]
[comm_ring S]
(f : R →+* S)
(hf : function.surjective ⇑f)
(I J : ideal S) :
ideal.map f (ideal.comap f I ⊔ ideal.comap f J) = I ⊔ J
theorem
ideal.map_supr_comap_of_surjective
{R : Type u}
{S : Type v}
[comm_ring R]
[comm_ring S]
(f : R →+* S)
{ι : Sort u_1}
(hf : function.surjective ⇑f)
(K : ι → ideal S) :
ideal.map f (⨆ (i : ι), ideal.comap f (K i)) = supr K
theorem
ideal.map_inf_comap_of_surjective
{R : Type u}
{S : Type v}
[comm_ring R]
[comm_ring S]
(f : R →+* S)
(hf : function.surjective ⇑f)
(I J : ideal S) :
ideal.map f (ideal.comap f I ⊓ ideal.comap f J) = I ⊓ J
theorem
ideal.map_infi_comap_of_surjective
{R : Type u}
{S : Type v}
[comm_ring R]
[comm_ring S]
(f : R →+* S)
{ι : Sort u_1}
(hf : function.surjective ⇑f)
(K : ι → ideal S) :
ideal.map f (⨅ (i : ι), ideal.comap f (K i)) = infi K
theorem
ideal.comap_map_of_surjective
{R : Type u}
{S : Type v}
[comm_ring R]
[comm_ring S]
(f : R →+* S)
(hf : function.surjective ⇑f)
(I : ideal R) :
ideal.comap f (ideal.map f I) = I ⊔ ideal.comap f ⊥
theorem
ideal.le_map_of_comap_le_of_surjective
{R : Type u}
{S : Type v}
[comm_ring R]
[comm_ring S]
(f : R →+* S)
{I : ideal R}
{K : ideal S} :
function.surjective ⇑f → ideal.comap f K ≤ I → K ≤ ideal.map f I
def
ideal.rel_iso_of_surjective
{R : Type u}
{S : Type v}
[comm_ring R]
[comm_ring S]
(f : R →+* S) :
function.surjective ⇑f → ideal S ≃o {p // ideal.comap f ⊥ ≤ p}
Correspondence theorem
Equations
- ideal.rel_iso_of_surjective f hf = {to_equiv := {to_fun := λ (J : ideal S), ⟨ideal.comap f J, _⟩, inv_fun := λ (I : {p // ideal.comap f ⊥ ≤ p}), ideal.map f I.val, left_inv := _, right_inv := _}, map_rel_iff' := _}
def
ideal.order_embedding_of_surjective
{R : Type u}
{S : Type v}
[comm_ring R]
[comm_ring S]
(f : R →+* S) :
function.surjective ⇑f → ideal S ↪o ideal R
The map on ideals induced by a surjective map preserves inclusion.
Equations
- ideal.order_embedding_of_surjective f hf = (rel_iso.to_rel_embedding (ideal.rel_iso_of_surjective f hf)).trans (subtype.rel_embedding has_le.le (λ (p : ideal R), ideal.comap f ⊥ ≤ p))
theorem
ideal.map_eq_top_or_is_maximal_of_surjective
{R : Type u}
{S : Type v}
[comm_ring R]
[comm_ring S]
(f : R →+* S)
{I : ideal R} :
function.surjective ⇑f → I.is_maximal → ideal.map f I = ⊤ ∨ (ideal.map f I).is_maximal
theorem
ideal.comap_is_maximal_of_surjective
{R : Type u}
{S : Type v}
[comm_ring R]
[comm_ring S]
(f : R →+* S)
{K : ideal S}
(hf : function.surjective ⇑f)
[H : K.is_maximal] :
(ideal.comap f K).is_maximal
theorem
ideal.mem_quotient_iff_mem
{R : Type u}
[comm_ring R]
{I J : ideal R}
(hIJ : I ≤ J)
{x : R} :
⇑(ideal.quotient.mk I) x ∈ ideal.map (ideal.quotient.mk I) J ↔ x ∈ J
theorem
ideal.comap_bot_le_of_injective
{R : Type u}
{S : Type v}
[comm_ring R]
[comm_ring S]
(f : R →+* S)
{I : ideal R} :
function.injective ⇑f → ideal.comap f ⊥ ≤ I
def
ideal.rel_iso_of_bijective
{R : Type u}
{S : Type v}
[comm_ring R]
[comm_ring S]
(f : R →+* S) :
function.bijective ⇑f → ideal S ≃o ideal R
Special case of the correspondence theorem for isomorphic rings
Equations
- ideal.rel_iso_of_bijective f hf = {to_equiv := {to_fun := ideal.comap f, inv_fun := ideal.map f, left_inv := _, right_inv := _}, map_rel_iff' := _}
theorem
ideal.comap_le_iff_le_map
{R : Type u}
{S : Type v}
[comm_ring R]
[comm_ring S]
(f : R →+* S)
{I : ideal R}
{K : ideal S} :
function.bijective ⇑f → (ideal.comap f K ≤ I ↔ K ≤ ideal.map f I)
theorem
ideal.map.is_maximal
{R : Type u}
{S : Type v}
[comm_ring R]
[comm_ring S]
(f : R →+* S)
{I : ideal R} :
function.bijective ⇑f → I.is_maximal → (ideal.map f I).is_maximal
theorem
ideal.is_primary.to_is_prime
{R : Type u}
[comm_ring R]
(I : ideal R) :
I.is_prime → I.is_primary
theorem
ideal.is_prime_radical
{R : Type u}
[comm_ring R]
{I : ideal R} :
I.is_primary → I.radical.is_prime
theorem
ideal.is_primary_inf
{R : Type u}
[comm_ring R]
{I J : ideal R} :
I.is_primary → J.is_primary → I.radical = J.radical → (I ⊓ J).is_primary
Kernel of a ring homomorphism as an ideal of the domain.
Equations
- f.ker = ideal.comap f ⊥
theorem
ring_hom.ker_eq
{R : Type u}
{S : Type v}
[comm_ring R]
[comm_ring S]
(f : R →+* S) :
↑(f.ker) = is_add_group_hom.ker ⇑f
theorem
ring_hom.ker_eq_comap_bot
{R : Type u}
{S : Type v}
[comm_ring R]
[comm_ring S]
(f : R →+* S) :
f.ker = ideal.comap f ⊥
theorem
ring_hom.not_one_mem_ker
{R : Type u}
{S : Type v}
[comm_ring R]
[comm_ring S]
[nontrivial S]
(f : R →+* S) :
If the target is not the zero ring, then one is not in the kernel.
theorem
ring_hom.ker_is_prime
{R : Type u}
{S : Type v}
[comm_ring R]
[integral_domain S]
(f : R →+* S) :
The kernel of a homomorphism to an integral domain is a prime ideal.
@[simp]
theorem
ideal.ker_le_comap
{R : Type u_1}
{S : Type u_2}
[comm_ring R]
[comm_ring S]
{K : ideal S}
(f : R →+* S) :
f.ker ≤ ideal.comap f K
theorem
ideal.map_Inf
{R : Type u_1}
{S : Type u_2}
[comm_ring R]
[comm_ring S]
{A : set (ideal R)}
{f : R →+* S} :
function.surjective ⇑f → (∀ (J : ideal R), J ∈ A → f.ker ≤ J) → ideal.map f (has_Inf.Inf A) = has_Inf.Inf (ideal.map f '' A)
@[instance]
def
submodule.semimodule_submodule
{R : Type u}
{M : Type v}
[comm_ring R]
[add_comm_group M]
[module R M] :
semimodule (ideal R) (submodule R M)
Equations
- submodule.semimodule_submodule = {to_distrib_mul_action := {to_mul_action := {to_has_scalar := submodule.has_scalar' _inst_3, one_smul := _, mul_smul := _}, smul_add := _, smul_zero := _}, add_smul := _, zero_smul := _}
def
ring_hom.lift_of_surjective
{A : Type u_1}
{B : Type u_2}
{C : Type u_3}
[comm_ring A]
[comm_ring B]
[comm_ring C]
(f : A →+* B)
(hf : function.surjective ⇑f)
(g : A →+* C) :
lift_of_surjective f hf g hg
is the unique ring homomorphism φ
- such that
φ.comp f = g
(lift_of_surjective_comp
), - where
f : A →+* B
is surjective (hf
), - and
g : B →+* C
satisfieshg : f.ker ≤ g.ker
.
See lift_of_surjective_eq
for the uniqueness lemma.
A .
| \
f | \ g
| \
v \⌟
B ----> C
∃!φ
Equations
- f.lift_of_surjective hf g hg = {to_fun := λ (b : B), ⇑g (classical.some _), map_one' := _, map_mul' := _, map_zero' := _, map_add' := _}
@[simp]
theorem
ring_hom.lift_of_surjective_comp
{A : Type u_1}
{B : Type u_2}
{C : Type u_3}
[comm_ring A]
[comm_ring B]
[comm_ring C]
(f : A →+* B)
(hf : function.surjective ⇑f)
(g : A →+* C)
(hg : f.ker ≤ g.ker) :
(f.lift_of_surjective hf g hg).comp f = g