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

Commit ac283c2

Browse files
feat(data/nat/prime): some lemmas about prime factors (#10385)
A few small lemmas about prime factors, for use in future PRs on prime factorisations and the Euler product formula for totient
1 parent eb8b1b8 commit ac283c2

File tree

1 file changed

+63
-0
lines changed

1 file changed

+63
-0
lines changed

src/data/nat/prime.lean

Lines changed: 63 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -704,6 +704,29 @@ begin
704704
rwa eq_of_mem_repeat hq },
705705
end
706706

707+
/-- For positive `a` and `b`, the prime factors of `a * b` are the union of those of `a` and `b` -/
708+
lemma perm_factors_mul_of_pos {a b : ℕ} (ha : 0 < a) (hb : 0 < b) :
709+
(a * b).factors ~ a.factors ++ b.factors :=
710+
begin
711+
refine (factors_unique _ _).symm,
712+
{ rw [list.prod_append, prod_factors ha, prod_factors hb] },
713+
{ intros p hp,
714+
rw list.mem_append at hp,
715+
cases hp;
716+
exact prime_of_mem_factors hp },
717+
end
718+
719+
/-- For coprime `a` and `b`, the prime factors of `a * b` are the union of those of `a` and `b` -/
720+
lemma perm_factors_mul_of_coprime {a b : ℕ} (hab : coprime a b) :
721+
(a * b).factors ~ a.factors ++ b.factors :=
722+
begin
723+
rcases a.eq_zero_or_pos with rfl | ha,
724+
{ simp [(coprime_zero_left _).mp hab] },
725+
rcases b.eq_zero_or_pos with rfl | hb,
726+
{ simp [(coprime_zero_right _).mp hab] },
727+
exact perm_factors_mul_of_pos ha hb,
728+
end
729+
707730
end
708731

709732
lemma succ_dvd_or_succ_dvd_of_succ_sum_dvd_mul {p : ℕ} (p_prime : prime p) {m n k l : ℕ}
@@ -1022,4 +1045,44 @@ begin
10221045
rw [mem_factors_mul_of_pos ha hb p, list.mem_union]
10231046
end
10241047

1048+
1049+
open list
1050+
1051+
/-- For `b > 0`, the power of `p` in `a * b` is at least that in `a` -/
1052+
lemma le_factors_count_mul_left {p a b : ℕ} (hb : 0 < b) :
1053+
list.count p a.factors ≤ list.count p (a * b).factors :=
1054+
begin
1055+
rcases a.eq_zero_or_pos with rfl | ha,
1056+
{ simp },
1057+
{ rw [perm.count_eq (perm_factors_mul_of_pos ha hb) p, count_append p], simp },
1058+
end
1059+
1060+
/-- For `a > 0`, the power of `p` in `a * b` is at least that in `b` -/
1061+
lemma le_factors_count_mul_right {p a b : ℕ} (ha : 0 < a) :
1062+
list.count p b.factors ≤ list.count p (a * b).factors :=
1063+
by { rw mul_comm, apply le_factors_count_mul_left ha }
1064+
1065+
/-- If `p` is a prime factor of `a` then `p` is also a prime factor of `a * b` for any `b > 0` -/
1066+
lemma mem_factors_mul_left {p a b : ℕ} (hpa : p ∈ a.factors) (hb : 0 < b) : p ∈ (a*b).factors :=
1067+
by { rw ←list.count_pos, exact gt_of_ge_of_gt (le_factors_count_mul_left hb) (count_pos.mpr hpa) }
1068+
1069+
/-- If `p` is a prime factor of `b` then `p` is also a prime factor of `a * b` for any `a > 0` -/
1070+
lemma mem_factors_mul_right {p a b : ℕ} (hpb : p ∈ b.factors) (ha : 0 < a) : p ∈ (a*b).factors :=
1071+
by { rw mul_comm, exact mem_factors_mul_left hpb ha }
1072+
1073+
/-- If `p` is a prime factor of `a` then the power of `p` in `a` is the same that in `a * b`,
1074+
for any `b` coprime to `a`. -/
1075+
lemma factors_count_eq_of_coprime_left {p a b : ℕ} (hab : coprime a b) (hpa : p ∈ a.factors) :
1076+
list.count p (a * b).factors = list.count p a.factors :=
1077+
begin
1078+
rw [perm.count_eq (perm_factors_mul_of_coprime hab) p, count_append],
1079+
simpa only [count_eq_zero_of_not_mem (coprime_factors_disjoint hab hpa)],
1080+
end
1081+
1082+
/-- If `p` is a prime factor of `b` then the power of `p` in `b` is the same that in `a * b`,
1083+
for any `a` coprime to `b`. -/
1084+
lemma factors_count_eq_of_coprime_right {p a b : ℕ} (hab : coprime a b) (hpb : p ∈ b.factors) :
1085+
list.count p (a * b).factors = list.count p b.factors :=
1086+
by { rw mul_comm, exact factors_count_eq_of_coprime_left (coprime_comm.mp hab) hpb }
1087+
10251088
end nat

0 commit comments

Comments
 (0)