Skip to content
This repository was archived by the owner on Jul 24, 2024. It is now read-only.

Commit 6ee6203

Browse files
feat(counterexample) : a homogeneous ideal that is not prime but homogeneously prime (#12485)
For graded rings, if the indexing set is nice, then a homogeneous ideal `I` is prime if and only if it is homogeneously prime, i.e. product of two homogeneous elements being in `I` implying at least one of them is in `I`. This fact is in `src/ring_theory/graded_algebra/radical.lean`. This counter example is meant to show that "nice condition" at least needs to include cancellation property by exhibiting an ideal in Zmod4^2 graded by a two element set being homogeneously prime but not prime. In #12277, Eric suggests that this counter example is worth pr-ing, so here is the pr. Co-authored-by: Eric Wieser <wieser.eric@gmail.com>
1 parent 2e7483d commit 6ee6203

File tree

2 files changed

+154
-2
lines changed

2 files changed

+154
-2
lines changed
Lines changed: 151 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,151 @@
1+
/-
2+
Copyright (c) 2022 Jujian Zhang. All rights reserved.
3+
Released under Apache 2.0 license as described in the file LICENSE.
4+
Authors: Johan Commelin, Eric Wieser, Jujian Zhang
5+
-/
6+
import ring_theory.graded_algebra.homogeneous_ideal
7+
import data.zmod.basic
8+
import tactic.derive_fintype
9+
10+
/-!
11+
# A homogeneous prime that is homogeneously prime but not prime
12+
13+
In `src/ring_theory/graded_algebra/radical.lean`, we assumed that the underline grading is inedxed
14+
by a `linear_ordered_cancel_add_comm_monoid` to prove that a homogeneous ideal is prime if and only
15+
if it is homogeneously prime. This file is aimed to show that even if this assumption isn't strictly
16+
necessary, the assumption of "being cancellative" is. We construct a counterexample where the
17+
underlying indexing set is a `linear_ordered_add_comm_monoid` but is not cancellative and the
18+
statement is false.
19+
20+
We achieve this by considering the ring `R=ℤ/4ℤ`. We first give the two element set `ι = {0, 1}` a
21+
structure of linear ordered additive commutative monoid by setting `0 + 0 = 0` and `_ + _ = 1` and
22+
`0 < 1`. Then we use `ι` to grade `R²` by setting `{(a, a) | a ∈ R}` to have grade `0`; and
23+
`{(0, b) | b ∈ R}` to have grade 1. Then the ideal `I = span {(0, 2)} ⊆ ℤ/4ℤ` is homogeneous and not
24+
prime. But it is homogeneously prime, i.e. if `(a, b), (c, d)` are two homogeneous elements then
25+
`(a, b) * (c, d) ∈ I` implies either `(a, b) ∈ I` or `(c, d) ∈ I`.
26+
27+
28+
## Tags
29+
30+
homogeneous, prime
31+
-/
32+
33+
namespace counterexample_not_prime_but_homogeneous_prime
34+
35+
open direct_sum
36+
local attribute [reducible] with_zero
37+
38+
abbreviation two := with_zero unit
39+
40+
instance : linear_ordered_add_comm_monoid two :=
41+
{ add_le_add_left := by delta two with_zero; dec_trivial,
42+
..(_ : linear_order two),
43+
..(_ : add_comm_monoid two)}
44+
45+
section
46+
47+
variables (R : Type*) [comm_ring R]
48+
49+
/-- The grade 0 part of `R²` is `{(a, a) | a ∈ R}`. -/
50+
def submodule_z : submodule R (R × R) :=
51+
{ carrier := { zz | zz.1 = zz.2 },
52+
zero_mem' := rfl,
53+
add_mem' := λ a b ha hb, congr_arg2 (+) ha hb,
54+
smul_mem' := λ a b hb, congr_arg ((*) a) hb }
55+
56+
/-- The grade 1 part of `R²` is `{(0, b) | b ∈ R}`. -/
57+
def submodule_o : submodule R (R × R) := (linear_map.fst R R R).ker
58+
59+
/-- Given the above grading (see `submodule_z` and `submodule_o`),
60+
we turn `R²` into a graded ring. -/
61+
def grading : two → submodule R (R × R)
62+
| 0 := submodule_z R
63+
| 1 := submodule_o R
64+
65+
lemma grading.one_mem : (1 : (R × R)) ∈ grading R 0 :=
66+
eq.refl (1, 1).fst
67+
68+
lemma grading.mul_mem : ∀ ⦃i j : two⦄ {a b : (R × R)} (ha : a ∈ grading R i) (hb : b ∈ grading R j),
69+
a * b ∈ grading R (i + j)
70+
| 0 0 a b (ha : a.1 = a.2) (hb : b.1 = b.2) := show a.1 * b.1 = a.2 * b.2, by rw [ha, hb]
71+
| 0 1 a b (ha : a.1 = a.2) (hb : b.1 = 0) := show a.1 * b.1 = 0, by rw [hb, mul_zero]
72+
| 1 0 a b (ha : a.1 = 0) hb := show a.1 * b.1 = 0, by rw [ha, zero_mul]
73+
| 1 1 a b (ha : a.1 = 0) hb := show a.1 * b.1 = 0, by rw [ha, zero_mul]
74+
75+
end
76+
77+
notation `R` := zmod 4
78+
79+
/-- `R² ≅ {(a, a) | a ∈ R} ⨁ {(0, b) | b ∈ R}` by `(x, y) ↦ (x, x) + (0, y - x)`. -/
80+
def grading.decompose : (R × R) →+ direct_sum two (λ i, grading R i) :=
81+
{ to_fun := λ zz, of (λ i, grading R i) 0 ⟨(zz.1, zz.1), rfl⟩ +
82+
of (λ i, grading R i) 1 ⟨(0, zz.2 - zz.1), rfl⟩,
83+
map_zero' := by { ext1 (_|⟨⟨⟩⟩); refl },
84+
map_add' := begin
85+
rintros ⟨a1, b1⟩ ⟨a2, b2⟩,
86+
have H : b1 + b2 - (a1 + a2) = b1 - a1 + (b2 - a2), by abel,
87+
ext (_|⟨⟨⟩⟩) : 3;
88+
simp only [prod.fst_add, prod.snd_add, add_apply, submodule.coe_add, prod.mk_add_mk, H];
89+
repeat { erw of_eq_same }; repeat { erw of_eq_of_ne }; try { apply option.no_confusion };
90+
dsimp; simp only [zero_add, add_zero]; refl,
91+
end }
92+
93+
lemma grading.left_inv :
94+
function.left_inverse grading.decompose (submodule_coe (grading R)) := λ zz,
95+
begin
96+
induction zz using direct_sum.induction_on with i zz d1 d2 ih1 ih2,
97+
{ simp only [map_zero],},
98+
{ rcases i with (_|⟨⟨⟩⟩); rcases zz with ⟨⟨a, b⟩, (hab : _ = _)⟩;
99+
dsimp at hab; cases hab; dec_trivial! },
100+
{ simp only [map_add, ih1, ih2], },
101+
end
102+
103+
lemma grading.right_inv :
104+
function.right_inverse grading.decompose (submodule_coe (grading R)) := λ zz,
105+
begin
106+
cases zz with a b,
107+
unfold grading.decompose,
108+
simp only [add_monoid_hom.coe_mk, map_add, submodule_coe_of, subtype.coe_mk, prod.mk_add_mk,
109+
add_zero, add_sub_cancel'_right],
110+
end
111+
112+
instance : graded_algebra (grading R) :=
113+
{ one_mem := grading.one_mem R,
114+
mul_mem := grading.mul_mem R,
115+
decompose' := grading.decompose,
116+
left_inv := by { convert grading.left_inv, },
117+
right_inv := by { convert grading.right_inv, } }
118+
119+
/-- The counterexample is the ideal `I = span {(2, 2)}`. -/
120+
def I : ideal (R × R) := ideal.span {((2, 2) : (R × R))}.
121+
122+
lemma I_not_prime : ¬ I.is_prime :=
123+
begin
124+
rintro ⟨rid1, rid2⟩,
125+
apply rid1, clear rid1, revert rid2,
126+
simp only [I, ideal.mem_span_singleton, ideal.eq_top_iff_one],
127+
dec_trivial,
128+
end
129+
130+
lemma I_is_homogeneous : I.is_homogeneous (grading R) :=
131+
begin
132+
rw ideal.is_homogeneous.iff_exists,
133+
refine ⟨{⟨(2, 2), ⟨0, rfl⟩⟩}, _⟩,
134+
rw set.image_singleton,
135+
refl,
136+
end
137+
138+
lemma homogeneous_mem_or_mem {x y : (R × R)} (hx : set_like.is_homogeneous (grading R) x)
139+
(hy : set_like.is_homogeneous (grading R) y)
140+
(hxy : x * y ∈ I) : x ∈ I ∨ y ∈ I :=
141+
begin
142+
simp only [I, ideal.mem_span_singleton] at hxy ⊢,
143+
cases x, cases y,
144+
obtain ⟨(_|⟨⟨⟩⟩), hx : _ = _⟩ := hx;
145+
obtain ⟨(_|⟨⟨⟩⟩), hy : _ = _⟩ := hy;
146+
dsimp at hx hy;
147+
cases hx; cases hy; clear hx hy;
148+
dec_trivial!,
149+
end
150+
151+
end counterexample_not_prime_but_homogeneous_prime

src/ring_theory/graded_algebra/radical.lean

Lines changed: 3 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -25,8 +25,9 @@ This file contains a proof that the radical of any homogeneous ideal is a homoge
2525
## Implementation details
2626
2727
Throughout this file, the indexing type `ι` of grading is assumed to be a
28-
`linear_ordered_cancel_add_comm_monoid`. This might be stronger than necessary and `linarith`
29-
does not work on `linear_ordered_cancel_add_comm_monoid`.
28+
`linear_ordered_cancel_add_comm_monoid`. This might be stronger than necessary but cancelling
29+
property is strictly necessary; for a counterexample of how `ideal.is_homogeneous.is_prime_iff`
30+
fails for a non-cancellative set see `counterexample/homogeneous_prime_not_prime.lean`.
3031
3132
## Tags
3233

0 commit comments

Comments
 (0)