@@ -64,6 +64,45 @@ multiset.induction_on s (by simp [mt is_unit_iff_dvd_one.2 hp.not_unit])
64
64
exact ⟨q, multiset.mem_cons.2 (or.inr hq₁), hq₂⟩ }
65
65
end )
66
66
67
+ lemma multiset.prod_primes_dvd
68
+ [cancel_comm_monoid_with_zero α] [Π a : α, decidable_pred (associated a)]
69
+ {s : multiset α} (n : α) (h : ∀ a ∈ s, prime a) (div : ∀ a ∈ s, a ∣ n)
70
+ (uniq : ∀ a, s.countp (associated a) ≤ 1 ) :
71
+ s.prod ∣ n :=
72
+ begin
73
+ induction s using multiset.induction_on with a s induct n primes divs generalizing n,
74
+ { simp only [multiset.prod_zero, one_dvd] },
75
+ { rw multiset.prod_cons,
76
+ obtain ⟨k, rfl⟩ : a ∣ n := div a (multiset.mem_cons_self a s),
77
+ apply mul_dvd_mul_left a,
78
+ refine induct
79
+ (λ a ha, h a (multiset.mem_cons_of_mem ha))
80
+ (λ a, (multiset.countp_le_of_le _ (multiset.le_cons_self _ _)).trans (uniq a))
81
+ k (λ b b_in_s, _),
82
+ { have b_div_n := div b (multiset.mem_cons_of_mem b_in_s),
83
+ have a_prime := h a (multiset.mem_cons_self a s),
84
+ have b_prime := h b (multiset.mem_cons_of_mem b_in_s),
85
+ refine (b_prime.dvd_or_dvd b_div_n).resolve_left (λ b_div_a, _),
86
+ have assoc := b_prime.associated_of_dvd a_prime b_div_a,
87
+ have := uniq a,
88
+ rw [multiset.countp_cons_of_pos _ (associated.refl _), nat.succ_le_succ_iff, ←not_lt,
89
+ multiset.countp_pos] at this ,
90
+ exact this ⟨b, b_in_s, assoc.symm⟩ } }
91
+ end
92
+
93
+ lemma finset.prod_primes_dvd
94
+ [cancel_comm_monoid_with_zero α] [unique αˣ]
95
+ {s : finset α} (n : α) (h : ∀ a ∈ s, prime a) (div : ∀ a ∈ s, a ∣ n) :
96
+ (∏ p in s, p) ∣ n :=
97
+ begin
98
+ classical,
99
+ exact multiset.prod_primes_dvd n
100
+ (by simpa only [multiset.map_id', finset.mem_def] using h)
101
+ (by simpa only [multiset.map_id', finset.mem_def] using div)
102
+ (by simp only [multiset.map_id', associated_eq_eq, multiset.countp_eq_card_filter,
103
+ ←multiset.count_eq_card_filter_eq, ←multiset.nodup_iff_count_le_one, s.nodup]),
104
+ end
105
+
67
106
namespace associates
68
107
69
108
section comm_monoid
0 commit comments