Skip to content
New issue

Have a question about this project? Sign up for a free GitHub account to open an issue and contact its maintainers and the community.

By clicking “Sign up for GitHub”, you agree to our terms of service and privacy statement. We’ll occasionally send you account related emails.

Already on GitHub? Sign in to your account

[Merged by Bors] - refactor(data/nat/factorization): Change definition of factorization to be computable #12301

Closed
wants to merge 62 commits into from
Closed
Show file tree
Hide file tree
Changes from 45 commits
Commits
Show all changes
62 commits
Select commit Hold shift + click to select a range
bacc815
initial
BoltonBailey Feb 26, 2022
42ee7fe
finish proof
BoltonBailey Feb 26, 2022
5c703d9
added lemma to part
BoltonBailey Feb 26, 2022
c1bd8eb
minor change
BoltonBailey Feb 26, 2022
2930c99
changes
BoltonBailey Mar 1, 2022
a74a858
Merge branch 'master' of github.com:leanprover-community/mathlib into…
BoltonBailey Mar 1, 2022
a88213a
change lemma
BoltonBailey Mar 1, 2022
c3b32f9
delete lemma
BoltonBailey Mar 1, 2022
45cc4cb
Merge branch 'master' of github.com:leanprover-community/mathlib into…
BoltonBailey Apr 21, 2022
8641730
resolve sorry
BoltonBailey Apr 21, 2022
240ac77
Delete part.lean
BoltonBailey Apr 21, 2022
8147e21
Merge branch 'master' of github.com:leanprover-community/mathlib into…
BoltonBailey Apr 23, 2022
325001f
uncomment lemma
BoltonBailey Apr 23, 2022
46730dc
uncomment import
BoltonBailey Apr 23, 2022
a64267a
added changes
BoltonBailey Apr 25, 2022
e78a4c5
Merge branch 'master' of github.com:leanprover-community/mathlib into…
BoltonBailey Apr 27, 2022
a0f5888
exploartory lemmas
BoltonBailey Apr 27, 2022
04837e9
Merge branch 'master' into BoltonBailey/computable-factorization
stuart-presnell Jun 25, 2022
e451b10
Remove imports of `factorization/basic`
stuart-presnell Jun 25, 2022
38df0a6
Golf `mem_support_to_fun`
stuart-presnell Jun 25, 2022
7a5d894
Golf `multiplicity_eq_factorization`
stuart-presnell Jun 25, 2022
37ff88c
Golf `padic_val_nat_eq_factorization`
stuart-presnell Jun 25, 2022
032699b
Update src/data/nat/factorization.lean
BoltonBailey Jun 25, 2022
fe7dd4a
Move `dvd_iff_padic_val_nat_ne_zero`
stuart-presnell Jun 25, 2022
38f529d
Merge branch 'BoltonBailey/computable-factorization' of https://githu…
stuart-presnell Jun 25, 2022
f7b3676
Update imports to `data.nat.totient`
stuart-presnell Jun 25, 2022
768aab0
Add lemma `factorization_def`
stuart-presnell Jun 25, 2022
b7389e1
Delete `noncomputable`
stuart-presnell Jun 25, 2022
18b8edb
Fix `support_factorization`
stuart-presnell Jun 25, 2022
acbc9cf
Fix `factorization_zero` and `factorization_one`
stuart-presnell Jun 26, 2022
d22c21c
Fix lemmas with `sorry` in `factors_count_eq`
stuart-presnell Jun 26, 2022
1424438
Prove `factors_count_eq`
stuart-presnell Jun 27, 2022
440e4c7
Move and golf `prod_pow_prime_padic_val_nat`
stuart-presnell Jun 27, 2022
014aa96
Move and prove `eq_iff_prime_padic_val_nat_eq`
stuart-presnell Jun 27, 2022
1c6dd00
Update basic.lean
stuart-presnell Jun 27, 2022
6aae01c
Fix imports
stuart-presnell Jun 27, 2022
d8d27ec
Merge branch 'master' into BoltonBailey/computable-factorization
stuart-presnell Jun 27, 2022
b5d1096
Delete `multiplicity_eq_factorization`
stuart-presnell Jun 28, 2022
6fcc3ed
Move `prod_factors_gcd_mul_prod_factors_mul`
stuart-presnell Jun 28, 2022
9ee17fe
Merge branch 'master' into BoltonBailey/computable-factorization
stuart-presnell Jun 28, 2022
59affc4
Restore `padic_val` import
stuart-presnell Jun 28, 2022
c99e1be
Add `algebra.char_p.two` import
stuart-presnell Jun 28, 2022
669cfc0
Fix imports in `group_theory/exponent`
stuart-presnell Jun 28, 2022
01cc590
Import in `analysis/special_functions/log/basic`
stuart-presnell Jun 28, 2022
385651e
Delete `padic_val_nat_eq_factorization`
stuart-presnell Jun 28, 2022
c02c287
Fix nonterminal simps
stuart-presnell Jun 28, 2022
9bfdcac
Merge branch 'master' into BoltonBailey/computable-factorization
stuart-presnell Jun 30, 2022
ecad72a
Merge branch 'master' into BoltonBailey/computable-factorization
stuart-presnell Jul 2, 2022
d3960c7
Merge branch 'master' into BoltonBailey/computable-factorization
stuart-presnell Jul 15, 2022
04fc75e
Change `enat` to `part_enat`
stuart-presnell Jul 17, 2022
37891c1
Update src/data/nat/factorization/basic.lean
BoltonBailey Jul 18, 2022
be62ca2
Update src/data/nat/factorization/basic.lean
BoltonBailey Jul 18, 2022
7a55cdb
Update src/data/nat/totient.lean
BoltonBailey Jul 18, 2022
b5cc6f5
Update src/data/nat/totient.lean
BoltonBailey Jul 18, 2022
107eb2e
Update src/group_theory/exponent.lean
BoltonBailey Jul 18, 2022
6c4ccf8
Update src/group_theory/exponent.lean
BoltonBailey Jul 18, 2022
11bf3fe
Update src/data/nat/factorization/basic.lean
BoltonBailey Jul 18, 2022
bb8bc47
Merge branch 'master' of github.com:leanprover-community/mathlib into…
BoltonBailey Jul 19, 2022
a004018
Merge branch 'master' into BoltonBailey/computable-factorization
stuart-presnell Jul 19, 2022
6cbf4d9
Merge branch 'master' into BoltonBailey/computable-factorization
stuart-presnell Jul 20, 2022
6d0fc7b
Fix error in merge
stuart-presnell Jul 20, 2022
b49c9b7
Merge branch 'BoltonBailey/computable-factorization' of https://githu…
stuart-presnell Jul 20, 2022
File filter

Filter by extension

Filter by extension

Conversations
Failed to load comments.
Loading
Jump to
Jump to file
Failed to load files.
Loading
Diff view
Diff view
1 change: 1 addition & 0 deletions src/analysis/special_functions/log/basic.lean
Original file line number Diff line number Diff line change
Expand Up @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE.
Authors: Chris Hughes, Abhimanyu Pallavi Sudhir, Jean Lo, Calle Sönne
-/
import analysis.special_functions.exp
import data.nat.factorization.basic

/-!
# Real logarithm
Expand Down
6 changes: 3 additions & 3 deletions src/data/nat/choose/factorization.lean
Original file line number Diff line number Diff line change
Expand Up @@ -5,7 +5,7 @@ Authors: Bolton Bailey, Patrick Stevens, Thomas Browning
-/

import data.nat.choose.central
import number_theory.padics.padic_norm
import data.nat.factorization.basic
import data.nat.multiplicity

/-!
Expand Down Expand Up @@ -38,7 +38,7 @@ begin
by_cases h : (choose n k).factorization p = 0, { simp [h] },
have hp : p.prime := not.imp_symm (choose n k).factorization_eq_zero_of_non_prime h,
have hkn : k ≤ n, { refine le_of_not_lt (λ hnk, h _), simp [choose_eq_zero_of_lt hnk] },
rw [←@padic_val_nat_eq_factorization p _ ⟨hp⟩, @padic_val_nat_def _ ⟨hp⟩ _ (choose_pos hkn)],
rw [factorization_def _ hp, @padic_val_nat_def _ ⟨hp⟩ _ (choose_pos hkn)],
simp only [hp.multiplicity_choose hkn (lt_add_one _), enat.get_coe],
refine (finset.card_filter_le _ _).trans (le_of_eq (nat.card_Ico _ _)),
end
Expand Down Expand Up @@ -73,7 +73,7 @@ begin
{ exact factorization_eq_zero_of_non_prime (choose n k) hp },
cases lt_or_le n k with hnk hkn,
{ simp [choose_eq_zero_of_lt hnk] },
rw [←@padic_val_nat_eq_factorization p _ ⟨hp⟩, @padic_val_nat_def _ ⟨hp⟩ _ (choose_pos hkn)],
rw [factorization_def _ hp, @padic_val_nat_def _ ⟨hp⟩ _ (choose_pos hkn)],
simp only [hp.multiplicity_choose hkn (lt_add_one _), enat.get_coe,
finset.card_eq_zero, finset.filter_eq_empty_iff, not_le],
intros i hi,
Expand Down
101 changes: 88 additions & 13 deletions src/data/nat/factorization/basic.lean
Original file line number Diff line number Diff line change
Expand Up @@ -6,6 +6,7 @@ Authors: Stuart Presnell
import data.nat.prime
import data.finsupp.multiset
import algebra.big_operators.finsupp
BoltonBailey marked this conversation as resolved.
Show resolved Hide resolved
import number_theory.padics.padic_val
import tactic.interval_cases

/-!
Expand Down Expand Up @@ -42,23 +43,52 @@ namespace nat

/-- `n.factorization` is the finitely supported function `ℕ →₀ ℕ`
mapping each prime factor of `n` to its multiplicity in `n`. -/
noncomputable def factorization (n : ℕ) : ℕ →₀ ℕ := (n.factors : multiset ℕ).to_finsupp
def factorization (n : ℕ) : ℕ →₀ ℕ :=
{ support := n.factors.to_finset,
to_fun := λ p, if p.prime then padic_val_nat p n else 0,
mem_support_to_fun :=
begin
rcases eq_or_ne n 0 with rfl | hn0, { simp },
simp only [mem_factors hn0, mem_to_finset, ne.def, ite_eq_right_iff, not_forall,
exists_prop, and.congr_right_iff],
rintro p hp,
haveI := fact_iff.mpr hp,
exact dvd_iff_padic_val_nat_ne_zero hn0,
end }

lemma factorization_def (n : ℕ) {p : ℕ} (pp : p.prime) : n.factorization p = padic_val_nat p n :=
by simpa [factorization] using absurd pp

/-- We can write both `n.factorization p` and `n.factors.count p` to represent the power
of `p` in the factorization of `n`: we declare the former to be the simp-normal form. -/
@[simp] lemma factors_count_eq {n p : ℕ} : n.factors.count p = n.factorization p :=
begin
rcases n.eq_zero_or_pos with rfl | hn0, { simp [factorization] },
by_cases pp : p.prime, swap,
{ rw count_eq_zero_of_not_mem (mt prime_of_mem_factors pp), simp [factorization, pp] },
simp only [factorization, coe_mk, pp, if_true],
rw [←enat.coe_inj, padic_val_nat_def' pp.ne_one hn0,
unique_factorization_monoid.multiplicity_eq_count_normalized_factors pp hn0.ne'],
simp [factors_eq],
end

lemma factorization_eq_factors_multiset (n : ℕ) :
n.factorization = (n.factors : multiset ℕ).to_finsupp :=
by { ext p, simp }

lemma multiplicity_eq_factorization {n p : ℕ} (pp : p.prime) (hn : n ≠ 0) :
multiplicity p n = n.factorization p :=
by simp [factorization, pp, (padic_val_nat_def' pp.ne_one hn.bot_lt)]

/-! ### Basic facts about factorization -/

@[simp] lemma factorization_prod_pow_eq_self {n : ℕ} (hn : n ≠ 0) : n.factorization.prod pow = n :=
begin
rw factorization_eq_factors_multiset n,
simp only [←prod_to_multiset, factorization, multiset.coe_prod, multiset.to_finsupp_to_multiset],
exact prod_factors hn,
end

/-- We can write both `n.factorization p` and `n.factors.count p` to represent the power
of `p` in the factorization of `n`: we declare the former to be the simp-normal form.
However, since `factorization` is a finsupp it's noncomputable. This theorem can also
be used in reverse to compute values of `factorization n p` when required. -/
@[simp] lemma factors_count_eq {n p : ℕ} : n.factors.count p = n.factorization p :=
by simp [factorization]

lemma eq_of_factorization_eq {a b : ℕ} (ha : a ≠ 0) (hb : b ≠ 0)
(h : ∀ p : ℕ, a.factorization p = b.factorization p) : a = b :=
eq_of_perm_factors ha hb (by simpa only [list.perm_iff_count, factors_count_eq] using h)
Expand All @@ -68,14 +98,14 @@ lemma factorization_inj : set.inj_on factorization { x : ℕ | x ≠ 0 } :=
λ a ha b hb h, eq_of_factorization_eq ha hb (λ p, by simp [h])

@[simp] lemma factorization_zero : factorization 0 = 0 :=
by simp [factorization]
by { simp [factorization], refl }
stuart-presnell marked this conversation as resolved.
Show resolved Hide resolved

@[simp] lemma factorization_one : factorization 1 = 0 :=
by simp [factorization]
by { simp [factorization], refl }

/-- The support of `n.factorization` is exactly `n.factors.to_finset` -/
@[simp] lemma support_factorization {n : ℕ} : n.factorization.support = n.factors.to_finset :=
by simpa [factorization, multiset.to_finsupp_support]
by simp [factorization]

lemma factor_iff_mem_factorization {n p : ℕ} : p ∈ n.factorization.support ↔ p ∈ n.factors :=
by simp only [support_factorization, list.mem_to_finset]
Expand Down Expand Up @@ -111,7 +141,10 @@ by rwa [←factors_count_eq, count_pos, mem_factors_iff_dvd hn hp]

/-- The only numbers with empty prime factorization are `0` and `1` -/
lemma factorization_eq_zero_iff (n : ℕ) : n.factorization = 0 ↔ n = 0 ∨ n = 1 :=
by simp [factorization, add_equiv.map_eq_zero_iff, multiset.coe_eq_zero]
begin
rw factorization_eq_factors_multiset n,
simp [factorization, add_equiv.map_eq_zero_iff, multiset.coe_eq_zero],
end

lemma factorization_eq_zero_iff' (n p : ℕ) :
n.factorization p = 0 ↔ ¬p.prime ∨ ¬p ∣ n ∨ n = 0 :=
Expand Down Expand Up @@ -203,7 +236,6 @@ lemma eq_factorization_iff {n : ℕ} {f : ℕ →₀ ℕ} (hn : n ≠ 0) (hf :
λ h, by rw [←h, prod_pow_factorization_eq_self hf]⟩

/-- The equiv between `ℕ+` and `ℕ →₀ ℕ` with support in the primes. -/
noncomputable
def factorization_equiv : ℕ+ ≃ {f : ℕ →₀ ℕ | ∀ p ∈ f.support, prime p} :=
{ to_fun := λ ⟨n, hn⟩, ⟨n.factorization, λ _, prime_of_mem_factorization⟩,
inv_fun := λ ⟨f, hf⟩, ⟨f.prod pow,
Expand Down Expand Up @@ -383,6 +415,19 @@ begin
simp [←factorization_le_iff_dvd he_pos hd_pos, h1, hea', heb'] },
end

@[to_additive sum_factors_gcd_add_sum_factors_mul]
lemma prod_factors_gcd_mul_prod_factors_mul {β : Type*} [comm_monoid β] (m n : ℕ) (f : ℕ → β) :
(m.gcd n).factors.to_finset.prod f * (m * n).factors.to_finset.prod f
= m.factors.to_finset.prod f * n.factors.to_finset.prod f :=
begin
rcases eq_or_ne n 0 with rfl | hm0, { simp },
rcases eq_or_ne m 0 with rfl | hn0, { simp },
rw [←@finset.prod_union_inter _ _ m.factors.to_finset n.factors.to_finset, mul_comm],
congr,
{ apply factors_mul_to_finset; assumption },
{ simp only [←support_factorization, factorization_gcd hn0 hm0, finsupp.support_inf] },
end

/-! ### Factorization and coprimes -/

/-- For coprime `a` and `b`, the power of `p` in `a * b` is the sum of the powers in `a` and `b` -/
Expand Down Expand Up @@ -530,4 +575,34 @@ begin
convert (factorization_disjoint_of_coprime hab) },
end

/-- Two positive naturals are equal if their prime padic valuations are equal -/
lemma eq_iff_prime_padic_val_nat_eq (a b : ℕ) (ha : a ≠ 0) (hb : b ≠ 0) :
a = b ↔ (∀ p : ℕ, p.prime → padic_val_nat p a = padic_val_nat p b) :=
begin
split,
{ rintros rfl, simp },
{ intro h,
refine eq_of_factorization_eq ha hb (λ p, _),
by_cases pp : p.prime,
{ simp [factorization_def, pp, h p pp] },
{ simp [factorization_eq_zero_of_non_prime, pp] } },
end

lemma prod_pow_prime_padic_val_nat (n : nat) (hn : n ≠ 0) (m : nat) (pr : n < m) :
∏ p in finset.filter nat.prime (finset.range m), p ^ (padic_val_nat p n) = n :=
begin
nth_rewrite_rhs 0 ←factorization_prod_pow_eq_self hn,
rw eq_comm,
apply finset.prod_subset_one_on_sdiff,
{ exact λ p hp, finset.mem_filter.mpr
⟨finset.mem_range.mpr (gt_of_gt_of_ge pr (le_of_mem_factorization hp)),
prime_of_mem_factorization hp⟩ },
{ intros p hp,
cases finset.mem_sdiff.mp hp with hp1 hp2,
rw ←factorization_def n (finset.mem_filter.mp hp1).2,
simp [finsupp.not_mem_support_iff.mp hp2] },
{ intros p hp,
simp [factorization_def n (prime_of_mem_factorization hp)] }
end

end nat
5 changes: 1 addition & 4 deletions src/data/nat/totient.lean
Original file line number Diff line number Diff line change
Expand Up @@ -3,13 +3,10 @@ Copyright (c) 2018 Chris Hughes. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Chris Hughes
-/
import algebra.big_operators.basic
import data.nat.prime
import data.zmod.basic
BoltonBailey marked this conversation as resolved.
Show resolved Hide resolved
import ring_theory.multiplicity
import data.nat.periodic
import data.nat.factorization.basic
import algebra.char_p.two
BoltonBailey marked this conversation as resolved.
Show resolved Hide resolved
import number_theory.divisors

/-!
# Euler's totient function
Expand Down
2 changes: 1 addition & 1 deletion src/group_theory/exponent.lean
Original file line number Diff line number Diff line change
Expand Up @@ -7,7 +7,7 @@ import group_theory.order_of_element
import algebra.punit_instances
import algebra.gcd_monoid.finset
BoltonBailey marked this conversation as resolved.
Show resolved Hide resolved
import tactic.by_contra
import number_theory.padics.padic_val
import data.nat.factorization.basic
BoltonBailey marked this conversation as resolved.
Show resolved Hide resolved

/-!
# Exponent of a group
Expand Down
33 changes: 5 additions & 28 deletions src/number_theory/padics/padic_val.lean
Original file line number Diff line number Diff line change
Expand Up @@ -9,7 +9,6 @@ import ring_theory.int.basic
import tactic.basic
import tactic.ring_exp
import number_theory.divisors
import data.nat.factorization.basic

/-!
# p-adic Valuation
Expand Down Expand Up @@ -237,6 +236,11 @@ begin
solve_by_elim,
end

lemma dvd_iff_padic_val_nat_ne_zero {p n : ℕ} [fact p.prime] (hn0 : n ≠ 0) :
(p ∣ n) ↔ padic_val_nat p n ≠ 0 :=
⟨λ h, one_le_iff_ne_zero.mp (one_le_padic_val_nat_of_dvd hn0.bot_lt h),
λ h, not_not.1 (mt padic_val_nat.eq_zero_of_not_dvd h)⟩

end padic_val_nat

namespace padic_val_rat
Expand Down Expand Up @@ -510,35 +514,8 @@ protected lemma padic_val_nat.div' {p : ℕ} [p_prime : fact p.prime] :
{ exact hc } },
end

lemma padic_val_nat_eq_factorization (p n : ℕ) [hp : fact p.prime] :
padic_val_nat p n = n.factorization p :=
begin
by_cases hn : n = 0, { subst hn, simp },
rw @padic_val_nat_def p _ n (nat.pos_of_ne_zero hn),
simp [@multiplicity_eq_factorization n p hp.elim hn],
end

open_locale big_operators

lemma prod_pow_prime_padic_val_nat (n : nat) (hn : n ≠ 0) (m : nat) (pr : n < m) :
∏ p in finset.filter nat.prime (finset.range m), p ^ (padic_val_nat p n) = n :=
begin
nth_rewrite_rhs 0 ←factorization_prod_pow_eq_self hn,
rw eq_comm,
apply finset.prod_subset_one_on_sdiff,
{ exact λ p hp, finset.mem_filter.mpr
⟨finset.mem_range.mpr (gt_of_gt_of_ge pr (le_of_mem_factorization hp)),
prime_of_mem_factorization hp⟩ },
{ intros p hp,
cases finset.mem_sdiff.mp hp with hp1 hp2,
haveI := fact_iff.mpr (finset.mem_filter.mp hp1).2,
rw padic_val_nat_eq_factorization p n,
simp [finsupp.not_mem_support_iff.mp hp2] },
{ intros p hp,
haveI := fact_iff.mpr (prime_of_mem_factorization hp),
simp [padic_val_nat_eq_factorization] }
end

lemma range_pow_padic_val_nat_subset_divisors {n : ℕ} (p : ℕ) (hn : n ≠ 0) :
(finset.range (padic_val_nat p n + 1)).image (pow p) ⊆ n.divisors :=
begin
Expand Down
18 changes: 0 additions & 18 deletions src/ring_theory/multiplicity.lean
Original file line number Diff line number Diff line change
Expand Up @@ -6,7 +6,6 @@ Authors: Robert Y. Lewis, Chris Hughes
import algebra.associated
import algebra.big_operators.basic
import ring_theory.valuation.basic
import data.nat.factorization.basic

/-!
# Multiplicity of a divisor
Expand Down Expand Up @@ -495,21 +494,4 @@ begin
exact hp this
end

lemma multiplicity_eq_factorization {n p : ℕ} (pp : p.prime) (hn : n ≠ 0) :
multiplicity p n = n.factorization p :=
multiplicity.eq_coe_iff.mpr ⟨pow_factorization_dvd n p, pow_succ_factorization_not_dvd hn pp⟩

@[to_additive sum_factors_gcd_add_sum_factors_mul]
lemma prod_factors_gcd_mul_prod_factors_mul {β : Type*} [comm_monoid β] (m n : ℕ) (f : ℕ → β) :
(m.gcd n).factors.to_finset.prod f * (m * n).factors.to_finset.prod f
= m.factors.to_finset.prod f * n.factors.to_finset.prod f :=
begin
rcases eq_or_ne n 0 with rfl | hm0, { simp },
rcases eq_or_ne m 0 with rfl | hn0, { simp },
rw [←@finset.prod_union_inter _ _ m.factors.to_finset n.factors.to_finset, mul_comm],
congr,
{ apply factors_mul_to_finset; assumption },
{ simp only [←support_factorization, factorization_gcd hn0 hm0, finsupp.support_inf] },
end

end nat