@@ -392,7 +392,7 @@ variable {ι : Type*} {p : ι → Ideal R} {S : Finset ι}
392
392
variable (hp : (S : Set ι).Pairwise fun i j => p i ⊔ p j = ⊤)
393
393
394
394
-- Porting note: mem_iSup_finset_iff_exists_sum now requires DecidableEq ι
395
- theorem iSup_torsionBySet_ideal_eq_torsionBySet_iInf [DecidableEq ι] :
395
+ theorem iSup_torsionBySet_ideal_eq_torsionBySet_iInf :
396
396
⨆ i ∈ S, torsionBySet R M (p i) = torsionBySet R M ↑(⨅ i ∈ S, p i) := by
397
397
rcases S.eq_empty_or_nonempty with h | h
398
398
· simp only [h]
@@ -430,7 +430,7 @@ theorem iSup_torsionBySet_ideal_eq_torsionBySet_iInf [DecidableEq ι] :
430
430
#align submodule.supr_torsion_by_ideal_eq_torsion_by_infi Submodule.iSup_torsionBySet_ideal_eq_torsionBySet_iInf
431
431
432
432
-- Porting note: iSup_torsionBySet_ideal_eq_torsionBySet_iInf now requires DecidableEq ι
433
- theorem supIndep_torsionBySet_ideal [DecidableEq ι] : S.SupIndep fun i => torsionBySet R M <| p i :=
433
+ theorem supIndep_torsionBySet_ideal : S.SupIndep fun i => torsionBySet R M <| p i :=
434
434
fun T hT i hi hiT => by
435
435
rw [disjoint_iff, Finset.sup_eq_iSup,
436
436
iSup_torsionBySet_ideal_eq_torsionBySet_iInf fun i hi j hj ij => hp (hT hi) (hT hj) ij]
@@ -443,7 +443,7 @@ theorem supIndep_torsionBySet_ideal [DecidableEq ι] : S.SupIndep fun i => torsi
443
443
444
444
variable {q : ι → R} (hq : (S : Set ι).Pairwise <| (IsCoprime on q))
445
445
446
- theorem iSup_torsionBy_eq_torsionBy_prod [DecidableEq ι] :
446
+ theorem iSup_torsionBy_eq_torsionBy_prod :
447
447
⨆ i ∈ S, torsionBy R M (q i) = torsionBy R M (∏ i in S, q i) := by
448
448
rw [← torsionBySet_span_singleton_eq, Ideal.submodule_span_eq, ←
449
449
Ideal.finset_inf_span_singleton _ _ hq, Finset.inf_eq_iInf, ←
@@ -456,8 +456,7 @@ theorem iSup_torsionBy_eq_torsionBy_prod [DecidableEq ι] :
456
456
exact fun i hi j hj ij => (Ideal.sup_eq_top_iff_isCoprime _ _).mpr (hq hi hj ij)
457
457
#align submodule.supr_torsion_by_eq_torsion_by_prod Submodule.iSup_torsionBy_eq_torsionBy_prod
458
458
459
- -- Porting note: supIndep_torsionBySet_ideal now requires DecidableEq ι
460
- theorem supIndep_torsionBy [DecidableEq ι] : S.SupIndep fun i => torsionBy R M <| q i := by
459
+ theorem supIndep_torsionBy : S.SupIndep fun i => torsionBy R M <| q i := by
461
460
convert supIndep_torsionBySet_ideal (M := M) fun i hi j hj ij =>
462
461
(Ideal.sup_eq_top_iff_isCoprime (q i) _).mpr <| hq hi hj ij
463
462
exact (torsionBySet_span_singleton_eq (R := R) (M := M) _).symm
0 commit comments