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

Commit 41d291c

Browse files
feat(algebra/big_operators/associated): generalize prod_primes_dvd (#12740)
1 parent 3ce4161 commit 41d291c

File tree

2 files changed

+43
-24
lines changed

2 files changed

+43
-24
lines changed

src/algebra/big_operators/associated.lean

Lines changed: 39 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -64,6 +64,45 @@ multiset.induction_on s (by simp [mt is_unit_iff_dvd_one.2 hp.not_unit])
6464
exact ⟨q, multiset.mem_cons.2 (or.inr hq₁), hq₂⟩ }
6565
end)
6666

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+
67106
namespace associates
68107

69108
section comm_monoid

src/number_theory/primorial.lean

Lines changed: 4 additions & 24 deletions
Original file line numberDiff line numberDiff line change
@@ -3,9 +3,10 @@ Copyright (c) 2020 Patrick Stevens. All rights reserved.
33
Released under Apache 2.0 license as described in the file LICENSE.
44
Authors: Patrick Stevens
55
-/
6-
import tactic.ring_exp
7-
import data.nat.parity
6+
import algebra.big_operators.associated
87
import data.nat.choose.sum
8+
import data.nat.parity
9+
import tactic.ring_exp
910

1011
/-!
1112
# Primorial
@@ -60,27 +61,6 @@ begin
6061
exacts [(s h).elim, (t h).elim],
6162
end
6263

63-
lemma prod_primes_dvd {s : finset ℕ} : ∀ (n : ℕ) (h : ∀ a ∈ s, nat.prime a) (div : ∀ a ∈ s, a ∣ n),
64-
(∏ p in s, p) ∣ n :=
65-
begin
66-
apply finset.induction_on s,
67-
{ simp, },
68-
{ intros a s a_not_in_s induct n primes divs,
69-
rw finset.prod_insert a_not_in_s,
70-
obtain ⟨k, rfl⟩ : a ∣ n := divs a (finset.mem_insert_self a s),
71-
apply mul_dvd_mul_left a,
72-
apply induct k,
73-
{ intros b b_in_s,
74-
exact primes b (finset.mem_insert_of_mem b_in_s), },
75-
{ intros b b_in_s,
76-
have b_div_n := divs b (finset.mem_insert_of_mem b_in_s),
77-
have a_prime := primes a (finset.mem_insert_self a s),
78-
have b_prime := primes b (finset.mem_insert_of_mem b_in_s),
79-
refine ((prime.dvd_mul b_prime).mp b_div_n).resolve_left (λ b_div_a, _),
80-
obtain rfl : b = a := ((nat.dvd_prime a_prime).1 b_div_a).resolve_left b_prime.ne_one,
81-
exact a_not_in_s b_in_s, } },
82-
end
83-
8464
lemma primorial_le_4_pow : ∀ (n : ℕ), n# ≤ 4 ^ n
8565
| 0 := le_rfl
8666
| 1 := le_of_inf_eq rfl
@@ -119,7 +99,7 @@ lemma primorial_le_4_pow : ∀ (n : ℕ), n# ≤ 4 ^ n
11999
have s : ∏ i in filter nat.prime (finset.Ico (m + 2) (2 * m + 2)),
120100
i ∣ choose (2 * m + 1) (m + 1),
121101
{ refine prod_primes_dvd (choose (2 * m + 1) (m + 1)) _ _,
122-
{ intros a, rw finset.mem_filter, apply and.right, },
102+
{ intros a, rw [finset.mem_filter, nat.prime_iff], apply and.right, },
123103
{ intros a, rw finset.mem_filter,
124104
intros pr,
125105
rcases pr with ⟨ size, is_prime ⟩,

0 commit comments

Comments
 (0)