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

Commit ccf3e37

Browse files
committed
feat(ring_theory/unique_factorization_domain): alternative specification for count (normalized_factors x) (#13161)
`count_normalized_factors_eq` specifies the number of times an irreducible factor `p` appears in `normalized_factors x`, namely the number of times it divides `x`. This is often easier to work with than going through `multiplicity` since this way we avoid coercing to `enat`.
1 parent 0c424e9 commit ccf3e37

File tree

1 file changed

+27
-1
lines changed

1 file changed

+27
-1
lines changed

src/ring_theory/unique_factorization_domain.lean

Lines changed: 27 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -695,9 +695,10 @@ in ⟨a', b', c', λ _ hpb hpa, no_factor hpa hpb, ha, hb⟩
695695

696696
section multiplicity
697697
variables [nontrivial R] [normalization_monoid R] [decidable_eq R]
698-
variables [decidable_rel (has_dvd.dvd : R → R → Prop)]
698+
variables [dec_dvd : decidable_rel (has_dvd.dvd : R → R → Prop)]
699699
open multiplicity multiset
700700

701+
include dec_dvd
701702
lemma le_multiplicity_iff_repeat_le_normalized_factors {a b : R} {n : ℕ}
702703
(ha : irreducible a) (hb : b ≠ 0) :
703704
↑n ≤ multiplicity a b ↔ repeat (normalize a) n ≤ normalized_factors b :=
@@ -718,6 +719,12 @@ begin
718719
exact (associated.pow_pow $ associated_normalize a).dvd.trans (dvd.intro u.prod rfl) }
719720
end
720721

722+
/-- The multiplicity of an irreducible factor of a nonzero element is exactly the number of times
723+
the normalized factor occurs in the `normalized_factors`.
724+
725+
See also `count_normalized_factors_eq` which expands the definition of `multiplicity`
726+
to produce a specification for `count (normalized_factors _) _`..
727+
-/
721728
lemma multiplicity_eq_count_normalized_factors {a b : R} (ha : irreducible a) (hb : b ≠ 0) :
722729
multiplicity a b = (normalized_factors b).count (normalize a) :=
723730
begin
@@ -729,6 +736,25 @@ begin
729736
rw [le_multiplicity_iff_repeat_le_normalized_factors ha hb, ← le_count_iff_repeat_le],
730737
end
731738

739+
omit dec_dvd
740+
/-- The number of times an irreducible factor `p` appears in `normalized_factors x` is defined by
741+
the number of times it divides `x`.
742+
743+
See also `multiplicity_eq_count_normalized_factors` if `n` is given by `multiplicity p x`.
744+
-/
745+
lemma count_normalized_factors_eq {p x : R} (hp : irreducible p) (hnorm : normalize p = p) {n : ℕ}
746+
(hle : p^n ∣ x) (hlt : ¬ (p^(n+1) ∣ x)) :
747+
(normalized_factors x).count p = n :=
748+
begin
749+
letI : decidable_rel ((∣) : R → R → Prop) := λ _ _, classical.prop_decidable _,
750+
by_cases hx0 : x = 0,
751+
{ simp [hx0] at hlt, contradiction },
752+
rw [← enat.coe_inj],
753+
convert (multiplicity_eq_count_normalized_factors hp hx0).symm,
754+
{ exact hnorm.symm },
755+
exact (multiplicity.eq_coe_iff.mpr ⟨hle, hlt⟩).symm
756+
end
757+
732758
end multiplicity
733759

734760
end unique_factorization_monoid

0 commit comments

Comments
 (0)