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

Commit cdc3448

Browse files
committed
refactor(algebra/module/dedekind_domain): eliminate existentials (#14734)
This extracts some constructions from some long existentials so that we don't have to prove everything about the construction in a single place. This makes some of the statements longer, so in places the existential version is kept in terms of the definition one.
1 parent 9ac7c0c commit cdc3448

File tree

3 files changed

+60
-29
lines changed

3 files changed

+60
-29
lines changed

src/algebra/module/dedekind_domain.lean

Lines changed: 18 additions & 8 deletions
Original file line numberDiff line numberDiff line change
@@ -26,17 +26,17 @@ namespace submodule
2626
variables [is_dedekind_domain R]
2727
open unique_factorization_monoid
2828

29+
open_locale classical
30+
2931
/--Over a Dedekind domain, a `I`-torsion module is the internal direct sum of its `p i ^ e i`-
3032
torsion submodules, where `I = ∏ i, p i ^ e i` is its unique decomposition in prime ideals.-/
3133
lemma is_internal_prime_power_torsion_of_is_torsion_by_ideal {I : ideal R} (hI : I ≠ ⊥)
3234
(hM : module.is_torsion_by_set R M I) :
33-
∃ (P : finset $ ideal R) [decidable_eq P] [∀ p ∈ P, prime p] (e : P → ℕ),
34-
by exactI direct_sum.is_internal (λ p : P, torsion_by_set R M (p ^ e p : ideal R)) :=
35+
direct_sum.is_internal (λ p : (factors I).to_finset,
36+
torsion_by_set R M (p ^ (factors I).count p : ideal R)) :=
3537
begin
36-
classical,
3738
let P := factors I,
3839
have prime_of_mem := λ p (hp : p ∈ P.to_finset), prime_of_factor p (multiset.mem_to_finset.mp hp),
39-
refine ⟨P.to_finset, infer_instance, prime_of_mem, λ i, P.count i, _⟩,
4040
apply @torsion_by_set_is_internal _ _ _ _ _ _ _ _ (λ p, p ^ P.count p) _,
4141
{ convert hM,
4242
rw [← finset.inf_eq_infi, is_dedekind_domain.inf_prime_pow_eq_prod,
@@ -55,15 +55,25 @@ begin
5555
end
5656

5757
/--A finitely generated torsion module over a Dedekind domain is an internal direct sum of its
58-
`p i ^ e i`-torsion submodules for some prime ideals `p i` and numbers `e i`.-/
58+
`p i ^ e i`-torsion submodules where `p i` are factors of `(⊤ : submodule R M).annihilator` and
59+
`e i` are their multiplicities. -/
5960
theorem is_internal_prime_power_torsion [module.finite R M] (hM : module.is_torsion R M) :
60-
∃ (P : finset $ ideal R) [decidable_eq P] [∀ p ∈ P, prime p] (e : P → ℕ),
61-
by exactI direct_sum.is_internal (λ p : P, torsion_by_set R M (p ^ e p : ideal R)) :=
61+
direct_sum.is_internal (λ p : (factors (⊤ : submodule R M).annihilator).to_finset,
62+
torsion_by_set R M (p ^ (factors (⊤ : submodule R M).annihilator).count p : ideal R)) :=
6263
begin
63-
obtain ⟨I, hI, hM'⟩ := is_torsion_by_ideal_of_finite_of_is_torsion hM,
64+
have hM' := module.is_torsion_by_set_annihilator_top R M,
65+
have hI := submodule.annihilator_top_inter_non_zero_divisors hM,
6466
refine is_internal_prime_power_torsion_of_is_torsion_by_ideal _ hM',
6567
rw ←set.nonempty_iff_ne_empty at hI, rw submodule.ne_bot_iff,
6668
obtain ⟨x, H, hx⟩ := hI, exact ⟨x, H, non_zero_divisors.ne_zero hx⟩
6769
end
6870

71+
/--A finitely generated torsion module over a Dedekind domain is an internal direct sum of its
72+
`p i ^ e i`-torsion submodules for some prime ideals `p i` and numbers `e i`.-/
73+
theorem exists_is_internal_prime_power_torsion [module.finite R M] (hM : module.is_torsion R M) :
74+
∃ (P : finset $ ideal R) [decidable_eq P] [∀ p ∈ P, prime p] (e : P → ℕ),
75+
by exactI direct_sum.is_internal (λ p : P, torsion_by_set R M (p ^ e p : ideal R)) :=
76+
⟨_, _, λ p hp, prime_of_factor p (multiset.mem_to_finset.mp hp), _,
77+
is_internal_prime_power_torsion hM⟩
78+
6979
end submodule

src/algebra/module/pid.lean

Lines changed: 24 additions & 9 deletions
Original file line numberDiff line numberDiff line change
@@ -47,7 +47,7 @@ Finitely generated module, principal ideal domain, classification, structure the
4747
-/
4848

4949
universes u v
50-
open_locale big_operators
50+
open_locale big_operators classical
5151

5252
variables {R : Type u} [comm_ring R] [is_domain R] [is_principal_ideal_ring R]
5353
variables {M : Type v} [add_comm_group M] [module R M]
@@ -56,21 +56,36 @@ variables {N : Type (max u v)} [add_comm_group N] [module R N]
5656
open_locale direct_sum
5757
open submodule
5858

59+
open unique_factorization_monoid
60+
5961
/--A finitely generated torsion module over a PID is an internal direct sum of its
6062
`p i ^ e i`-torsion submodules for some primes `p i` and numbers `e i`.-/
6163
theorem submodule.is_internal_prime_power_torsion_of_pid
64+
[module.finite R M] (hM : module.is_torsion R M) :
65+
direct_sum.is_internal (λ p : (factors (⊤ : submodule R M).annihilator).to_finset,
66+
torsion_by R M
67+
(is_principal.generator (p : ideal R)
68+
^ (factors (⊤ : submodule R M).annihilator).count p)) :=
69+
begin
70+
convert is_internal_prime_power_torsion hM,
71+
ext p : 1,
72+
rw [← torsion_by_span_singleton_eq, ideal.submodule_span_eq, ← ideal.span_singleton_pow,
73+
ideal.span_singleton_generator],
74+
end
75+
76+
/--A finitely generated torsion module over a PID is an internal direct sum of its
77+
`p i ^ e i`-torsion submodules for some primes `p i` and numbers `e i`.-/
78+
theorem submodule.exists_is_internal_prime_power_torsion_of_pid
6279
[module.finite R M] (hM : module.is_torsion R M) :
6380
∃ (ι : Type u) [fintype ι] [decidable_eq ι] (p : ι → R) (h : ∀ i, irreducible $ p i) (e : ι → ℕ),
6481
by exactI direct_sum.is_internal (λ i, torsion_by R M $ p i ^ e i) :=
6582
begin
66-
obtain ⟨P, dec, hP, e, this⟩ := is_internal_prime_power_torsion hM,
67-
refine ⟨P, infer_instance, dec, λ p, is_principal.generator (p : ideal R), _, e, _⟩,
83+
refine ⟨_, _, _, _, _, _, submodule.is_internal_prime_power_torsion_of_pid hM,
84+
exact finset.fintype_coe_sort _,
6885
{ rintro ⟨p, hp⟩,
69-
haveI := ideal.is_prime_of_prime (hP p hp),
70-
exact (is_principal.prime_generator_of_is_prime p (hP p hp).ne_zero).irreducible },
71-
{ convert this, ext p : 1,
72-
rw [← torsion_by_span_singleton_eq, ideal.submodule_span_eq, ← ideal.span_singleton_pow,
73-
ideal.span_singleton_generator] }
86+
have hP := prime_of_factor p (multiset.mem_to_finset.mp hp),
87+
haveI := ideal.is_prime_of_prime hP,
88+
exact (is_principal.prime_generator_of_is_prime p hP.ne_zero).irreducible },
7489
end
7590

7691
namespace module
@@ -205,7 +220,7 @@ theorem equiv_direct_sum_of_is_torsion [h' : module.finite R N] (hN : module.is_
205220
∃ (ι : Type u) [fintype ι] (p : ι → R) (h : ∀ i, irreducible $ p i) (e : ι → ℕ),
206221
nonempty $ N ≃ₗ[R] ⨁ (i : ι), R ⧸ R ∙ (p i ^ e i) :=
207222
begin
208-
obtain ⟨I, fI, _, p, hp, e, h⟩ := submodule.is_internal_prime_power_torsion_of_pid hN,
223+
obtain ⟨I, fI, _, p, hp, e, h⟩ := submodule.exists_is_internal_prime_power_torsion_of_pid hN,
209224
haveI := fI,
210225
have : ∀ i, ∃ (d : ℕ) (k : fin d → ℕ),
211226
nonempty $ torsion_by R N (p i ^ e i) ≃ₗ[R] ⨁ j, R ⧸ R ∙ (p i ^ k j),

src/algebra/module/torsion.lean

Lines changed: 18 additions & 12 deletions
Original file line numberDiff line numberDiff line change
@@ -496,19 +496,25 @@ section torsion
496496
variables [comm_semiring R] [add_comm_monoid M] [module R M]
497497
open_locale big_operators
498498

499-
lemma is_torsion_by_ideal_of_finite_of_is_torsion [module.finite R M] (hM : module.is_torsion R M) :
500-
∃ I : ideal R, (I : set R) ∩ R⁰ ≠ ∅ ∧ module.is_torsion_by_set R M I :=
499+
variables (R M)
500+
501+
lemma _root_.module.is_torsion_by_set_annihilator_top :
502+
module.is_torsion_by_set R M (⊤ : submodule R M).annihilator :=
503+
λ x ha, mem_annihilator.mp ha.prop x mem_top
504+
505+
variables {R M}
506+
507+
lemma _root_.submodule.annihilator_top_inter_non_zero_divisors [module.finite R M]
508+
(hM : module.is_torsion R M) :
509+
((⊤ : submodule R M).annihilator : set R) ∩ R⁰ ≠ ∅:=
501510
begin
502-
cases (module.finite_def.mp infer_instance : (⊤ : submodule R M).fg) with S h,
503-
refine ⟨∏ x in S, ideal.torsion_of R M x, _, _⟩,
504-
{ refine set.nonempty.ne_empty ⟨_, _, (∏ x in S, (@hM x).some : R⁰).2⟩,
505-
rw [subtype.val_eq_coe, submonoid.coe_finset_prod],
506-
apply ideal.prod_mem_prod,
507-
exact λ x _, (@hM x).some_spec },
508-
{ rw [module.is_torsion_by_set_iff_torsion_by_set_eq_top, eq_top_iff, ← h, span_le],
509-
intros x hx, apply torsion_by_set_le_torsion_by_set_of_subset,
510-
{ apply ideal.le_of_dvd, exact finset.dvd_prod_of_mem _ hx },
511-
{ rw mem_torsion_by_set_iff, rintro ⟨a, ha⟩, exact ha } }
511+
obtain ⟨S, hS⟩ := ‹module.finite R M›.out,
512+
refine set.nonempty.ne_empty ⟨_, _, (∏ x in S, (@hM x).some : R⁰).prop⟩,
513+
rw [submonoid.coe_finset_prod, set_like.mem_coe, ←hS, mem_annihilator_span],
514+
intro n,
515+
letI := classical.dec_eq M,
516+
rw [←finset.prod_erase_mul _ _ n.prop, mul_smul, ←submonoid.smul_def, (@hM n).some_spec,
517+
smul_zero],
512518
end
513519

514520
variables [no_zero_divisors R] [nontrivial R]

0 commit comments

Comments
 (0)