From 62e346a8a1f67d71b54f833968c4835e42f40ed2 Mon Sep 17 00:00:00 2001 From: Koundinya Vajjha Date: Thu, 7 Sep 2023 22:03:11 -0700 Subject: [PATCH 01/74] feat: stated the main theorem --- Counterexamples/HarmonicNotInteger.lean | 29 +++++++++++++++++++++++++ 1 file changed, 29 insertions(+) create mode 100644 Counterexamples/HarmonicNotInteger.lean diff --git a/Counterexamples/HarmonicNotInteger.lean b/Counterexamples/HarmonicNotInteger.lean new file mode 100644 index 0000000000000..122e75aec571b --- /dev/null +++ b/Counterexamples/HarmonicNotInteger.lean @@ -0,0 +1,29 @@ +/- +Copyright (c) 2023 Koundinya Vajjha. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Koundinya Vajjha +-/ + +import Mathlib.NumberTheory.Padics.PadicNorm +import Mathlib.NumberTheory.Padics.PadicNumbers +import Mathlib.NumberTheory.Padics.PadicIntegers + +/-! + +The nth Harmonic number is not an integer. We formalize the proof using +2-adic valuations. This proof is due to Theisinger. + +Reference: +https://kconrad.math.uconn.edu/blurbs/gradnumthy/padicharmonicsum.pdf + +-/ +open Rat + +variable (p : ℕ) (x : ℤ) [Fact p.Prime] + + +def harmonic : ℕ → ℚ +| 0 => 0 +| (k+1) => 1 / (k+1) + harmonic k + +theorem harmonic_not_int : ∀ n, n ≥ 2 -> (harmonic n).isInt = False := sorry From 5254a3e847f8cb8abc8e90fc80226888c7e89ba2 Mon Sep 17 00:00:00 2001 From: Koundinya Vajjha Date: Thu, 7 Sep 2023 22:22:12 -0700 Subject: [PATCH 02/74] wip: progress --- Counterexamples/HarmonicNotInteger.lean | 20 +++++++++++++++++++- 1 file changed, 19 insertions(+), 1 deletion(-) diff --git a/Counterexamples/HarmonicNotInteger.lean b/Counterexamples/HarmonicNotInteger.lean index 122e75aec571b..b0215f1f03f49 100644 --- a/Counterexamples/HarmonicNotInteger.lean +++ b/Counterexamples/HarmonicNotInteger.lean @@ -26,4 +26,22 @@ def harmonic : ℕ → ℚ | 0 => 0 | (k+1) => 1 / (k+1) + harmonic k -theorem harmonic_not_int : ∀ n, n ≥ 2 -> (harmonic n).isInt = False := sorry +theorem not_int_of_not_padic_int (a : ℚ) : padicNorm 2 a >1 → a.isInt = False := by { + intro H + by_cases (a = 0) + { + rw [h] at H h + dsimp [padicNorm] at * + simp at H + } + { + dsimp [padicNorm] at * + sorry + } +} + +theorem harmonic_not_int : ∀ n, n ≥ 2 -> (harmonic n).isInt = False := by { + intro n Hn + apply not_int_of_not_padic_int + +} From 886117ab226aef116bb6b1219db9ed266073b47d Mon Sep 17 00:00:00 2001 From: Koundinya Vajjha Date: Fri, 8 Sep 2023 15:55:26 -0700 Subject: [PATCH 03/74] feat: prove not_int_of_not_padic_int --- Counterexamples/HarmonicNotInteger.lean | 21 +++++++++++++++++---- 1 file changed, 17 insertions(+), 4 deletions(-) diff --git a/Counterexamples/HarmonicNotInteger.lean b/Counterexamples/HarmonicNotInteger.lean index b0215f1f03f49..51386dd6918db 100644 --- a/Counterexamples/HarmonicNotInteger.lean +++ b/Counterexamples/HarmonicNotInteger.lean @@ -26,21 +26,34 @@ def harmonic : ℕ → ℚ | 0 => 0 | (k+1) => 1 / (k+1) + harmonic k -theorem not_int_of_not_padic_int (a : ℚ) : padicNorm 2 a >1 → a.isInt = False := by { +theorem not_int_of_not_padic_int (a : ℚ) : padicNorm 2 a > 1 → ¬ a.isInt := by { intro H + suffices : a.den ≠ 1; simpa [Rat.isInt] by_cases (a = 0) { rw [h] at H h dsimp [padicNorm] at * - simp at H + simp only at H } { dsimp [padicNorm] at * - sorry + split_ifs at H; try contradiction + suffices : padicValRat 2 a < 0 + unfold padicValRat at this + intro Hden + rw [Hden] at this + simp only [padicValNat.one, CharP.cast_eq_zero, sub_zero] at this + norm_cast + have Hz : 0 = -0 := by {norm_num} + rw [Hz] + apply lt_neg_of_lt_neg + have hx : (1 : ℚ) < 2 := by {norm_num} + rw [← zpow_lt_iff_lt hx] + simpa only [zpow_zero] } } -theorem harmonic_not_int : ∀ n, n ≥ 2 -> (harmonic n).isInt = False := by { +theorem harmonic_not_int : ∀ n, n ≥ 2 -> ¬ (harmonic n).isInt := by { intro n Hn apply not_int_of_not_padic_int From 7e789be616bf1eb7d256bfb424803358ae6323e3 Mon Sep 17 00:00:00 2001 From: Koundinya Vajjha Date: Sat, 9 Sep 2023 00:15:29 -0700 Subject: [PATCH 04/74] progress, add harmonic ne_zero --- Counterexamples/HarmonicNotInteger.lean | 39 +++++++++++++++++++++++-- 1 file changed, 37 insertions(+), 2 deletions(-) diff --git a/Counterexamples/HarmonicNotInteger.lean b/Counterexamples/HarmonicNotInteger.lean index 51386dd6918db..d465a10185eeb 100644 --- a/Counterexamples/HarmonicNotInteger.lean +++ b/Counterexamples/HarmonicNotInteger.lean @@ -7,7 +7,6 @@ Authors: Koundinya Vajjha import Mathlib.NumberTheory.Padics.PadicNorm import Mathlib.NumberTheory.Padics.PadicNumbers import Mathlib.NumberTheory.Padics.PadicIntegers - /-! The nth Harmonic number is not an integer. We formalize the proof using @@ -17,6 +16,9 @@ Reference: https://kconrad.math.uconn.edu/blurbs/gradnumthy/padicharmonicsum.pdf -/ + +namespace Counterexample + open Rat variable (p : ℕ) (x : ℤ) [Fact p.Prime] @@ -26,6 +28,27 @@ def harmonic : ℕ → ℚ | 0 => 0 | (k+1) => 1 / (k+1) + harmonic k + +lemma harmonic_ne_zero : ∀ n, n ≠ 0 → harmonic n > 0 := by { + intros n Hn + induction' n with k ih; try contradiction + dsimp [harmonic] + by_cases (k = 0) + { + rw [h] + dsimp [harmonic] + norm_num + } + { + specialize (ih h) + have H₀ : (0 : ℚ) = 0 + 0 := by rfl + rw [H₀] + refine' add_lt_add _ ih + apply div_pos; try norm_num + norm_cast; simp only [add_pos_iff, or_true] + } +} + theorem not_int_of_not_padic_int (a : ℚ) : padicNorm 2 a > 1 → ¬ a.isInt := by { intro H suffices : a.den ≠ 1; simpa [Rat.isInt] @@ -42,7 +65,7 @@ theorem not_int_of_not_padic_int (a : ℚ) : padicNorm 2 a > 1 → ¬ a.isInt := unfold padicValRat at this intro Hden rw [Hden] at this - simp only [padicValNat.one, CharP.cast_eq_zero, sub_zero] at this + simp only [padicValNat.one, sub_zero] at this norm_cast have Hz : 0 = -0 := by {norm_num} rw [Hz] @@ -56,5 +79,17 @@ theorem not_int_of_not_padic_int (a : ℚ) : padicNorm 2 a > 1 → ¬ a.isInt := theorem harmonic_not_int : ∀ n, n ≥ 2 -> ¬ (harmonic n).isInt := by { intro n Hn apply not_int_of_not_padic_int + unfold padicNorm + split_ifs with h + { + have := harmonic_ne_zero n + rw [h] at this + apply this + linarith + } + { + norm_cast + + } } From d8f7d3eeb12f2a4d60aa76dfb9c88fae5959359c Mon Sep 17 00:00:00 2001 From: Koundinya Vajjha Date: Sat, 9 Sep 2023 09:07:17 -0700 Subject: [PATCH 05/74] slight progress. --- Counterexamples/HarmonicNotInteger.lean | 12 +++++++++--- 1 file changed, 9 insertions(+), 3 deletions(-) diff --git a/Counterexamples/HarmonicNotInteger.lean b/Counterexamples/HarmonicNotInteger.lean index d465a10185eeb..d87dcbc96c8ad 100644 --- a/Counterexamples/HarmonicNotInteger.lean +++ b/Counterexamples/HarmonicNotInteger.lean @@ -49,7 +49,8 @@ lemma harmonic_ne_zero : ∀ n, n ≠ 0 → harmonic n > 0 := by { } } -theorem not_int_of_not_padic_int (a : ℚ) : padicNorm 2 a > 1 → ¬ a.isInt := by { +theorem not_int_of_not_padic_int (a : ℚ) : + padicNorm 2 a > 1 → ¬ a.isInt := by { intro H suffices : a.den ≠ 1; simpa [Rat.isInt] by_cases (a = 0) @@ -66,7 +67,7 @@ theorem not_int_of_not_padic_int (a : ℚ) : padicNorm 2 a > 1 → ¬ a.isInt := intro Hden rw [Hden] at this simp only [padicValNat.one, sub_zero] at this - norm_cast + norm_cast at H have Hz : 0 = -0 := by {norm_num} rw [Hz] apply lt_neg_of_lt_neg @@ -88,7 +89,12 @@ theorem harmonic_not_int : ∀ n, n ≥ 2 -> ¬ (harmonic n).isInt := by { linarith } { - norm_cast + apply one_lt_zpow; try simp + apply lt_neg.mp + rw [neg_zero] + + + } From c7c345cf1ec1a2da67b9edc0be23b9794a8e58f6 Mon Sep 17 00:00:00 2001 From: Koundinya Vajjha Date: Sat, 9 Sep 2023 22:44:05 -0700 Subject: [PATCH 06/74] add finset defn of harmonic numbers --- Counterexamples/HarmonicNotInteger.lean | 25 ++++++++++++++++++------- 1 file changed, 18 insertions(+), 7 deletions(-) diff --git a/Counterexamples/HarmonicNotInteger.lean b/Counterexamples/HarmonicNotInteger.lean index d87dcbc96c8ad..671f6bb443631 100644 --- a/Counterexamples/HarmonicNotInteger.lean +++ b/Counterexamples/HarmonicNotInteger.lean @@ -7,6 +7,7 @@ Authors: Koundinya Vajjha import Mathlib.NumberTheory.Padics.PadicNorm import Mathlib.NumberTheory.Padics.PadicNumbers import Mathlib.NumberTheory.Padics.PadicIntegers +import Mathlib.Data.Int.Log /-! The nth Harmonic number is not an integer. We formalize the proof using @@ -20,14 +21,22 @@ https://kconrad.math.uconn.edu/blurbs/gradnumthy/padicharmonicsum.pdf namespace Counterexample open Rat - -variable (p : ℕ) (x : ℤ) [Fact p.Prime] - +open BigOperators def harmonic : ℕ → ℚ | 0 => 0 | (k+1) => 1 / (k+1) + harmonic k +def harmonic' : ℕ → ℚ := fun n => ∑ i in Finset.range n, 1 / (i + 1) + +lemma harmonic_harmonic' n : harmonic n = harmonic' n := by { + induction' n with k ih ; try simp + dsimp [harmonic, harmonic'] + rw [Finset.sum_range_succ, ih] + dsimp [harmonic'] + rw [add_comm] +} + lemma harmonic_ne_zero : ∀ n, n ≠ 0 → harmonic n > 0 := by { intros n Hn @@ -77,6 +86,7 @@ theorem not_int_of_not_padic_int (a : ℚ) : } } + theorem harmonic_not_int : ∀ n, n ≥ 2 -> ¬ (harmonic n).isInt := by { intro n Hn apply not_int_of_not_padic_int @@ -92,10 +102,11 @@ theorem harmonic_not_int : ∀ n, n ≥ 2 -> ¬ (harmonic n).isInt := by { apply one_lt_zpow; try simp apply lt_neg.mp rw [neg_zero] - - - - + suffices Hlog : padicValRat 2 (harmonic n) = -Nat.log 2 n + rw [Hlog] + simpa only [Left.neg_neg_iff, Nat.cast_pos, Nat.log_pos_iff, and_true] + norm_cast + sorry } } From 25b4e914f10add5553e8c2c41eb634b7bea8cca3 Mon Sep 17 00:00:00 2001 From: Koundinya Vajjha Date: Sun, 10 Sep 2023 10:35:54 -0700 Subject: [PATCH 07/74] wip min_eq_padicValRat --- Counterexamples/HarmonicNotInteger.lean | 26 +++++++++++++++++++++++-- 1 file changed, 24 insertions(+), 2 deletions(-) diff --git a/Counterexamples/HarmonicNotInteger.lean b/Counterexamples/HarmonicNotInteger.lean index 671f6bb443631..bfac45e0a0308 100644 --- a/Counterexamples/HarmonicNotInteger.lean +++ b/Counterexamples/HarmonicNotInteger.lean @@ -11,7 +11,7 @@ import Mathlib.Data.Int.Log /-! The nth Harmonic number is not an integer. We formalize the proof using -2-adic valuations. This proof is due to Theisinger. +2-adic valuations. This proof is due to K¨ursch´ak. Reference: https://kconrad.math.uconn.edu/blurbs/gradnumthy/padicharmonicsum.pdf @@ -23,6 +23,28 @@ namespace Counterexample open Rat open BigOperators +namespace padicValRat + + lemma min_eq_padicValRat {p : ℕ} [hp : Fact (Nat.Prime p)] {q r : ℚ} {hqr : q + r ≠ 0} (hval : padicValRat p q ≠ padicValRat p r) : padicValRat p (q + r) = min (padicValRat p q) (padicValRat p r) := by { + have Hmin := @padicValRat.min_le_padicValRat_add p hp q r hqr + wlog h : padicValRat p q < padicValRat p r generalizing q r with H + { + sorry + } + { + rw [min_eq_left (le_of_lt h)] at Hmin ⊢ + suffices Hreq : padicValRat p q ≥ padicValRat p (q + r) by linarith + calc padicValRat p q = padicValRat p (q + r - r) := by congr; simp + _ ≥ min (padicValRat p (q +r)) (padicValRat p r):= by _ + _ = padicValRat p (q + r) := _ + } + + } + +end padicValRat + + + def harmonic : ℕ → ℚ | 0 => 0 | (k+1) => 1 / (k+1) + harmonic k @@ -37,7 +59,6 @@ lemma harmonic_harmonic' n : harmonic n = harmonic' n := by { rw [add_comm] } - lemma harmonic_ne_zero : ∀ n, n ≠ 0 → harmonic n > 0 := by { intros n Hn induction' n with k ih; try contradiction @@ -107,6 +128,7 @@ theorem harmonic_not_int : ∀ n, n ≥ 2 -> ¬ (harmonic n).isInt := by { simpa only [Left.neg_neg_iff, Nat.cast_pos, Nat.log_pos_iff, and_true] norm_cast sorry + } } From d193c37be6eb14039b931e59661e8076e0df4fc7 Mon Sep 17 00:00:00 2001 From: Koundinya Vajjha Date: Mon, 11 Sep 2023 17:30:00 -0700 Subject: [PATCH 08/74] feat: min_eq_padicValRat --- Counterexamples/HarmonicNotInteger.lean | 30 +++++++++++++++++-------- 1 file changed, 21 insertions(+), 9 deletions(-) diff --git a/Counterexamples/HarmonicNotInteger.lean b/Counterexamples/HarmonicNotInteger.lean index bfac45e0a0308..092dd768084fe 100644 --- a/Counterexamples/HarmonicNotInteger.lean +++ b/Counterexamples/HarmonicNotInteger.lean @@ -25,21 +25,33 @@ open BigOperators namespace padicValRat - lemma min_eq_padicValRat {p : ℕ} [hp : Fact (Nat.Prime p)] {q r : ℚ} {hqr : q + r ≠ 0} (hval : padicValRat p q ≠ padicValRat p r) : padicValRat p (q + r) = min (padicValRat p q) (padicValRat p r) := by { - have Hmin := @padicValRat.min_le_padicValRat_add p hp q r hqr + lemma min_eq_padicValRat {p : ℕ} [hp : Fact (Nat.Prime p)] {q r : ℚ} {hqr : q + r ≠ 0} (hq : q ≠ 0) (hr : r ≠ 0) (hval : padicValRat p q ≠ padicValRat p r) : padicValRat p (q + r) = min (padicValRat p q) (padicValRat p r) := by { + have Hmin := padicValRat.min_le_padicValRat_add (p := p) (hp := hp) hqr wlog h : padicValRat p q < padicValRat p r generalizing q r with H { - sorry + push_neg at h + rw [add_comm, min_comm] + refine' (H hr hq _ _ _) + rwa [add_comm] + exact hval.symm + rwa [min_comm,add_comm] + exact Ne.lt_of_le (Ne.symm hval) h } { rw [min_eq_left (le_of_lt h)] at Hmin ⊢ suffices Hreq : padicValRat p q ≥ padicValRat p (q + r) by linarith - calc padicValRat p q = padicValRat p (q + r - r) := by congr; simp - _ ≥ min (padicValRat p (q +r)) (padicValRat p r):= by _ - _ = padicValRat p (q + r) := _ - } - - } + have Haux : padicValRat p q ≥ min (padicValRat p (q + r)) (padicValRat p r) := by { + calc padicValRat p q = padicValRat p ((q + r) - r) := by congr; simp + _ ≥ min (padicValRat p (q + r)) (padicValRat p (-r)):= by { + have hqr' : q + r + - r = q + r - r := by simp + have := padicValRat.min_le_padicValRat_add (p := p) (hp := hp) (q := q+r) (r := -r) (by simpa) + rwa [ge_iff_le,←hqr'] + } + _ = min (padicValRat p (q + r)) (padicValRat p r) := by simp only [padicValRat.neg] + } + rw [min_def] at Haux; split_ifs at Haux with Hspl; try assumption + linarith + }} end padicValRat From a67c69f792978533b226def55d2c49ea441ce7bc Mon Sep 17 00:00:00 2001 From: Koundinya Vajjha Date: Mon, 11 Sep 2023 19:12:42 -0700 Subject: [PATCH 09/74] feat: golf attempt --- Counterexamples/HarmonicNotInteger.lean | 19 ++++++++----------- 1 file changed, 8 insertions(+), 11 deletions(-) diff --git a/Counterexamples/HarmonicNotInteger.lean b/Counterexamples/HarmonicNotInteger.lean index 092dd768084fe..62d64bddcb7f0 100644 --- a/Counterexamples/HarmonicNotInteger.lean +++ b/Counterexamples/HarmonicNotInteger.lean @@ -25,15 +25,16 @@ open BigOperators namespace padicValRat - lemma min_eq_padicValRat {p : ℕ} [hp : Fact (Nat.Prime p)] {q r : ℚ} {hqr : q + r ≠ 0} (hq : q ≠ 0) (hr : r ≠ 0) (hval : padicValRat p q ≠ padicValRat p r) : padicValRat p (q + r) = min (padicValRat p q) (padicValRat p r) := by { + lemma min_eq_padicValRat {p : ℕ} [hp : Fact (Nat.Prime p)] {q r : ℚ} (hqr : q + r ≠ 0) (hq : q ≠ 0) (hr : r ≠ 0) + (hval : padicValRat p q ≠ padicValRat p r) : + padicValRat p (q + r) = min (padicValRat p q) (padicValRat p r) := by { have Hmin := padicValRat.min_le_padicValRat_add (p := p) (hp := hp) hqr wlog h : padicValRat p q < padicValRat p r generalizing q r with H { push_neg at h rw [add_comm, min_comm] - refine' (H hr hq _ _ _) + refine' (H _ hr hq hval.symm _ _) rwa [add_comm] - exact hval.symm rwa [min_comm,add_comm] exact Ne.lt_of_le (Ne.symm hval) h } @@ -42,21 +43,17 @@ namespace padicValRat suffices Hreq : padicValRat p q ≥ padicValRat p (q + r) by linarith have Haux : padicValRat p q ≥ min (padicValRat p (q + r)) (padicValRat p r) := by { calc padicValRat p q = padicValRat p ((q + r) - r) := by congr; simp - _ ≥ min (padicValRat p (q + r)) (padicValRat p (-r)):= by { - have hqr' : q + r + - r = q + r - r := by simp - have := padicValRat.min_le_padicValRat_add (p := p) (hp := hp) (q := q+r) (r := -r) (by simpa) - rwa [ge_iff_le,←hqr'] + _ ≥ min (padicValRat p (q + r)) (padicValRat p (-r)) := ge_iff_le.mp <| le_trans (padicValRat.min_le_padicValRat_add (q := q+r) (r := -r) (by simpa)) (by rw [add_neg_cancel_right, add_sub_cancel]) + _ = min (padicValRat p (q + r)) (padicValRat p r) := by rw [padicValRat.neg] } - _ = min (padicValRat p (q + r)) (padicValRat p r) := by simp only [padicValRat.neg] - } - rw [min_def] at Haux; split_ifs at Haux with Hspl; try assumption + rw [min_def] at Haux; + split_ifs at Haux with Hspl; try assumption linarith }} end padicValRat - def harmonic : ℕ → ℚ | 0 => 0 | (k+1) => 1 / (k+1) + harmonic k From 0ed3c48f9149f66d5efe3123b7cc249677a9d467 Mon Sep 17 00:00:00 2001 From: Koundinya Vajjha Date: Tue, 12 Sep 2023 19:02:46 -0700 Subject: [PATCH 10/74] add le_nat_log + wip padicValRat_2_sum --- Counterexamples/HarmonicNotInteger.lean | 66 +++++++++++++++++++++++-- 1 file changed, 61 insertions(+), 5 deletions(-) diff --git a/Counterexamples/HarmonicNotInteger.lean b/Counterexamples/HarmonicNotInteger.lean index 62d64bddcb7f0..a29f31c7ea20e 100644 --- a/Counterexamples/HarmonicNotInteger.lean +++ b/Counterexamples/HarmonicNotInteger.lean @@ -7,6 +7,7 @@ Authors: Koundinya Vajjha import Mathlib.NumberTheory.Padics.PadicNorm import Mathlib.NumberTheory.Padics.PadicNumbers import Mathlib.NumberTheory.Padics.PadicIntegers +import Mathlib.NumberTheory.Padics.PadicVal import Mathlib.Data.Int.Log /-! @@ -46,13 +47,29 @@ namespace padicValRat _ ≥ min (padicValRat p (q + r)) (padicValRat p (-r)) := ge_iff_le.mp <| le_trans (padicValRat.min_le_padicValRat_add (q := q+r) (r := -r) (by simpa)) (by rw [add_neg_cancel_right, add_sub_cancel]) _ = min (padicValRat p (q + r)) (padicValRat p r) := by rw [padicValRat.neg] } - rw [min_def] at Haux; + rw [min_def] at Haux split_ifs at Haux with Hspl; try assumption linarith }} end padicValRat +namespace padicValNat + + -- TODO: prove for when strict inequality holds. + lemma le_nat_log {p : ℕ} [hp : Fact (Nat.Prime p)] (n : ℕ): + padicValNat p n ≤ Nat.log p n := by { + by_cases (n = 0); simp [h] + apply Nat.le_log_of_pow_le (Nat.Prime.one_lt hp.elim) + by_contra Hnot + have H₁ := Nat.eq_zero_of_dvd_of_lt (@pow_padicValNat_dvd p n) + push_neg at Hnot + apply h + apply H₁ + exact Hnot + } + +end padicValNat def harmonic : ℕ → ℚ | 0 => 0 @@ -116,8 +133,47 @@ theorem not_int_of_not_padic_int (a : ℚ) : } } +lemma padicValRat_2_pow (r : ℕ) : padicValRat 2 (1 / 2^r) = -r := by { + rw [one_div,padicValRat.inv,neg_inj,padicValRat.pow (by simp)] + suffices : padicValRat 2 2 = 1 + simp only [this, mul_one] + rw [←padicValRat.self (p := 2)]; simp only [Nat.cast_ofNat] + norm_num +} + +#eval padicValRat 2 (1/4 + 1/7) + +lemma padicValRat_2_sum {r n : ℕ} (hr₁ : 2^r < n)(hr₂ : n < 2^(r+1)) : padicValRat 2 (1 / 2^r + 1 / n) = -r := by { + rw [padicValRat.min_eq_padicValRat] + { + rw [min_eq_left] + apply padicValRat_2_pow + rw [padicValRat_2_pow, one_div,padicValRat.inv,padicValRat.of_nat, neg_le_neg_iff, Nat.cast_le] + suffices Hr : r = Nat.log 2 n + rw [Hr] + apply padicValNat.le_nat_log + symm + rw [Nat.log_eq_iff] + exact ⟨le_of_lt hr₁,hr₂⟩ + right + constructor; trivial + rw [← Nat.pos_iff_ne_zero] + eapply lt_of_lt_of_le _ (le_of_lt hr₁) + apply pow_pos; trivial + } + sorry + sorry + sorry + rw [padicValRat_2_pow, one_div, padicValRat.inv,padicValRat.of_nat] + intro Hnot + simp only [padicValRat.of_nat, neg_inj, Nat.cast_inj] at Hnot + rw [Hnot] at hr₁ + have Hdvd := pow_padicValNat_dvd (p := 2) (n := n) + sorry + +} -theorem harmonic_not_int : ∀ n, n ≥ 2 -> ¬ (harmonic n).isInt := by { +theorem harmonic_not_int : ∀ n, n ≥ 2 → ¬ (harmonic n).isInt := by { intro n Hn apply not_int_of_not_padic_int unfold padicNorm @@ -132,10 +188,10 @@ theorem harmonic_not_int : ∀ n, n ≥ 2 -> ¬ (harmonic n).isInt := by { apply one_lt_zpow; try simp apply lt_neg.mp rw [neg_zero] - suffices Hlog : padicValRat 2 (harmonic n) = -Nat.log 2 n + suffices Hlog : padicValRat 2 (harmonic n) = -Int.log 2 (n : ℚ) rw [Hlog] - simpa only [Left.neg_neg_iff, Nat.cast_pos, Nat.log_pos_iff, and_true] - norm_cast + simp only [Int.log_natCast, Left.neg_neg_iff, Nat.cast_pos, Nat.log_pos_iff, and_true, Hn] + simp only [Int.log_natCast] sorry } From 39e7c9ba9406a047afb667cda923b769bed9206e Mon Sep 17 00:00:00 2001 From: Koundinya Vajjha Date: Wed, 13 Sep 2023 16:25:22 -0700 Subject: [PATCH 11/74] add padicvalnat le log lemmas --- Counterexamples/HarmonicNotInteger.lean | 69 ++++++++++++++----------- 1 file changed, 39 insertions(+), 30 deletions(-) diff --git a/Counterexamples/HarmonicNotInteger.lean b/Counterexamples/HarmonicNotInteger.lean index a29f31c7ea20e..2ed182d3f25ec 100644 --- a/Counterexamples/HarmonicNotInteger.lean +++ b/Counterexamples/HarmonicNotInteger.lean @@ -69,6 +69,9 @@ namespace padicValNat exact Hnot } + lemma le_nat_log_gen {p n₁ n₂ : ℕ} [Fact (Nat.Prime p)] (hn : n₁ ≤ n₂): + padicValNat p n₁ ≤ Nat.log p n₂ := le_trans (le_nat_log n₁) (Nat.log_mono_right hn) + end padicValNat def harmonic : ℕ → ℚ @@ -141,36 +144,10 @@ lemma padicValRat_2_pow (r : ℕ) : padicValRat 2 (1 / 2^r) = -r := by { norm_num } -#eval padicValRat 2 (1/4 + 1/7) - -lemma padicValRat_2_sum {r n : ℕ} (hr₁ : 2^r < n)(hr₂ : n < 2^(r+1)) : padicValRat 2 (1 / 2^r + 1 / n) = -r := by { - rw [padicValRat.min_eq_padicValRat] - { - rw [min_eq_left] - apply padicValRat_2_pow - rw [padicValRat_2_pow, one_div,padicValRat.inv,padicValRat.of_nat, neg_le_neg_iff, Nat.cast_le] - suffices Hr : r = Nat.log 2 n - rw [Hr] - apply padicValNat.le_nat_log - symm - rw [Nat.log_eq_iff] - exact ⟨le_of_lt hr₁,hr₂⟩ - right - constructor; trivial - rw [← Nat.pos_iff_ne_zero] - eapply lt_of_lt_of_le _ (le_of_lt hr₁) - apply pow_pos; trivial - } - sorry - sorry - sorry - rw [padicValRat_2_pow, one_div, padicValRat.inv,padicValRat.of_nat] - intro Hnot - simp only [padicValRat.of_nat, neg_inj, Nat.cast_inj] at Hnot - rw [Hnot] at hr₁ - have Hdvd := pow_padicValNat_dvd (p := 2) (n := n) - sorry - +lemma padicValRat_ge_neg_Nat_log {p n : ℕ}[hp : Fact (Nat.Prime p)]: ∀ q, q ≤ n → -Nat.log p n ≤ padicValRat p (1 / q) := by { + intros q Hq + rw [one_div,padicValRat.inv,padicValRat.of_nat, neg_le_neg_iff, Nat.cast_le] + apply padicValNat.le_nat_log_gen Hq } theorem harmonic_not_int : ∀ n, n ≥ 2 → ¬ (harmonic n).isInt := by { @@ -197,3 +174,35 @@ theorem harmonic_not_int : ∀ n, n ≥ 2 → ¬ (harmonic n).isInt := by { } } + + +-- lemma padicValRat_2_sum {r n : ℕ} (hr₁ : 2^r < n)(hr₂ : n < 2^(r+1)) : padicValRat 2 (1 / 2^r + 1 / n) = -r := by { + + -- have Hr : r = Nat.log 2 n := by { + -- symm + -- rw [Nat.log_eq_iff] + -- exact ⟨le_of_lt hr₁,hr₂⟩ + -- right + -- constructor; trivial + -- rw [← Nat.pos_iff_ne_zero] + -- eapply lt_of_lt_of_le _ (le_of_lt hr₁) + -- apply pow_pos; trivial + -- } + -- rw [padicValRat.min_eq_padicValRat] + -- { + -- rw [min_eq_left] + -- apply padicValRat_2_pow + -- rw [padicValRat_2_pow, one_div,padicValRat.inv,padicValRat.of_nat, neg_le_neg_iff, Nat.cast_le, Hr] + -- apply padicValNat.le_nat_log + -- } + -- sorry + -- sorry + -- sorry + + -- rw [padicValRat_2_pow, one_div, padicValRat.inv,padicValRat.of_nat] + -- intro Hnot + -- simp only [padicValRat.of_nat, neg_inj, Nat.cast_inj] at Hnot + -- rw [Hnot] at hr₁ + -- have Hdvd := pow_padicValNat_dvd (p := 2) (n := n) + -- sorry +-- } From 258d3cf457ceee6bd22f5c572d4f2e3c1d7ec54f Mon Sep 17 00:00:00 2001 From: Koundinya Vajjha Date: Thu, 14 Sep 2023 17:29:06 -0700 Subject: [PATCH 12/74] prove nat_log_eq_padicvalnat_iff --- Counterexamples/HarmonicNotInteger.lean | 18 ++++++++++++++++++ 1 file changed, 18 insertions(+) diff --git a/Counterexamples/HarmonicNotInteger.lean b/Counterexamples/HarmonicNotInteger.lean index 2ed182d3f25ec..ef7210638965d 100644 --- a/Counterexamples/HarmonicNotInteger.lean +++ b/Counterexamples/HarmonicNotInteger.lean @@ -72,6 +72,16 @@ namespace padicValNat lemma le_nat_log_gen {p n₁ n₂ : ℕ} [Fact (Nat.Prime p)] (hn : n₁ ≤ n₂): padicValNat p n₁ ≤ Nat.log p n₂ := le_trans (le_nat_log n₁) (Nat.log_mono_right hn) + lemma nat_log_eq_padicvalnat_iff {p : ℕ} [hp : Fact (Nat.Prime p)] (n : ℕ)(hn : 0 < n): Nat.log p n = padicValNat p n ↔ n < p^(padicValNat p n + 1) := by { + rw [Nat.log_eq_iff] + have Hdiv: p^padicValNat p n ≤ n := Nat.le_of_dvd hn pow_padicValNat_dvd + simp only [and_iff_right_iff_imp, Hdiv] + intros; trivial + right + refine' ⟨(Nat.Prime.one_lt' p).out,_⟩ + linarith + } + end padicValNat def harmonic : ℕ → ℚ @@ -150,6 +160,13 @@ lemma padicValRat_ge_neg_Nat_log {p n : ℕ}[hp : Fact (Nat.Prime p)]: ∀ q, q apply padicValNat.le_nat_log_gen Hq } +lemma padicValRat_ge_neg_Nat_log' {n : ℕ}: ∀ q, q ≤ n → q ≠ Nat.log 2 n → padicValRat 2 (1 / q) ≠ -Nat.log 2 n := by { + intros q Hq₁ Hq₂ + rw [one_div,padicValRat.inv,padicValRat.of_nat] + simp only [ne_eq, neg_inj, Nat.cast_inj] + +} + theorem harmonic_not_int : ∀ n, n ≥ 2 → ¬ (harmonic n).isInt := by { intro n Hn apply not_int_of_not_padic_int @@ -169,6 +186,7 @@ theorem harmonic_not_int : ∀ n, n ≥ 2 → ¬ (harmonic n).isInt := by { rw [Hlog] simp only [Int.log_natCast, Left.neg_neg_iff, Nat.cast_pos, Nat.log_pos_iff, and_true, Hn] simp only [Int.log_natCast] + sorry } From 57f8e457ea7984e71ea765fb582b2c394ecadf90 Mon Sep 17 00:00:00 2001 From: Koundinya Vajjha Date: Sat, 16 Sep 2023 12:23:09 -0700 Subject: [PATCH 13/74] add nat log not dvd --- Counterexamples/HarmonicNotInteger.lean | 70 ++++++++++++++++++++++++- 1 file changed, 69 insertions(+), 1 deletion(-) diff --git a/Counterexamples/HarmonicNotInteger.lean b/Counterexamples/HarmonicNotInteger.lean index ef7210638965d..855e28892a461 100644 --- a/Counterexamples/HarmonicNotInteger.lean +++ b/Counterexamples/HarmonicNotInteger.lean @@ -24,6 +24,8 @@ namespace Counterexample open Rat open BigOperators + + namespace padicValRat lemma min_eq_padicValRat {p : ℕ} [hp : Fact (Nat.Prime p)] {q r : ℚ} (hqr : q + r ≠ 0) (hq : q ≠ 0) (hr : r ≠ 0) @@ -52,6 +54,18 @@ namespace padicValRat linarith }} + -- lemma sum_eq_of_lt {p : ℕ} [hp : Fact (Nat.Prime p)] {q r : ℚ} (hqr : q + r ≠ 0) (hq : q ≠ 0) (hr : r ≠ 0) + -- (hval : padicValRat p q < padicValRat p r) : + -- padicValRat p (q + r) = padicValRat p q := by rw [min_eq_padicValRat hqr hq hr (ne_of_lt hval),min_eq_left (le_of_lt hval)] + + -- lemma of_sum_eq_min {p n j : ℕ} [hp : Fact (Nat.Prime p)] {F : ℕ → ℚ} (hF : ∀ i, i ≤ n → padicValRat p (F j) < padicValRat p (F i)) + -- (hn : ∀ i, F i > 0): + -- padicValRat p (∑ i in Finset.range (n+1), F i) = padicValRat p (F j) := by { + -- induction' n with d hd + -- simp at * + + -- } + end padicValRat namespace padicValNat @@ -82,6 +96,20 @@ namespace padicValNat linarith } + -- lemma nat_log_eq_padicvalnat_iff' {p : ℕ} [hp : Fact (Nat.Prime p)] (n : ℕ)(hn : 0 < n): Nat.log p n = padicValNat p n ↔ n < p^(Nat.log p n + 1) := by { + -- rw [Nat.log_eq_iff] + -- have Hdiv: p^padicValNat p n ≤ n := Nat.le_of_dvd hn pow_padicValNat_dvd + -- simp only [and_iff_right_iff_imp, Hdiv, true_and] + -- suffices : padicValNat p n = Nat.log p n; rw [this] + -- apply Nat.le_antisymm (le_nat_log _) _ + -- have Hp := (Nat.Prime.one_lt' p).out + -- rw [← pow_le_pow_iff Hp] + + -- -- right + -- -- refine' ⟨(Nat.Prime.one_lt' p).out,_⟩ + -- -- linarith + -- } + end padicValNat def harmonic : ℕ → ℚ @@ -90,6 +118,9 @@ def harmonic : ℕ → ℚ def harmonic' : ℕ → ℚ := fun n => ∑ i in Finset.range n, 1 / (i + 1) +#eval let x := 19; (padicValRat 2 (harmonic x), padicValRat 2 (x)) + + lemma harmonic_harmonic' n : harmonic n = harmonic' n := by { induction' n with k ih ; try simp dsimp [harmonic, harmonic'] @@ -160,11 +191,48 @@ lemma padicValRat_ge_neg_Nat_log {p n : ℕ}[hp : Fact (Nat.Prime p)]: ∀ q, q apply padicValNat.le_nat_log_gen Hq } +lemma nat_log_not_dvd {n : ℕ} : ∀ i, 0 < i ∧ i ≠ 2^(Nat.log 2 n) ∧ i < n → ¬ 2^(Nat.log 2 n) ∣ i := by { + rintro i ⟨Hi₁,Hi₂,Hi₃⟩ + simp only [Nat.instDvdNat] + push_neg + intros c Hc + have Hle : n < 2 ^ (Nat.log 2 n + 1) := Nat.lt_pow_succ_log_self (by simp) n + have Hpow : 2 ^ Nat.log 2 n * 2 = 2 ^(Nat.log 2 n + 1) := by exact rfl + obtain ⟨k,Hk⟩ := Nat.even_or_odd' c + by_cases (k = 0) + { + rw [h, mul_zero] at Hk + rcases Hk with rfl | rfl + linarith + simp only [zero_add, mul_one] at Hc + exact Hi₂ Hc + } + rcases Hk with rfl | rfl + { + rw [← mul_assoc, Hpow] at Hc + { + suffices Hnk : n < i by linarith + calc n ≤ n * k := Nat.le_mul_of_pos_right (Nat.pos_of_ne_zero h) + _ < 2^(Nat.log 2 n + 1) * k := Nat.mul_lt_mul Hle (le_refl _) (Nat.pos_of_ne_zero h) + _ = i := Hc.symm + } + } + { + rw [mul_add, ← mul_assoc, Hpow, mul_one] at Hc + suffices Hnk : n < i by linarith + calc n < 2 ^ (Nat.log 2 n + 1) := Hle + _ ≤ 2 ^ (Nat.log 2 n + 1)*k := Nat.le_mul_of_pos_right (Nat.pos_of_ne_zero h) + _ < 2 ^ (Nat.log 2 n + 1)*k + 2^ Nat.log 2 n := by linarith + _ = i := Hc.symm + } + } + + lemma padicValRat_ge_neg_Nat_log' {n : ℕ}: ∀ q, q ≤ n → q ≠ Nat.log 2 n → padicValRat 2 (1 / q) ≠ -Nat.log 2 n := by { intros q Hq₁ Hq₂ rw [one_div,padicValRat.inv,padicValRat.of_nat] simp only [ne_eq, neg_inj, Nat.cast_inj] - + sorry } theorem harmonic_not_int : ∀ n, n ≥ 2 → ¬ (harmonic n).isInt := by { From 4df29c5c30a0e969de99687a78a7d8a356a04a76 Mon Sep 17 00:00:00 2001 From: Koundinya Vajjha Date: Sat, 16 Sep 2023 12:31:37 -0700 Subject: [PATCH 14/74] padicValRat_ge_neg_Nat_log' --- Counterexamples/HarmonicNotInteger.lean | 11 +++++++---- 1 file changed, 7 insertions(+), 4 deletions(-) diff --git a/Counterexamples/HarmonicNotInteger.lean b/Counterexamples/HarmonicNotInteger.lean index 855e28892a461..d211ceab220b8 100644 --- a/Counterexamples/HarmonicNotInteger.lean +++ b/Counterexamples/HarmonicNotInteger.lean @@ -191,7 +191,7 @@ lemma padicValRat_ge_neg_Nat_log {p n : ℕ}[hp : Fact (Nat.Prime p)]: ∀ q, q apply padicValNat.le_nat_log_gen Hq } -lemma nat_log_not_dvd {n : ℕ} : ∀ i, 0 < i ∧ i ≠ 2^(Nat.log 2 n) ∧ i < n → ¬ 2^(Nat.log 2 n) ∣ i := by { +lemma nat_log_not_dvd {n : ℕ} : ∀ i, 0 < i ∧ i ≠ 2^(Nat.log 2 n) ∧ i ≤ n → ¬ 2^(Nat.log 2 n) ∣ i := by { rintro i ⟨Hi₁,Hi₂,Hi₃⟩ simp only [Nat.instDvdNat] push_neg @@ -228,11 +228,14 @@ lemma nat_log_not_dvd {n : ℕ} : ∀ i, 0 < i ∧ i ≠ 2^(Nat.log 2 n) ∧ i < } -lemma padicValRat_ge_neg_Nat_log' {n : ℕ}: ∀ q, q ≤ n → q ≠ Nat.log 2 n → padicValRat 2 (1 / q) ≠ -Nat.log 2 n := by { - intros q Hq₁ Hq₂ +lemma padicValRat_ge_neg_Nat_log' {n : ℕ}: ∀ q, 0 < q ∧ q ≤ n → q ≠ 2^Nat.log 2 n → padicValRat 2 (1 / q) ≠ -Nat.log 2 n := by { + rintro q ⟨Hq₁,Hq₂⟩ Hq₃ rw [one_div,padicValRat.inv,padicValRat.of_nat] simp only [ne_eq, neg_inj, Nat.cast_inj] - sorry + intro Hnot + have := pow_padicValNat_dvd (p := 2) (n := q) + rw [Hnot] at this + exact nat_log_not_dvd (n := n) q ⟨Hq₁,Hq₃,Hq₂⟩ this } theorem harmonic_not_int : ∀ n, n ≥ 2 → ¬ (harmonic n).isInt := by { From 3f020f945a7154ab348aa0ba1e682155f80aee3e Mon Sep 17 00:00:00 2001 From: Koundinya Vajjha Date: Sat, 16 Sep 2023 22:49:38 -0700 Subject: [PATCH 15/74] down to one sorry --- Counterexamples/HarmonicNotInteger.lean | 142 ++++++++++++------------ 1 file changed, 71 insertions(+), 71 deletions(-) diff --git a/Counterexamples/HarmonicNotInteger.lean b/Counterexamples/HarmonicNotInteger.lean index d211ceab220b8..cffdec0d26f0a 100644 --- a/Counterexamples/HarmonicNotInteger.lean +++ b/Counterexamples/HarmonicNotInteger.lean @@ -54,23 +54,15 @@ namespace padicValRat linarith }} - -- lemma sum_eq_of_lt {p : ℕ} [hp : Fact (Nat.Prime p)] {q r : ℚ} (hqr : q + r ≠ 0) (hq : q ≠ 0) (hr : r ≠ 0) - -- (hval : padicValRat p q < padicValRat p r) : - -- padicValRat p (q + r) = padicValRat p q := by rw [min_eq_padicValRat hqr hq hr (ne_of_lt hval),min_eq_left (le_of_lt hval)] + lemma sum_eq_of_lt {p : ℕ} [hp : Fact (Nat.Prime p)] {q r : ℚ} (hqr : q + r ≠ 0) (hq : q ≠ 0) (hr : r ≠ 0) + (hval : padicValRat p q < padicValRat p r) : + padicValRat p (q + r) = padicValRat p q := by rw [min_eq_padicValRat hqr hq hr (ne_of_lt hval),min_eq_left (le_of_lt hval)] - -- lemma of_sum_eq_min {p n j : ℕ} [hp : Fact (Nat.Prime p)] {F : ℕ → ℚ} (hF : ∀ i, i ≤ n → padicValRat p (F j) < padicValRat p (F i)) - -- (hn : ∀ i, F i > 0): - -- padicValRat p (∑ i in Finset.range (n+1), F i) = padicValRat p (F j) := by { - -- induction' n with d hd - -- simp at * - - -- } end padicValRat namespace padicValNat - -- TODO: prove for when strict inequality holds. lemma le_nat_log {p : ℕ} [hp : Fact (Nat.Prime p)] (n : ℕ): padicValNat p n ≤ Nat.log p n := by { by_cases (n = 0); simp [h] @@ -86,7 +78,8 @@ namespace padicValNat lemma le_nat_log_gen {p n₁ n₂ : ℕ} [Fact (Nat.Prime p)] (hn : n₁ ≤ n₂): padicValNat p n₁ ≤ Nat.log p n₂ := le_trans (le_nat_log n₁) (Nat.log_mono_right hn) - lemma nat_log_eq_padicvalnat_iff {p : ℕ} [hp : Fact (Nat.Prime p)] (n : ℕ)(hn : 0 < n): Nat.log p n = padicValNat p n ↔ n < p^(padicValNat p n + 1) := by { + lemma nat_log_eq_padicvalnat_iff {p : ℕ} [hp : Fact (Nat.Prime p)] (n : ℕ)(hn : 0 < n): + Nat.log p n = padicValNat p n ↔ n < p^(padicValNat p n + 1) := by { rw [Nat.log_eq_iff] have Hdiv: p^padicValNat p n ≤ n := Nat.le_of_dvd hn pow_padicValNat_dvd simp only [and_iff_right_iff_imp, Hdiv] @@ -96,20 +89,6 @@ namespace padicValNat linarith } - -- lemma nat_log_eq_padicvalnat_iff' {p : ℕ} [hp : Fact (Nat.Prime p)] (n : ℕ)(hn : 0 < n): Nat.log p n = padicValNat p n ↔ n < p^(Nat.log p n + 1) := by { - -- rw [Nat.log_eq_iff] - -- have Hdiv: p^padicValNat p n ≤ n := Nat.le_of_dvd hn pow_padicValNat_dvd - -- simp only [and_iff_right_iff_imp, Hdiv, true_and] - -- suffices : padicValNat p n = Nat.log p n; rw [this] - -- apply Nat.le_antisymm (le_nat_log _) _ - -- have Hp := (Nat.Prime.one_lt' p).out - -- rw [← pow_le_pow_iff Hp] - - -- -- right - -- -- refine' ⟨(Nat.Prime.one_lt' p).out,_⟩ - -- -- linarith - -- } - end padicValNat def harmonic : ℕ → ℚ @@ -118,9 +97,6 @@ def harmonic : ℕ → ℚ def harmonic' : ℕ → ℚ := fun n => ∑ i in Finset.range n, 1 / (i + 1) -#eval let x := 19; (padicValRat 2 (harmonic x), padicValRat 2 (x)) - - lemma harmonic_harmonic' n : harmonic n = harmonic' n := by { induction' n with k ih ; try simp dsimp [harmonic, harmonic'] @@ -149,6 +125,17 @@ lemma harmonic_ne_zero : ∀ n, n ≠ 0 → harmonic n > 0 := by { } } +lemma harmonic_singleton {n c : ℕ} (hc : c ∈ Finset.range n): harmonic n =1 / ((c + 1):ℚ) + ∑ x in Finset.range n \ {c}, 1 / ((x : ℚ) + 1) := by { + rw [add_comm,harmonic_harmonic'] + unfold harmonic' + rwa [Finset.sum_eq_sum_diff_singleton_add (i := c)] +} + + +lemma harmonic_singleton_ne_zero {c n : ℕ} (hn : 1 ≤ n)(hc : c ∈ Finset.range n) : (∑ x in Finset.range n \ {c}, 1 / ((x + 1 : ℚ)) ≠ 0) := by { + sorry +} + theorem not_int_of_not_padic_int (a : ℚ) : padicNorm 2 a > 1 → ¬ a.isInt := by { intro H @@ -185,11 +172,6 @@ lemma padicValRat_2_pow (r : ℕ) : padicValRat 2 (1 / 2^r) = -r := by { norm_num } -lemma padicValRat_ge_neg_Nat_log {p n : ℕ}[hp : Fact (Nat.Prime p)]: ∀ q, q ≤ n → -Nat.log p n ≤ padicValRat p (1 / q) := by { - intros q Hq - rw [one_div,padicValRat.inv,padicValRat.of_nat, neg_le_neg_iff, Nat.cast_le] - apply padicValNat.le_nat_log_gen Hq -} lemma nat_log_not_dvd {n : ℕ} : ∀ i, 0 < i ∧ i ≠ 2^(Nat.log 2 n) ∧ i ≤ n → ¬ 2^(Nat.log 2 n) ∣ i := by { rintro i ⟨Hi₁,Hi₂,Hi₃⟩ @@ -228,7 +210,8 @@ lemma nat_log_not_dvd {n : ℕ} : ∀ i, 0 < i ∧ i ≠ 2^(Nat.log 2 n) ∧ i } -lemma padicValRat_ge_neg_Nat_log' {n : ℕ}: ∀ q, 0 < q ∧ q ≤ n → q ≠ 2^Nat.log 2 n → padicValRat 2 (1 / q) ≠ -Nat.log 2 n := by { +lemma padicValRat_ge_neg_Nat_log_ne {n : ℕ}: +∀ q, 0 < q ∧ q ≤ n → q ≠ 2^Nat.log 2 n → padicValRat 2 (1 / q) ≠ -Nat.log 2 n := by { rintro q ⟨Hq₁,Hq₂⟩ Hq₃ rw [one_div,padicValRat.inv,padicValRat.of_nat] simp only [ne_eq, neg_inj, Nat.cast_inj] @@ -238,6 +221,46 @@ lemma padicValRat_ge_neg_Nat_log' {n : ℕ}: ∀ q, 0 < q ∧ q ≤ n → q ≠ exact nat_log_not_dvd (n := n) q ⟨Hq₁,Hq₃,Hq₂⟩ this } +lemma padicValRat_ge_neg_Nat_log_ge {p n : ℕ}[hp : Fact (Nat.Prime p)]: + ∀ q, q ≤ n → -Nat.log p n ≤ padicValRat p (1 / q) := by { + intros q Hq + rw [one_div,padicValRat.inv,padicValRat.of_nat, neg_le_neg_iff, Nat.cast_le] + apply padicValNat.le_nat_log_gen Hq +} + +lemma padicValRat_ge_neg_Nat_log_lt {n : ℕ}: +∀ q, 0 < q ∧ q ≤ n → q ≠ 2^Nat.log 2 n → -Nat.log 2 n < padicValRat 2 (1 / q) := by { + rintro q ⟨Hq₁,Hq₂⟩ Hq₃ + have H₁ := padicValRat_ge_neg_Nat_log_ne q ⟨Hq₁,Hq₂⟩ Hq₃ + have H₂ := padicValRat_ge_neg_Nat_log_ge (p := 2) q Hq₂ + exact lt_of_le_of_ne H₂ H₁.symm +} + +lemma pow2_log_in_finset {n : ℕ} (hn : 2 ≤ n) : 2^(Nat.log 2 n) - 1 ∈ Finset.range n := by { + simp only [ge_iff_le, Finset.mem_range] + have H := Nat.pow_log_le_add_one 2 n + rw [le_iff_lt_or_eq] at H + rcases H with H₁ | H₂ + { + simp only [ge_iff_le, gt_iff_lt] + refine' Nat.sub_lt_right_of_lt_add _ H₁ + calc 1 ≤ 2 := by norm_num + _ = 2^1 := by norm_num + _ <= 2 ^ Nat.log 2 n := by { + refine' Nat.pow_le_pow_of_le_right (by norm_num) _ + rw [←Nat.pow_le_iff_le_log (by norm_num) _] + simpa + linarith + } + } + have Habs : n + 1 ≤ n := by { + rw [← H₂] + apply Nat.pow_log_le_self; linarith + } + linarith +} + + theorem harmonic_not_int : ∀ n, n ≥ 2 → ¬ (harmonic n).isInt := by { intro n Hn apply not_int_of_not_padic_int @@ -257,41 +280,18 @@ theorem harmonic_not_int : ∀ n, n ≥ 2 → ¬ (harmonic n).isInt := by { rw [Hlog] simp only [Int.log_natCast, Left.neg_neg_iff, Nat.cast_pos, Nat.log_pos_iff, and_true, Hn] simp only [Int.log_natCast] - - sorry - + rw [harmonic_singleton (pow2_log_in_finset Hn)] + simp only [ge_iff_le, Finset.mem_range, not_lt, Finset.singleton_subset_iff, gt_iff_lt, pow_pos, + Nat.cast_pred, Nat.cast_pow, Nat.cast_ofNat, sub_add_cancel] + rw [padicValRat.sum_eq_of_lt (p := 2) _ (one_div_ne_zero <| pow_ne_zero _ two_ne_zero) (harmonic_singleton_ne_zero (le_trans (one_le_two) Hn) (pow2_log_in_finset Hn)) _]; apply padicValRat_2_pow + { + rw [harmonic_singleton (pow2_log_in_finset Hn)] at h + simpa only [ge_iff_le, gt_iff_lt, pow_pos, Nat.cast_pred, Nat.cast_pow, Nat.cast_ofNat, sub_add_cancel, Finset.mem_range, not_lt, Finset.singleton_subset_iff] using h + } + { + rw [padicValRat_2_pow] + -- refine' lt_of_le_of_lt _ _ + sorry + } } - } - - --- lemma padicValRat_2_sum {r n : ℕ} (hr₁ : 2^r < n)(hr₂ : n < 2^(r+1)) : padicValRat 2 (1 / 2^r + 1 / n) = -r := by { - - -- have Hr : r = Nat.log 2 n := by { - -- symm - -- rw [Nat.log_eq_iff] - -- exact ⟨le_of_lt hr₁,hr₂⟩ - -- right - -- constructor; trivial - -- rw [← Nat.pos_iff_ne_zero] - -- eapply lt_of_lt_of_le _ (le_of_lt hr₁) - -- apply pow_pos; trivial - -- } - -- rw [padicValRat.min_eq_padicValRat] - -- { - -- rw [min_eq_left] - -- apply padicValRat_2_pow - -- rw [padicValRat_2_pow, one_div,padicValRat.inv,padicValRat.of_nat, neg_le_neg_iff, Nat.cast_le, Hr] - -- apply padicValNat.le_nat_log - -- } - -- sorry - -- sorry - -- sorry - - -- rw [padicValRat_2_pow, one_div, padicValRat.inv,padicValRat.of_nat] - -- intro Hnot - -- simp only [padicValRat.of_nat, neg_inj, Nat.cast_inj] at Hnot - -- rw [Hnot] at hr₁ - -- have Hdvd := pow_padicValNat_dvd (p := 2) (n := n) - -- sorry --- } From e4538fbc5e67561408dd7ab43f928687ae417cba Mon Sep 17 00:00:00 2001 From: Koundinya Vajjha Date: Sun, 17 Sep 2023 22:10:42 -0700 Subject: [PATCH 16/74] wip --- Counterexamples/HarmonicNotInteger.lean | 35 ++++++++++++++++++++++--- 1 file changed, 31 insertions(+), 4 deletions(-) diff --git a/Counterexamples/HarmonicNotInteger.lean b/Counterexamples/HarmonicNotInteger.lean index cffdec0d26f0a..08277204f850e 100644 --- a/Counterexamples/HarmonicNotInteger.lean +++ b/Counterexamples/HarmonicNotInteger.lean @@ -58,6 +58,32 @@ namespace padicValRat (hval : padicValRat p q < padicValRat p r) : padicValRat p (q + r) = padicValRat p q := by rw [min_eq_padicValRat hqr hq hr (ne_of_lt hval),min_eq_left (le_of_lt hval)] +-- ⊢ padicValRat 2 (1 / 2 ^ Nat.log 2 n) < padicValRat 2 (∑ x in Finset.range n \ {2 ^ Nat.log 2 n - 1}, 1 / (↑x + 1)) + + theorem sum_ge_of_ge {p : ℕ} [hp : Fact (Nat.Prime p)] {n j : ℕ} {F : ℕ → ℚ} (hF : ∀ i, i ≤ n → padicValRat p (F j) ≤ padicValRat p (F i)) (hn0 : ∑ i in Finset.range n, F i ≠ 0) : padicValRat p (F j) ≤ padicValRat p (∑ i in Finset.range n, F i) := by + induction' n with d hd + · exact False.elim (hn0 rfl) + · rw [Finset.sum_range_succ] at hn0 ⊢ + by_cases h : ∑ x : ℕ in Finset.range d, F x = 0 + · rw [h, zero_add] + exact hF d (by linarith) + · refine' le_trans _ (padicValRat.min_le_padicValRat_add (p := p) hn0) + · exact le_min (hd (fun i hi => hF _ (by linarith)) h) (hF _ (by linarith)) + + theorem finset_sum_eq_of_lt {p : ℕ} [hp : Fact (Nat.Prime p)] {n j : ℕ} {F : ℕ → ℚ} (hF : ∀ i, i ≤ n → i ≠ j → padicValRat p (F j) < padicValRat p (F i)) (hn0 : ∑ i in Finset.range n, F i ≠ 0) : padicValRat p (F j) < padicValRat p (∑ i in Finset.range n \ {j}, F i) := by + { + have hF' : (∀ i, i ≤ n → padicValRat p (F j) ≤ padicValRat p (F i)) := sorry + refine' lt_of_le_of_lt (sum_ge_of_ge hF' hn0) _ + sorry + } + -- induction' n with d hd + -- · exact False.elim (hn0 rfl) + -- · rw [Finset.sum_range_succ] at hn0 ⊢ + -- by_cases h : ∑ x : ℕ in Finset.range d, F x = 0 + -- · rw [h, zero_add] + -- exact hF d (by linarith) + -- · refine' le_trans _ (padicValRat.min_le_padicValRat_add (p := p) hn0) + -- · exact le_min (hd (fun i hi => hF _ (by linarith)) h) (hF _ (by linarith)) end padicValRat @@ -228,11 +254,12 @@ lemma padicValRat_ge_neg_Nat_log_ge {p n : ℕ}[hp : Fact (Nat.Prime p)]: apply padicValNat.le_nat_log_gen Hq } -lemma padicValRat_ge_neg_Nat_log_lt {n : ℕ}: -∀ q, 0 < q ∧ q ≤ n → q ≠ 2^Nat.log 2 n → -Nat.log 2 n < padicValRat 2 (1 / q) := by { +lemma padicValRat_ge_neg_Nat_log_lt (n : ℕ): +∀ q, 0 < q ∧ q ≤ n → q ≠ 2^Nat.log 2 n → padicValRat 2 (1 / 2^Nat.log 2 n) < padicValRat 2 (1 / q) := by { rintro q ⟨Hq₁,Hq₂⟩ Hq₃ have H₁ := padicValRat_ge_neg_Nat_log_ne q ⟨Hq₁,Hq₂⟩ Hq₃ have H₂ := padicValRat_ge_neg_Nat_log_ge (p := 2) q Hq₂ + rw [padicValRat_2_pow] exact lt_of_le_of_ne H₂ H₁.symm } @@ -289,8 +316,8 @@ theorem harmonic_not_int : ∀ n, n ≥ 2 → ¬ (harmonic n).isInt := by { simpa only [ge_iff_le, gt_iff_lt, pow_pos, Nat.cast_pred, Nat.cast_pow, Nat.cast_ofNat, sub_add_cancel, Finset.mem_range, not_lt, Finset.singleton_subset_iff] using h } { - rw [padicValRat_2_pow] - -- refine' lt_of_le_of_lt _ _ + + -- refine ?_ sorry } } From 8ec3340f7979e2f659ad6345da7ccb0b85962e03 Mon Sep 17 00:00:00 2001 From: Koundinya Vajjha Date: Tue, 19 Sep 2023 18:58:48 -0700 Subject: [PATCH 17/74] progress? --- Counterexamples/HarmonicNotInteger.lean | 17 ++++++++++++----- 1 file changed, 12 insertions(+), 5 deletions(-) diff --git a/Counterexamples/HarmonicNotInteger.lean b/Counterexamples/HarmonicNotInteger.lean index 08277204f850e..6d041b217dcc8 100644 --- a/Counterexamples/HarmonicNotInteger.lean +++ b/Counterexamples/HarmonicNotInteger.lean @@ -58,6 +58,12 @@ namespace padicValRat (hval : padicValRat p q < padicValRat p r) : padicValRat p (q + r) = padicValRat p q := by rw [min_eq_padicValRat hqr hq hr (ne_of_lt hval),min_eq_left (le_of_lt hval)] + lemma sum_lt_of_lt {p : ℕ} [hp : Fact (Nat.Prime p)] {q r₁ r₂ : ℚ} (hqr : r₁ + r₂ ≠ 0) + (hval₁ : padicValRat p q < padicValRat p r₁) (hval₂ : padicValRat p q < padicValRat p r₂): + padicValRat p q < padicValRat p (r₁ + r₂) := by { + have H₁ : padicValRat p q < min (padicValRat p r₁) (padicValRat p r₂) := lt_min hval₁ hval₂ + refine' lt_of_lt_of_le H₁ (padicValRat.min_le_padicValRat_add (p := p) hqr) + } -- ⊢ padicValRat 2 (1 / 2 ^ Nat.log 2 n) < padicValRat 2 (∑ x in Finset.range n \ {2 ^ Nat.log 2 n - 1}, 1 / (↑x + 1)) theorem sum_ge_of_ge {p : ℕ} [hp : Fact (Nat.Prime p)] {n j : ℕ} {F : ℕ → ℚ} (hF : ∀ i, i ≤ n → padicValRat p (F j) ≤ padicValRat p (F i)) (hn0 : ∑ i in Finset.range n, F i ≠ 0) : padicValRat p (F j) ≤ padicValRat p (∑ i in Finset.range n, F i) := by @@ -70,11 +76,11 @@ namespace padicValRat · refine' le_trans _ (padicValRat.min_le_padicValRat_add (p := p) hn0) · exact le_min (hd (fun i hi => hF _ (by linarith)) h) (hF _ (by linarith)) - theorem finset_sum_eq_of_lt {p : ℕ} [hp : Fact (Nat.Prime p)] {n j : ℕ} {F : ℕ → ℚ} (hF : ∀ i, i ≤ n → i ≠ j → padicValRat p (F j) < padicValRat p (F i)) (hn0 : ∑ i in Finset.range n, F i ≠ 0) : padicValRat p (F j) < padicValRat p (∑ i in Finset.range n \ {j}, F i) := by + theorem finset_sum_eq_of_lt {p : ℕ} [hp : Fact (Nat.Prime p)] {n j : ℕ} {F : ℕ → ℚ} (hF : ∀ i, i ≤ n → i ≠ j → padicValRat p (F j) < padicValRat p (F i)) (hn0 : ∑ i in Finset.range n, F i ≠ 0) (hj₁ : padicValRat p (F j) < 0) (hj₂ : j ∈ Finset.range n): padicValRat p (F j) < padicValRat p (∑ i in Finset.range n \ {j}, F i) := by { - have hF' : (∀ i, i ≤ n → padicValRat p (F j) ≤ padicValRat p (F i)) := sorry - refine' lt_of_le_of_lt (sum_ge_of_ge hF' hn0) _ - sorry + induction' j with d hd + simp at * + } -- induction' n with d hd -- · exact False.elim (hn0 rfl) @@ -190,6 +196,7 @@ theorem not_int_of_not_padic_int (a : ℚ) : } } +@[simp] lemma padicValRat_2_pow (r : ℕ) : padicValRat 2 (1 / 2^r) = -r := by { rw [one_div,padicValRat.inv,neg_inj,padicValRat.pow (by simp)] suffices : padicValRat 2 2 = 1 @@ -316,8 +323,8 @@ theorem harmonic_not_int : ∀ n, n ≥ 2 → ¬ (harmonic n).isInt := by { simpa only [ge_iff_le, gt_iff_lt, pow_pos, Nat.cast_pred, Nat.cast_pow, Nat.cast_ofNat, sub_add_cancel, Finset.mem_range, not_lt, Finset.singleton_subset_iff] using h } { + have := padicValRat_ge_neg_Nat_log_lt n - -- refine ?_ sorry } } From 60ceb83c8475577ead368c7200bc273e71d4ed94 Mon Sep 17 00:00:00 2001 From: Koundinya Vajjha Date: Tue, 19 Sep 2023 22:58:28 -0700 Subject: [PATCH 18/74] more lemmas, maybe progress --- Counterexamples/HarmonicNotInteger.lean | 60 ++++++++++++++++++++----- 1 file changed, 49 insertions(+), 11 deletions(-) diff --git a/Counterexamples/HarmonicNotInteger.lean b/Counterexamples/HarmonicNotInteger.lean index 6d041b217dcc8..8d90603c0a5e2 100644 --- a/Counterexamples/HarmonicNotInteger.lean +++ b/Counterexamples/HarmonicNotInteger.lean @@ -59,14 +59,11 @@ namespace padicValRat padicValRat p (q + r) = padicValRat p q := by rw [min_eq_padicValRat hqr hq hr (ne_of_lt hval),min_eq_left (le_of_lt hval)] lemma sum_lt_of_lt {p : ℕ} [hp : Fact (Nat.Prime p)] {q r₁ r₂ : ℚ} (hqr : r₁ + r₂ ≠ 0) - (hval₁ : padicValRat p q < padicValRat p r₁) (hval₂ : padicValRat p q < padicValRat p r₂): - padicValRat p q < padicValRat p (r₁ + r₂) := by { - have H₁ : padicValRat p q < min (padicValRat p r₁) (padicValRat p r₂) := lt_min hval₁ hval₂ - refine' lt_of_lt_of_le H₁ (padicValRat.min_le_padicValRat_add (p := p) hqr) - } + (hval₁ : padicValRat p q < padicValRat p r₁) (hval₂ : padicValRat p q < padicValRat p r₂): padicValRat p q < padicValRat p (r₁ + r₂) := lt_of_lt_of_le (lt_min hval₁ hval₂) (padicValRat.min_le_padicValRat_add (p := p) (hp := hp) hqr) + -- ⊢ padicValRat 2 (1 / 2 ^ Nat.log 2 n) < padicValRat 2 (∑ x in Finset.range n \ {2 ^ Nat.log 2 n - 1}, 1 / (↑x + 1)) - theorem sum_ge_of_ge {p : ℕ} [hp : Fact (Nat.Prime p)] {n j : ℕ} {F : ℕ → ℚ} (hF : ∀ i, i ≤ n → padicValRat p (F j) ≤ padicValRat p (F i)) (hn0 : ∑ i in Finset.range n, F i ≠ 0) : padicValRat p (F j) ≤ padicValRat p (∑ i in Finset.range n, F i) := by + theorem finset_sum_ge_of_ge {p : ℕ} [hp : Fact (Nat.Prime p)] {n j : ℕ} {F : ℕ → ℚ} (hF : ∀ i, i ≤ n → padicValRat p (F j) ≤ padicValRat p (F i)) (hn0 : ∑ i in Finset.range n, F i ≠ 0) : padicValRat p (F j) ≤ padicValRat p (∑ i in Finset.range n, F i) := by induction' n with d hd · exact False.elim (hn0 rfl) · rw [Finset.sum_range_succ] at hn0 ⊢ @@ -76,10 +73,50 @@ namespace padicValRat · refine' le_trans _ (padicValRat.min_le_padicValRat_add (p := p) hn0) · exact le_min (hd (fun i hi => hF _ (by linarith)) h) (hF _ (by linarith)) - theorem finset_sum_eq_of_lt {p : ℕ} [hp : Fact (Nat.Prime p)] {n j : ℕ} {F : ℕ → ℚ} (hF : ∀ i, i ≤ n → i ≠ j → padicValRat p (F j) < padicValRat p (F i)) (hn0 : ∑ i in Finset.range n, F i ≠ 0) (hj₁ : padicValRat p (F j) < 0) (hj₂ : j ∈ Finset.range n): padicValRat p (F j) < padicValRat p (∑ i in Finset.range n \ {j}, F i) := by + theorem finset_sum_lt_of_lt {p : ℕ} [hp : Fact (Nat.Prime p)] {n j : ℕ} {F : ℕ → ℚ} (hF : ∀ i, i ≤ n → padicValRat p (F j) < padicValRat p (F i)) (hn0 : ∑ i in Finset.range n, F i ≠ 0) : padicValRat p (F j) < padicValRat p (∑ i in Finset.range n, F i) := by + induction' n with d hd + · exact False.elim (hn0 rfl) + · rw [Finset.sum_range_succ] at hn0 ⊢ + by_cases h : ∑ x : ℕ in Finset.range d, F x = 0 + · rw [h, zero_add] + exact hF d (by linarith) + · refine' le_trans _ (padicValRat.min_le_padicValRat_add (p := p) hn0) + · exact le_min (hd (fun i hi => hF _ (by linarith)) h) (hF _ (by linarith)) + + theorem finset_gen_sum_lt_of_lt {p : ℕ} [hp : Fact (Nat.Prime p)] {j : ℕ} {F : ℕ → ℚ} (S : Finset ℕ) (hS : S.Nonempty) (hF : ∀ i, i ∈ S → padicValRat p (F j) < padicValRat p (F i)) + (hn1 : ∀ i : ℕ, 0 < F i): padicValRat p (F j) < padicValRat p (∑ i in S, F i) := by { + induction' hS using Finset.Nonempty.cons_induction with k s S' Hnot Hne Hind + simp only [Finset.sum_singleton] + exact hF k (by simp) + { + rw [Finset.cons_eq_insert, Finset.sum_insert Hnot] + refine' sum_lt_of_lt (_) (hF _ (by simp [Finset.mem_insert, true_or])) (Hind (fun i hi => hF _ (by rw [Finset.cons_eq_insert,Finset.mem_insert]; exact Or.inr hi))) + sorry + } + } + + theorem finset_gen_sum_lt_of_lt' {p : ℕ} [hp : Fact (Nat.Prime p)] {j : ℕ} {F : ℕ → ℚ} (S : Finset ℕ) (hS : S.Nonempty) (hF : ∀ i, i ∈ S → padicValRat p (F j) < padicValRat p (F i)) + (hn1 : ∀ i : ℕ, 0 < F i): padicValRat p (F j) < padicValRat p (∑ i in S, F i) := by { + -- have hn0 : 0 < ∑ i in S, F i := Finset.sum_pos hn1 hS + induction' S using Finset.induction_on with k S' hS' H + simp only [ne_eq, Finset.not_mem_empty, IsEmpty.forall_iff, forall_const, not_false_eq_true, Finset.sum_empty] at * + rw [Finset.sum_insert hS'] + refine' sum_lt_of_lt _ (hF _ (by simp [Finset.mem_insert, true_or])) (H _ (fun i hi => hF _ (Finset.mem_insert.mpr (Or.inr hi))) ) + { + sorry + } + { + sorry + } + -- simp at hj + -- push_neg at hj + -- exact hj.2 + } + + theorem finset_sum_eq_of_lt {p : ℕ} [hp : Fact (Nat.Prime p)] {n j : ℕ} {F : ℕ → ℚ} (hF : ∀ i, i ≤ n → i ≠ j → padicValRat p (F j) < padicValRat p (F i)) (hn0 : ∑ i in Finset.range n \ {j}, F i ≠ 0) (hj₂ : j ∈ Finset.range n): padicValRat p (F j) < padicValRat p (∑ i in Finset.range n \ {j}, F i) := by { - induction' j with d hd - simp at * + + } -- induction' n with d hd @@ -323,9 +360,10 @@ theorem harmonic_not_int : ∀ n, n ≥ 2 → ¬ (harmonic n).isInt := by { simpa only [ge_iff_le, gt_iff_lt, pow_pos, Nat.cast_pred, Nat.cast_pow, Nat.cast_ofNat, sub_add_cancel, Finset.mem_range, not_lt, Finset.singleton_subset_iff] using h } { - have := padicValRat_ge_neg_Nat_log_lt n + refine' padicValRat.finset_gen_sum_lt_of_lt (p := 2) _ _ _ _ + -- have := padicValRat_ge_neg_Nat_log_lt n - sorry + -- sorry } } } From 761aee7aed65552f1ff21197e4f0dede0b8482d4 Mon Sep 17 00:00:00 2001 From: Koundinya Vajjha Date: Wed, 20 Sep 2023 18:35:39 -0700 Subject: [PATCH 19/74] sorry free proof of harmonic not int --- Counterexamples/HarmonicNotInteger.lean | 87 ++++++++++++------------- 1 file changed, 42 insertions(+), 45 deletions(-) diff --git a/Counterexamples/HarmonicNotInteger.lean b/Counterexamples/HarmonicNotInteger.lean index 8d90603c0a5e2..5ef7fca7166b9 100644 --- a/Counterexamples/HarmonicNotInteger.lean +++ b/Counterexamples/HarmonicNotInteger.lean @@ -83,51 +83,22 @@ namespace padicValRat · refine' le_trans _ (padicValRat.min_le_padicValRat_add (p := p) hn0) · exact le_min (hd (fun i hi => hF _ (by linarith)) h) (hF _ (by linarith)) - theorem finset_gen_sum_lt_of_lt {p : ℕ} [hp : Fact (Nat.Prime p)] {j : ℕ} {F : ℕ → ℚ} (S : Finset ℕ) (hS : S.Nonempty) (hF : ∀ i, i ∈ S → padicValRat p (F j) < padicValRat p (F i)) - (hn1 : ∀ i : ℕ, 0 < F i): padicValRat p (F j) < padicValRat p (∑ i in S, F i) := by { + theorem finset_gen_sum_lt_of_lt {p : ℕ} [hp : Fact (Nat.Prime p)] {j : ℕ} + {F : ℕ → ℚ} {S : Finset ℕ} (hS : S.Nonempty) + (hF : ∀ i, i ∈ S → padicValRat p (F j) < padicValRat p (F i)) + (hn1 : ∀ i : ℕ, 0 < F i): + padicValRat p (F j) < padicValRat p (∑ i in S, F i) := by { induction' hS using Finset.Nonempty.cons_induction with k s S' Hnot Hne Hind simp only [Finset.sum_singleton] exact hF k (by simp) { rw [Finset.cons_eq_insert, Finset.sum_insert Hnot] - refine' sum_lt_of_lt (_) (hF _ (by simp [Finset.mem_insert, true_or])) (Hind (fun i hi => hF _ (by rw [Finset.cons_eq_insert,Finset.mem_insert]; exact Or.inr hi))) - sorry + refine' sum_lt_of_lt (_) (hF _ (by simp [Finset.mem_insert, true_or])) + (Hind (fun i hi => hF _ (by rw [Finset.cons_eq_insert,Finset.mem_insert]; exact Or.inr hi))) + refine' ne_of_gt (add_pos (hn1 s) (Finset.sum_pos (fun i _ => hn1 i) Hne)) } } - theorem finset_gen_sum_lt_of_lt' {p : ℕ} [hp : Fact (Nat.Prime p)] {j : ℕ} {F : ℕ → ℚ} (S : Finset ℕ) (hS : S.Nonempty) (hF : ∀ i, i ∈ S → padicValRat p (F j) < padicValRat p (F i)) - (hn1 : ∀ i : ℕ, 0 < F i): padicValRat p (F j) < padicValRat p (∑ i in S, F i) := by { - -- have hn0 : 0 < ∑ i in S, F i := Finset.sum_pos hn1 hS - induction' S using Finset.induction_on with k S' hS' H - simp only [ne_eq, Finset.not_mem_empty, IsEmpty.forall_iff, forall_const, not_false_eq_true, Finset.sum_empty] at * - rw [Finset.sum_insert hS'] - refine' sum_lt_of_lt _ (hF _ (by simp [Finset.mem_insert, true_or])) (H _ (fun i hi => hF _ (Finset.mem_insert.mpr (Or.inr hi))) ) - { - sorry - } - { - sorry - } - -- simp at hj - -- push_neg at hj - -- exact hj.2 - } - - theorem finset_sum_eq_of_lt {p : ℕ} [hp : Fact (Nat.Prime p)] {n j : ℕ} {F : ℕ → ℚ} (hF : ∀ i, i ≤ n → i ≠ j → padicValRat p (F j) < padicValRat p (F i)) (hn0 : ∑ i in Finset.range n \ {j}, F i ≠ 0) (hj₂ : j ∈ Finset.range n): padicValRat p (F j) < padicValRat p (∑ i in Finset.range n \ {j}, F i) := by - { - - - - } - -- induction' n with d hd - -- · exact False.elim (hn0 rfl) - -- · rw [Finset.sum_range_succ] at hn0 ⊢ - -- by_cases h : ∑ x : ℕ in Finset.range d, F x = 0 - -- · rw [h, zero_add] - -- exact hF d (by linarith) - -- · refine' le_trans _ (padicValRat.min_le_padicValRat_add (p := p) hn0) - -- · exact le_min (hd (fun i hi => hF _ (by linarith)) h) (hF _ (by linarith)) - end padicValRat namespace padicValNat @@ -200,9 +171,20 @@ lemma harmonic_singleton {n c : ℕ} (hc : c ∈ Finset.range n): harmonic n =1 rwa [Finset.sum_eq_sum_diff_singleton_add (i := c)] } +lemma finset_range_sdiff_singleton_nonempty {c n : ℕ} (hn : 2 ≤ n) : Finset.Nonempty (Finset.range n \ {c}) := by { + rw [Finset.sdiff_nonempty,Finset.subset_singleton_iff, Finset.range_eq_empty_iff] + push_neg + refine' ⟨by linarith,fun Hnot => _⟩ + suffices : n = 1; linarith + rw [← Finset.card_range n, ← Finset.card_singleton c, Hnot] +} -lemma harmonic_singleton_ne_zero {c n : ℕ} (hn : 1 ≤ n)(hc : c ∈ Finset.range n) : (∑ x in Finset.range n \ {c}, 1 / ((x + 1 : ℚ)) ≠ 0) := by { - sorry +lemma harmonic_singleton_ne_zero {c n : ℕ} (hn : 2 ≤ n) : (∑ x in Finset.range n \ {c}, 1 / ((x + 1 : ℚ)) ≠ 0) := by { + refine' ne_of_gt _ + refine' Finset.sum_pos (fun i _ => div_pos zero_lt_one _) _ + norm_cast + simp only [add_pos_iff, or_true] + apply finset_range_sdiff_singleton_nonempty hn } theorem not_int_of_not_padic_int (a : ℚ) : @@ -344,7 +326,7 @@ theorem harmonic_not_int : ∀ n, n ≥ 2 → ¬ (harmonic n).isInt := by { linarith } { - apply one_lt_zpow; try simp + apply one_lt_zpow; try norm_num apply lt_neg.mp rw [neg_zero] suffices Hlog : padicValRat 2 (harmonic n) = -Int.log 2 (n : ℚ) @@ -354,16 +336,31 @@ theorem harmonic_not_int : ∀ n, n ≥ 2 → ¬ (harmonic n).isInt := by { rw [harmonic_singleton (pow2_log_in_finset Hn)] simp only [ge_iff_le, Finset.mem_range, not_lt, Finset.singleton_subset_iff, gt_iff_lt, pow_pos, Nat.cast_pred, Nat.cast_pow, Nat.cast_ofNat, sub_add_cancel] - rw [padicValRat.sum_eq_of_lt (p := 2) _ (one_div_ne_zero <| pow_ne_zero _ two_ne_zero) (harmonic_singleton_ne_zero (le_trans (one_le_two) Hn) (pow2_log_in_finset Hn)) _]; apply padicValRat_2_pow + rw [padicValRat.sum_eq_of_lt (p := 2) _ (one_div_ne_zero <| pow_ne_zero _ two_ne_zero) (harmonic_singleton_ne_zero (Hn)) _]; apply padicValRat_2_pow { rw [harmonic_singleton (pow2_log_in_finset Hn)] at h simpa only [ge_iff_le, gt_iff_lt, pow_pos, Nat.cast_pred, Nat.cast_pow, Nat.cast_ofNat, sub_add_cancel, Finset.mem_range, not_lt, Finset.singleton_subset_iff] using h } { - refine' padicValRat.finset_gen_sum_lt_of_lt (p := 2) _ _ _ _ - -- have := padicValRat_ge_neg_Nat_log_lt n - - -- sorry + have := @padicValRat.finset_gen_sum_lt_of_lt 2 _ (2 ^ Nat.log 2 n - 1) (fun n => 1 / ((n + 1) : ℚ)) (Finset.range n \ {2^Nat.log 2 n - 1}) + simp at * + apply this + { + apply finset_range_sdiff_singleton_nonempty Hn + } + { + intros i Hi₁ Hi₂ + have := padicValRat_ge_neg_Nat_log_lt (n := n) (i+1) + simp at * + refine' this (by linarith) _ + intro Hnot + rw [←Hnot] at Hi₂ + simp only [ge_iff_le, add_le_iff_nonpos_left, nonpos_iff_eq_zero, add_tsub_cancel_right, not_true] at Hi₂ + } + { + intros i + norm_cast; linarith + } } } } From 7cac0afcaeb5798bde12b6a26d3de33bfae8b2a5 Mon Sep 17 00:00:00 2001 From: Koundinya Vajjha Date: Wed, 20 Sep 2023 22:00:33 -0700 Subject: [PATCH 20/74] clean up some proofs --- Counterexamples/HarmonicNotInteger.lean | 202 ++++++++++-------------- 1 file changed, 81 insertions(+), 121 deletions(-) diff --git a/Counterexamples/HarmonicNotInteger.lean b/Counterexamples/HarmonicNotInteger.lean index 5ef7fca7166b9..d34ceac9961f5 100644 --- a/Counterexamples/HarmonicNotInteger.lean +++ b/Counterexamples/HarmonicNotInteger.lean @@ -28,146 +28,109 @@ open BigOperators namespace padicValRat - lemma min_eq_padicValRat {p : ℕ} [hp : Fact (Nat.Prime p)] {q r : ℚ} (hqr : q + r ≠ 0) (hq : q ≠ 0) (hr : r ≠ 0) - (hval : padicValRat p q ≠ padicValRat p r) : - padicValRat p (q + r) = min (padicValRat p q) (padicValRat p r) := by { +-- Keep + lemma min_eq_padicValRat_add {p : ℕ} [hp : Fact (Nat.Prime p)] {q r : ℚ} + (hqr : q + r ≠ 0) (hq : q ≠ 0) (hr : r ≠ 0) (hval : padicValRat p q ≠ padicValRat p r) : + padicValRat p (q + r) = min (padicValRat p q) (padicValRat p r) := by have Hmin := padicValRat.min_le_padicValRat_add (p := p) (hp := hp) hqr - wlog h : padicValRat p q < padicValRat p r generalizing q r with H - { - push_neg at h - rw [add_comm, min_comm] - refine' (H _ hr hq hval.symm _ _) - rwa [add_comm] - rwa [min_comm,add_comm] - exact Ne.lt_of_le (Ne.symm hval) h - } - { - rw [min_eq_left (le_of_lt h)] at Hmin ⊢ + wlog h : padicValRat p q < padicValRat p r generalizing q r with Hgen + · push_neg at h; rw [add_comm, min_comm] + exact (Hgen (by rwa [add_comm]) hr hq hval.symm + (by rwa [min_comm,add_comm]) + (Ne.lt_of_le (Ne.symm hval) h)) + · rw [min_eq_left (le_of_lt h)] at Hmin ⊢ suffices Hreq : padicValRat p q ≥ padicValRat p (q + r) by linarith - have Haux : padicValRat p q ≥ min (padicValRat p (q + r)) (padicValRat p r) := by { + suffices Haux : padicValRat p q ≥ min (padicValRat p (q + r)) (padicValRat p r) by + rw [min_def] at Haux; split_ifs at Haux with Hspl; try assumption; linarith calc padicValRat p q = padicValRat p ((q + r) - r) := by congr; simp - _ ≥ min (padicValRat p (q + r)) (padicValRat p (-r)) := ge_iff_le.mp <| le_trans (padicValRat.min_le_padicValRat_add (q := q+r) (r := -r) (by simpa)) (by rw [add_neg_cancel_right, add_sub_cancel]) + _ ≥ min (padicValRat p (q + r)) (padicValRat p (-r)) := + ge_iff_le.mp <| le_trans (padicValRat.min_le_padicValRat_add (q := q+r) (r := -r) (by simpa)) + (by rw [add_neg_cancel_right, add_sub_cancel]) _ = min (padicValRat p (q + r)) (padicValRat p r) := by rw [padicValRat.neg] - } - rw [min_def] at Haux - split_ifs at Haux with Hspl; try assumption - linarith - }} - - lemma sum_eq_of_lt {p : ℕ} [hp : Fact (Nat.Prime p)] {q r : ℚ} (hqr : q + r ≠ 0) (hq : q ≠ 0) (hr : r ≠ 0) - (hval : padicValRat p q < padicValRat p r) : - padicValRat p (q + r) = padicValRat p q := by rw [min_eq_padicValRat hqr hq hr (ne_of_lt hval),min_eq_left (le_of_lt hval)] - - lemma sum_lt_of_lt {p : ℕ} [hp : Fact (Nat.Prime p)] {q r₁ r₂ : ℚ} (hqr : r₁ + r₂ ≠ 0) - (hval₁ : padicValRat p q < padicValRat p r₁) (hval₂ : padicValRat p q < padicValRat p r₂): padicValRat p q < padicValRat p (r₁ + r₂) := lt_of_lt_of_le (lt_min hval₁ hval₂) (padicValRat.min_le_padicValRat_add (p := p) (hp := hp) hqr) - --- ⊢ padicValRat 2 (1 / 2 ^ Nat.log 2 n) < padicValRat 2 (∑ x in Finset.range n \ {2 ^ Nat.log 2 n - 1}, 1 / (↑x + 1)) - - theorem finset_sum_ge_of_ge {p : ℕ} [hp : Fact (Nat.Prime p)] {n j : ℕ} {F : ℕ → ℚ} (hF : ∀ i, i ≤ n → padicValRat p (F j) ≤ padicValRat p (F i)) (hn0 : ∑ i in Finset.range n, F i ≠ 0) : padicValRat p (F j) ≤ padicValRat p (∑ i in Finset.range n, F i) := by - induction' n with d hd - · exact False.elim (hn0 rfl) - · rw [Finset.sum_range_succ] at hn0 ⊢ - by_cases h : ∑ x : ℕ in Finset.range d, F x = 0 - · rw [h, zero_add] - exact hF d (by linarith) - · refine' le_trans _ (padicValRat.min_le_padicValRat_add (p := p) hn0) - · exact le_min (hd (fun i hi => hF _ (by linarith)) h) (hF _ (by linarith)) - - theorem finset_sum_lt_of_lt {p : ℕ} [hp : Fact (Nat.Prime p)] {n j : ℕ} {F : ℕ → ℚ} (hF : ∀ i, i ≤ n → padicValRat p (F j) < padicValRat p (F i)) (hn0 : ∑ i in Finset.range n, F i ≠ 0) : padicValRat p (F j) < padicValRat p (∑ i in Finset.range n, F i) := by - induction' n with d hd - · exact False.elim (hn0 rfl) - · rw [Finset.sum_range_succ] at hn0 ⊢ - by_cases h : ∑ x : ℕ in Finset.range d, F x = 0 - · rw [h, zero_add] - exact hF d (by linarith) - · refine' le_trans _ (padicValRat.min_le_padicValRat_add (p := p) hn0) - · exact le_min (hd (fun i hi => hF _ (by linarith)) h) (hF _ (by linarith)) - - theorem finset_gen_sum_lt_of_lt {p : ℕ} [hp : Fact (Nat.Prime p)] {j : ℕ} + +-- Keep + lemma add_eq_of_lt {p : ℕ} [hp : Fact (Nat.Prime p)] {q r : ℚ} (hqr : q + r ≠ 0) + (hq : q ≠ 0) (hr : r ≠ 0) (hval : padicValRat p q < padicValRat p r) : + padicValRat p (q + r) = padicValRat p q := + by rw [min_eq_padicValRat_add hqr hq hr (ne_of_lt hval),min_eq_left (le_of_lt hval)] + +-- Keep + lemma add_lt_of_lt {p : ℕ} [hp : Fact (Nat.Prime p)] {q r₁ r₂ : ℚ} (hqr : r₁ + r₂ ≠ 0) + (hval₁ : padicValRat p q < padicValRat p r₁) (hval₂ : padicValRat p q < padicValRat p r₂) : + padicValRat p q < padicValRat p (r₁ + r₂) := + lt_of_lt_of_le (lt_min hval₁ hval₂) (padicValRat.min_le_padicValRat_add (p := p) (hp := hp) hqr) + +-- Keep + -- theorem finset_sum_ge_of_ge {p : ℕ} [hp : Fact (Nat.Prime p)] {n j : ℕ} {F : ℕ → ℚ} + -- (hF : ∀ i, i ≤ n → padicValRat p (F j) ≤ padicValRat p (F i)) + -- (hn0 : ∑ i in Finset.range n, F i ≠ 0) : padicValRat p (F j) ≤ + -- padicValRat p (∑ i in Finset.range n, F i) := by + -- induction' n with d hd + -- · exact False.elim (hn0 rfl) + -- · rw [Finset.sum_range_succ] at hn0 ⊢ + -- by_cases h : ∑ x : ℕ in Finset.range d, F x = 0 + -- · rw [h, zero_add] + -- exact hF d (by linarith) + -- · refine' le_trans _ (padicValRat.min_le_padicValRat_add (p := p) hn0) + -- · exact le_min (hd (fun i hi => hF _ (by linarith)) h) (hF _ (by linarith)) + + -- theorem finset_sum_lt_of_lt {p : ℕ} [hp : Fact (Nat.Prime p)] {n j : ℕ} {F : ℕ → ℚ} + -- (hF : ∀ i, i ≤ n → padicValRat p (F j) < padicValRat p (F i)) + -- (hn0 : ∑ i in Finset.range n, F i ≠ 0) : + -- padicValRat p (F j) < padicValRat p (∑ i in Finset.range n, F i) := by + -- induction' n with d hd + -- · exact False.elim (hn0 rfl) + -- · rw [Finset.sum_range_succ] at hn0 ⊢ + -- by_cases h : ∑ x : ℕ in Finset.range d, F x = 0 + -- · rw [h, zero_add] + -- exact hF d (by linarith) + -- · refine' le_trans _ (padicValRat.min_le_padicValRat_add (p := p) hn0) + -- · exact le_min (hd (fun i hi => hF _ (by linarith)) h) (hF _ (by linarith)) + + theorem finset_gen_sum_lt_of_lt {p j : ℕ} [hp : Fact (Nat.Prime p)] {F : ℕ → ℚ} {S : Finset ℕ} (hS : S.Nonempty) (hF : ∀ i, i ∈ S → padicValRat p (F j) < padicValRat p (F i)) (hn1 : ∀ i : ℕ, 0 < F i): - padicValRat p (F j) < padicValRat p (∑ i in S, F i) := by { + padicValRat p (F j) < padicValRat p (∑ i in S, F i) := by induction' hS using Finset.Nonempty.cons_induction with k s S' Hnot Hne Hind - simp only [Finset.sum_singleton] - exact hF k (by simp) - { - rw [Finset.cons_eq_insert, Finset.sum_insert Hnot] - refine' sum_lt_of_lt (_) (hF _ (by simp [Finset.mem_insert, true_or])) - (Hind (fun i hi => hF _ (by rw [Finset.cons_eq_insert,Finset.mem_insert]; exact Or.inr hi))) - refine' ne_of_gt (add_pos (hn1 s) (Finset.sum_pos (fun i _ => hn1 i) Hne)) - } - } + · rw [Finset.sum_singleton] + exact hF k (by simp) + · rw [Finset.cons_eq_insert, Finset.sum_insert Hnot] + exact add_lt_of_lt + (ne_of_gt (add_pos (hn1 s) (Finset.sum_pos (fun i _ => hn1 i) Hne))) + (hF _ (by simp [Finset.mem_insert, true_or])) + (Hind (fun i hi => hF _ (by rw [Finset.cons_eq_insert,Finset.mem_insert]; exact Or.inr hi))) end padicValRat namespace padicValNat lemma le_nat_log {p : ℕ} [hp : Fact (Nat.Prime p)] (n : ℕ): - padicValNat p n ≤ Nat.log p n := by { - by_cases (n = 0); simp [h] - apply Nat.le_log_of_pow_le (Nat.Prime.one_lt hp.elim) - by_contra Hnot - have H₁ := Nat.eq_zero_of_dvd_of_lt (@pow_padicValNat_dvd p n) - push_neg at Hnot - apply h - apply H₁ - exact Hnot - } + padicValNat p n ≤ Nat.log p n := by + by_cases (n = 0) + · simp only [h, padicValNat.zero, Nat.log_zero_right, le_refl] + · refine' Nat.le_log_of_pow_le (Nat.Prime.one_lt hp.elim) _ + by_contra Hnot; push_neg at Hnot + exact h (Nat.eq_zero_of_dvd_of_lt (pow_padicValNat_dvd (p := p) (n := n)) Hnot) lemma le_nat_log_gen {p n₁ n₂ : ℕ} [Fact (Nat.Prime p)] (hn : n₁ ≤ n₂): padicValNat p n₁ ≤ Nat.log p n₂ := le_trans (le_nat_log n₁) (Nat.log_mono_right hn) lemma nat_log_eq_padicvalnat_iff {p : ℕ} [hp : Fact (Nat.Prime p)] (n : ℕ)(hn : 0 < n): - Nat.log p n = padicValNat p n ↔ n < p^(padicValNat p n + 1) := by { - rw [Nat.log_eq_iff] - have Hdiv: p^padicValNat p n ≤ n := Nat.le_of_dvd hn pow_padicValNat_dvd - simp only [and_iff_right_iff_imp, Hdiv] - intros; trivial - right - refine' ⟨(Nat.Prime.one_lt' p).out,_⟩ - linarith - } + Nat.log p n = padicValNat p n ↔ n < p^(padicValNat p n + 1) := by + · rw [Nat.log_eq_iff (Or.inr ⟨(Nat.Prime.one_lt' p).out, by linarith⟩)] + · rw [and_iff_right_iff_imp] + exact (fun _ => Nat.le_of_dvd hn pow_padicValNat_dvd) end padicValNat -def harmonic : ℕ → ℚ -| 0 => 0 -| (k+1) => 1 / (k+1) + harmonic k - -def harmonic' : ℕ → ℚ := fun n => ∑ i in Finset.range n, 1 / (i + 1) +def harmonic : ℕ → ℚ := fun n => ∑ i in Finset.range n, 1 / (i + 1) -lemma harmonic_harmonic' n : harmonic n = harmonic' n := by { - induction' n with k ih ; try simp - dsimp [harmonic, harmonic'] - rw [Finset.sum_range_succ, ih] - dsimp [harmonic'] - rw [add_comm] -} - -lemma harmonic_ne_zero : ∀ n, n ≠ 0 → harmonic n > 0 := by { - intros n Hn - induction' n with k ih; try contradiction - dsimp [harmonic] - by_cases (k = 0) - { - rw [h] - dsimp [harmonic] - norm_num - } - { - specialize (ih h) - have H₀ : (0 : ℚ) = 0 + 0 := by rfl - rw [H₀] - refine' add_lt_add _ ih - apply div_pos; try norm_num - norm_cast; simp only [add_pos_iff, or_true] - } -} +lemma harmonic_ne_zero : ∀ n, n ≠ 0 → harmonic n > 0 := fun n Hn => Finset.sum_pos (fun i _ => div_pos zero_lt_one (by norm_cast; linarith)) (by (rwa [Finset.nonempty_range_iff])) lemma harmonic_singleton {n c : ℕ} (hc : c ∈ Finset.range n): harmonic n =1 / ((c + 1):ℚ) + ∑ x in Finset.range n \ {c}, 1 / ((x : ℚ) + 1) := by { - rw [add_comm,harmonic_harmonic'] - unfold harmonic' + rw [add_comm] + unfold harmonic rwa [Finset.sum_eq_sum_diff_singleton_add (i := c)] } @@ -336,18 +299,15 @@ theorem harmonic_not_int : ∀ n, n ≥ 2 → ¬ (harmonic n).isInt := by { rw [harmonic_singleton (pow2_log_in_finset Hn)] simp only [ge_iff_le, Finset.mem_range, not_lt, Finset.singleton_subset_iff, gt_iff_lt, pow_pos, Nat.cast_pred, Nat.cast_pow, Nat.cast_ofNat, sub_add_cancel] - rw [padicValRat.sum_eq_of_lt (p := 2) _ (one_div_ne_zero <| pow_ne_zero _ two_ne_zero) (harmonic_singleton_ne_zero (Hn)) _]; apply padicValRat_2_pow + rw [padicValRat.add_eq_of_lt (p := 2) _ (one_div_ne_zero <| pow_ne_zero _ two_ne_zero) (harmonic_singleton_ne_zero (Hn)) _]; apply padicValRat_2_pow { rw [harmonic_singleton (pow2_log_in_finset Hn)] at h simpa only [ge_iff_le, gt_iff_lt, pow_pos, Nat.cast_pred, Nat.cast_pow, Nat.cast_ofNat, sub_add_cancel, Finset.mem_range, not_lt, Finset.singleton_subset_iff] using h } { - have := @padicValRat.finset_gen_sum_lt_of_lt 2 _ (2 ^ Nat.log 2 n - 1) (fun n => 1 / ((n + 1) : ℚ)) (Finset.range n \ {2^Nat.log 2 n - 1}) + have := @padicValRat.finset_gen_sum_lt_of_lt (p := 2) (j := 2 ^ Nat.log 2 n - 1) _ (F := fun n => 1 / ((n + 1) : ℚ)) (Finset.range n \ {2^Nat.log 2 n - 1}) simp at * - apply this - { - apply finset_range_sdiff_singleton_nonempty Hn - } + apply this (finset_range_sdiff_singleton_nonempty Hn) { intros i Hi₁ Hi₂ have := padicValRat_ge_neg_Nat_log_lt (n := n) (i+1) From 25695222e2c29ae7c79a0c877fb4dfd2c5de26da Mon Sep 17 00:00:00 2001 From: Koundinya Vajjha Date: Wed, 20 Sep 2023 22:09:49 -0700 Subject: [PATCH 21/74] slight more refactor --- Counterexamples/HarmonicNotInteger.lean | 21 ++++++++++++--------- 1 file changed, 12 insertions(+), 9 deletions(-) diff --git a/Counterexamples/HarmonicNotInteger.lean b/Counterexamples/HarmonicNotInteger.lean index d34ceac9961f5..58ec60383f259 100644 --- a/Counterexamples/HarmonicNotInteger.lean +++ b/Counterexamples/HarmonicNotInteger.lean @@ -126,15 +126,18 @@ end padicValNat def harmonic : ℕ → ℚ := fun n => ∑ i in Finset.range n, 1 / (i + 1) -lemma harmonic_ne_zero : ∀ n, n ≠ 0 → harmonic n > 0 := fun n Hn => Finset.sum_pos (fun i _ => div_pos zero_lt_one (by norm_cast; linarith)) (by (rwa [Finset.nonempty_range_iff])) - -lemma harmonic_singleton {n c : ℕ} (hc : c ∈ Finset.range n): harmonic n =1 / ((c + 1):ℚ) + ∑ x in Finset.range n \ {c}, 1 / ((x : ℚ) + 1) := by { - rw [add_comm] - unfold harmonic - rwa [Finset.sum_eq_sum_diff_singleton_add (i := c)] -} - -lemma finset_range_sdiff_singleton_nonempty {c n : ℕ} (hn : 2 ≤ n) : Finset.Nonempty (Finset.range n \ {c}) := by { +lemma harmonic_ne_zero : ∀ n, n ≠ 0 → harmonic n > 0 := fun n Hn => + Finset.sum_pos (fun _ _ => div_pos zero_lt_one + (by norm_cast; linarith)) + (by (rwa [Finset.nonempty_range_iff])) + +lemma harmonic_singleton {n c : ℕ} (hc : c ∈ Finset.range n): + harmonic n = 1 / (c + 1) + ∑ x in Finset.range n \ {c}, 1 / ((x : ℚ) + 1) := by + unfold harmonic + rwa [add_comm, Finset.sum_eq_sum_diff_singleton_add (i := c)] + +lemma finset_range_sdiff_singleton_nonempty {c n : ℕ} (hn : 2 ≤ n): + Finset.Nonempty (Finset.range n \ {c}) := by { rw [Finset.sdiff_nonempty,Finset.subset_singleton_iff, Finset.range_eq_empty_iff] push_neg refine' ⟨by linarith,fun Hnot => _⟩ From 16095829805029d034a69803cb15d402bb67c1ad Mon Sep 17 00:00:00 2001 From: Koundinya Vajjha Date: Wed, 20 Sep 2023 22:10:26 -0700 Subject: [PATCH 22/74] cleanup --- Counterexamples/HarmonicNotInteger.lean | 30 ------------------------- 1 file changed, 30 deletions(-) diff --git a/Counterexamples/HarmonicNotInteger.lean b/Counterexamples/HarmonicNotInteger.lean index 58ec60383f259..3ac6d432c2a30 100644 --- a/Counterexamples/HarmonicNotInteger.lean +++ b/Counterexamples/HarmonicNotInteger.lean @@ -28,7 +28,6 @@ open BigOperators namespace padicValRat --- Keep lemma min_eq_padicValRat_add {p : ℕ} [hp : Fact (Nat.Prime p)] {q r : ℚ} (hqr : q + r ≠ 0) (hq : q ≠ 0) (hr : r ≠ 0) (hval : padicValRat p q ≠ padicValRat p r) : padicValRat p (q + r) = min (padicValRat p q) (padicValRat p r) := by @@ -48,45 +47,16 @@ namespace padicValRat (by rw [add_neg_cancel_right, add_sub_cancel]) _ = min (padicValRat p (q + r)) (padicValRat p r) := by rw [padicValRat.neg] --- Keep lemma add_eq_of_lt {p : ℕ} [hp : Fact (Nat.Prime p)] {q r : ℚ} (hqr : q + r ≠ 0) (hq : q ≠ 0) (hr : r ≠ 0) (hval : padicValRat p q < padicValRat p r) : padicValRat p (q + r) = padicValRat p q := by rw [min_eq_padicValRat_add hqr hq hr (ne_of_lt hval),min_eq_left (le_of_lt hval)] --- Keep lemma add_lt_of_lt {p : ℕ} [hp : Fact (Nat.Prime p)] {q r₁ r₂ : ℚ} (hqr : r₁ + r₂ ≠ 0) (hval₁ : padicValRat p q < padicValRat p r₁) (hval₂ : padicValRat p q < padicValRat p r₂) : padicValRat p q < padicValRat p (r₁ + r₂) := lt_of_lt_of_le (lt_min hval₁ hval₂) (padicValRat.min_le_padicValRat_add (p := p) (hp := hp) hqr) --- Keep - -- theorem finset_sum_ge_of_ge {p : ℕ} [hp : Fact (Nat.Prime p)] {n j : ℕ} {F : ℕ → ℚ} - -- (hF : ∀ i, i ≤ n → padicValRat p (F j) ≤ padicValRat p (F i)) - -- (hn0 : ∑ i in Finset.range n, F i ≠ 0) : padicValRat p (F j) ≤ - -- padicValRat p (∑ i in Finset.range n, F i) := by - -- induction' n with d hd - -- · exact False.elim (hn0 rfl) - -- · rw [Finset.sum_range_succ] at hn0 ⊢ - -- by_cases h : ∑ x : ℕ in Finset.range d, F x = 0 - -- · rw [h, zero_add] - -- exact hF d (by linarith) - -- · refine' le_trans _ (padicValRat.min_le_padicValRat_add (p := p) hn0) - -- · exact le_min (hd (fun i hi => hF _ (by linarith)) h) (hF _ (by linarith)) - - -- theorem finset_sum_lt_of_lt {p : ℕ} [hp : Fact (Nat.Prime p)] {n j : ℕ} {F : ℕ → ℚ} - -- (hF : ∀ i, i ≤ n → padicValRat p (F j) < padicValRat p (F i)) - -- (hn0 : ∑ i in Finset.range n, F i ≠ 0) : - -- padicValRat p (F j) < padicValRat p (∑ i in Finset.range n, F i) := by - -- induction' n with d hd - -- · exact False.elim (hn0 rfl) - -- · rw [Finset.sum_range_succ] at hn0 ⊢ - -- by_cases h : ∑ x : ℕ in Finset.range d, F x = 0 - -- · rw [h, zero_add] - -- exact hF d (by linarith) - -- · refine' le_trans _ (padicValRat.min_le_padicValRat_add (p := p) hn0) - -- · exact le_min (hd (fun i hi => hF _ (by linarith)) h) (hF _ (by linarith)) - theorem finset_gen_sum_lt_of_lt {p j : ℕ} [hp : Fact (Nat.Prime p)] {F : ℕ → ℚ} {S : Finset ℕ} (hS : S.Nonempty) (hF : ∀ i, i ∈ S → padicValRat p (F j) < padicValRat p (F i)) From 45c0d30fb0305ab7e490a291701d6ad6ce422bc7 Mon Sep 17 00:00:00 2001 From: Koundinya Vajjha Date: Wed, 20 Sep 2023 22:12:49 -0700 Subject: [PATCH 23/74] remove unneeded imports --- Counterexamples/HarmonicNotInteger.lean | 3 --- 1 file changed, 3 deletions(-) diff --git a/Counterexamples/HarmonicNotInteger.lean b/Counterexamples/HarmonicNotInteger.lean index 3ac6d432c2a30..db31e33c740ec 100644 --- a/Counterexamples/HarmonicNotInteger.lean +++ b/Counterexamples/HarmonicNotInteger.lean @@ -4,10 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Koundinya Vajjha -/ -import Mathlib.NumberTheory.Padics.PadicNorm -import Mathlib.NumberTheory.Padics.PadicNumbers import Mathlib.NumberTheory.Padics.PadicIntegers -import Mathlib.NumberTheory.Padics.PadicVal import Mathlib.Data.Int.Log /-! From 1311ab68c415de5cbb2230e47f68d0320c7323e1 Mon Sep 17 00:00:00 2001 From: Koundinya Vajjha Date: Wed, 20 Sep 2023 22:22:00 -0700 Subject: [PATCH 24/74] document --- Counterexamples/HarmonicNotInteger.lean | 4 ++++ 1 file changed, 4 insertions(+) diff --git a/Counterexamples/HarmonicNotInteger.lean b/Counterexamples/HarmonicNotInteger.lean index db31e33c740ec..8054093597006 100644 --- a/Counterexamples/HarmonicNotInteger.lean +++ b/Counterexamples/HarmonicNotInteger.lean @@ -54,6 +54,10 @@ namespace padicValRat padicValRat p q < padicValRat p (r₁ + r₂) := lt_of_lt_of_le (lt_min hval₁ hval₂) (padicValRat.min_le_padicValRat_add (p := p) (hp := hp) hqr) + /-- + If the p-adic valuation of a finite set of positive rationals is greater than a given rational + number, then the p-adic valuation of their sum is also greater than the same rational number. + -/ theorem finset_gen_sum_lt_of_lt {p j : ℕ} [hp : Fact (Nat.Prime p)] {F : ℕ → ℚ} {S : Finset ℕ} (hS : S.Nonempty) (hF : ∀ i, i ∈ S → padicValRat p (F j) < padicValRat p (F i)) From a7fece2b43ae71d828419e25882e8bc857516589 Mon Sep 17 00:00:00 2001 From: Koundinya Vajjha Date: Wed, 20 Sep 2023 22:26:33 -0700 Subject: [PATCH 25/74] add file to counterexamples --- Counterexamples.lean | 1 + 1 file changed, 1 insertion(+) diff --git a/Counterexamples.lean b/Counterexamples.lean index 18cd76a5ceb5f..b3b4e15cc9478 100644 --- a/Counterexamples.lean +++ b/Counterexamples.lean @@ -13,3 +13,4 @@ import Counterexamples.QuadraticForm import Counterexamples.SeminormLatticeNotDistrib import Counterexamples.SorgenfreyLine import Counterexamples.ZeroDivisorsInAddMonoidAlgebras +import Counterexamples.HarmonicNotInteger From 15b79f9d12f9fceffde0be591619bae2aa33d53f Mon Sep 17 00:00:00 2001 From: Koundinya Vajjha Date: Thu, 21 Sep 2023 16:35:09 -0700 Subject: [PATCH 26/74] more changes --- Counterexamples/HarmonicNotInteger.lean | 56 +++++++++---------------- 1 file changed, 20 insertions(+), 36 deletions(-) diff --git a/Counterexamples/HarmonicNotInteger.lean b/Counterexamples/HarmonicNotInteger.lean index 8054093597006..1ba9a07401d6a 100644 --- a/Counterexamples/HarmonicNotInteger.lean +++ b/Counterexamples/HarmonicNotInteger.lean @@ -25,6 +25,7 @@ open BigOperators namespace padicValRat + /-- Ultrametric property of a p-adic valuation. -/ lemma min_eq_padicValRat_add {p : ℕ} [hp : Fact (Nat.Prime p)] {q r : ℚ} (hqr : q + r ≠ 0) (hq : q ≠ 0) (hr : r ≠ 0) (hval : padicValRat p q ≠ padicValRat p r) : padicValRat p (q + r) = min (padicValRat p q) (padicValRat p r) := by @@ -108,60 +109,43 @@ lemma harmonic_singleton {n c : ℕ} (hc : c ∈ Finset.range n): rwa [add_comm, Finset.sum_eq_sum_diff_singleton_add (i := c)] lemma finset_range_sdiff_singleton_nonempty {c n : ℕ} (hn : 2 ≤ n): - Finset.Nonempty (Finset.range n \ {c}) := by { - rw [Finset.sdiff_nonempty,Finset.subset_singleton_iff, Finset.range_eq_empty_iff] - push_neg + Finset.Nonempty (Finset.range n \ {c}) := by + rw [Finset.sdiff_nonempty,Finset.subset_singleton_iff, Finset.range_eq_empty_iff,not_or] refine' ⟨by linarith,fun Hnot => _⟩ suffices : n = 1; linarith rw [← Finset.card_range n, ← Finset.card_singleton c, Hnot] -} -lemma harmonic_singleton_ne_zero {c n : ℕ} (hn : 2 ≤ n) : (∑ x in Finset.range n \ {c}, 1 / ((x + 1 : ℚ)) ≠ 0) := by { - refine' ne_of_gt _ - refine' Finset.sum_pos (fun i _ => div_pos zero_lt_one _) _ - norm_cast - simp only [add_pos_iff, or_true] - apply finset_range_sdiff_singleton_nonempty hn -} -theorem not_int_of_not_padic_int (a : ℚ) : - padicNorm 2 a > 1 → ¬ a.isInt := by { - intro H +lemma harmonic_singleton_ne_zero {c n : ℕ} (hn : 2 ≤ n) : + (∑ x in Finset.range n \ {c}, 1 / ((x + 1 : ℚ)) ≠ 0) := by + refine' ne_of_gt <| Finset.sum_pos (fun i _ => div_pos zero_lt_one _) + (finset_range_sdiff_singleton_nonempty hn) + norm_cast; simp only [add_pos_iff, or_true] + +theorem not_int_of_not_padic_int {a : ℚ} (H : 1 < padicNorm 2 a): + ¬ a.isInt := by suffices : a.den ≠ 1; simpa [Rat.isInt] by_cases (a = 0) - { - rw [h] at H h - dsimp [padicNorm] at * - simp only at H - } - { - dsimp [padicNorm] at * + · subst h; try contradiction + · unfold padicNorm at H split_ifs at H; try contradiction - suffices : padicValRat 2 a < 0 - unfold padicValRat at this + suffices : padicValRat 2 a < 0; unfold padicValRat at this intro Hden - rw [Hden] at this - simp only [padicValNat.one, sub_zero] at this - norm_cast at H - have Hz : 0 = -0 := by {norm_num} - rw [Hz] - apply lt_neg_of_lt_neg - have hx : (1 : ℚ) < 2 := by {norm_num} + rw [Hden] at this; simp only [padicValNat.one, sub_zero] at this + norm_cast + refine' neg_of_neg_pos _ + have hx : (1 : ℚ) < 2 := by norm_num rw [← zpow_lt_iff_lt hx] simpa only [zpow_zero] - } -} -@[simp] + lemma padicValRat_2_pow (r : ℕ) : padicValRat 2 (1 / 2^r) = -r := by { rw [one_div,padicValRat.inv,neg_inj,padicValRat.pow (by simp)] suffices : padicValRat 2 2 = 1 simp only [this, mul_one] - rw [←padicValRat.self (p := 2)]; simp only [Nat.cast_ofNat] - norm_num + rw [←padicValRat.self (p := 2) one_lt_two]; simp only [Nat.cast_ofNat] } - lemma nat_log_not_dvd {n : ℕ} : ∀ i, 0 < i ∧ i ≠ 2^(Nat.log 2 n) ∧ i ≤ n → ¬ 2^(Nat.log 2 n) ∣ i := by { rintro i ⟨Hi₁,Hi₂,Hi₃⟩ simp only [Nat.instDvdNat] From 2a27077e5771a8ba57255006a3ab09d1260638a0 Mon Sep 17 00:00:00 2001 From: Koundinya Vajjha Date: Thu, 21 Sep 2023 21:35:32 -0700 Subject: [PATCH 27/74] simplify all proofs and add some documentation --- Counterexamples/HarmonicNotInteger.lean | 180 +++++++++--------------- 1 file changed, 68 insertions(+), 112 deletions(-) diff --git a/Counterexamples/HarmonicNotInteger.lean b/Counterexamples/HarmonicNotInteger.lean index 1ba9a07401d6a..d760dfe66a1cd 100644 --- a/Counterexamples/HarmonicNotInteger.lean +++ b/Counterexamples/HarmonicNotInteger.lean @@ -96,6 +96,7 @@ namespace padicValNat end padicValNat +/-- The nth-harmonic number defined as a finset sum of reciprocals. -/ def harmonic : ℕ → ℚ := fun n => ∑ i in Finset.range n, 1 / (i + 1) lemma harmonic_ne_zero : ∀ n, n ≠ 0 → harmonic n > 0 := fun n Hn => @@ -122,6 +123,7 @@ lemma harmonic_singleton_ne_zero {c n : ℕ} (hn : 2 ≤ n) : (finset_range_sdiff_singleton_nonempty hn) norm_cast; simp only [add_pos_iff, or_true] +/-- If a rational is not a 2-adic integer, it is not an integer. -/ theorem not_int_of_not_padic_int {a : ℚ} (H : 1 < padicNorm 2 a): ¬ a.isInt := by suffices : a.den ≠ 1; simpa [Rat.isInt] @@ -139,146 +141,100 @@ theorem not_int_of_not_padic_int {a : ℚ} (H : 1 < padicNorm 2 a): simpa only [zpow_zero] -lemma padicValRat_2_pow (r : ℕ) : padicValRat 2 (1 / 2^r) = -r := by { +lemma padicValRat_2_pow (r : ℕ) : padicValRat 2 (1 / 2^r) = -r := by rw [one_div,padicValRat.inv,neg_inj,padicValRat.pow (by simp)] suffices : padicValRat 2 2 = 1 simp only [this, mul_one] rw [←padicValRat.self (p := 2) one_lt_two]; simp only [Nat.cast_ofNat] -} -lemma nat_log_not_dvd {n : ℕ} : ∀ i, 0 < i ∧ i ≠ 2^(Nat.log 2 n) ∧ i ≤ n → ¬ 2^(Nat.log 2 n) ∣ i := by { + +lemma nat_log_not_dvd {n : ℕ} : + ∀ i, 0 < i ∧ i ≠ 2^(Nat.log 2 n) ∧ i ≤ n → ¬ 2^(Nat.log 2 n) ∣ i := by rintro i ⟨Hi₁,Hi₂,Hi₃⟩ simp only [Nat.instDvdNat] - push_neg - intros c Hc + push_neg; intros c Hc have Hle : n < 2 ^ (Nat.log 2 n + 1) := Nat.lt_pow_succ_log_self (by simp) n have Hpow : 2 ^ Nat.log 2 n * 2 = 2 ^(Nat.log 2 n + 1) := by exact rfl obtain ⟨k,Hk⟩ := Nat.even_or_odd' c by_cases (k = 0) - { - rw [h, mul_zero] at Hk + · subst h + rw [mul_zero, zero_add] at Hk rcases Hk with rfl | rfl - linarith - simp only [zero_add, mul_one] at Hc - exact Hi₂ Hc - } + · linarith + · exact Hi₂ (Hc ▸ mul_one (2^Nat.log 2 n)) rcases Hk with rfl | rfl - { - rw [← mul_assoc, Hpow] at Hc - { - suffices Hnk : n < i by linarith - calc n ≤ n * k := Nat.le_mul_of_pos_right (Nat.pos_of_ne_zero h) - _ < 2^(Nat.log 2 n + 1) * k := Nat.mul_lt_mul Hle (le_refl _) (Nat.pos_of_ne_zero h) - _ = i := Hc.symm - } - } - { - rw [mul_add, ← mul_assoc, Hpow, mul_one] at Hc + · rw [← mul_assoc, Hpow] at Hc + suffices Hnk : n < i by linarith + calc n ≤ n * k := Nat.le_mul_of_pos_right (Nat.pos_of_ne_zero h) + _ < 2^(Nat.log 2 n + 1) * k := Nat.mul_lt_mul Hle (le_refl _) (Nat.pos_of_ne_zero h) + _ = i := Hc.symm + · rw [mul_add, ← mul_assoc, Hpow, mul_one] at Hc suffices Hnk : n < i by linarith calc n < 2 ^ (Nat.log 2 n + 1) := Hle _ ≤ 2 ^ (Nat.log 2 n + 1)*k := Nat.le_mul_of_pos_right (Nat.pos_of_ne_zero h) _ < 2 ^ (Nat.log 2 n + 1)*k + 2^ Nat.log 2 n := by linarith _ = i := Hc.symm - } - } - lemma padicValRat_ge_neg_Nat_log_ne {n : ℕ}: -∀ q, 0 < q ∧ q ≤ n → q ≠ 2^Nat.log 2 n → padicValRat 2 (1 / q) ≠ -Nat.log 2 n := by { + ∀ q, 0 < q ∧ q ≤ n → q ≠ 2^Nat.log 2 n → padicValRat 2 (1 / q) ≠ -Nat.log 2 n := by rintro q ⟨Hq₁,Hq₂⟩ Hq₃ - rw [one_div,padicValRat.inv,padicValRat.of_nat] - simp only [ne_eq, neg_inj, Nat.cast_inj] - intro Hnot - have := pow_padicValNat_dvd (p := 2) (n := q) - rw [Hnot] at this - exact nat_log_not_dvd (n := n) q ⟨Hq₁,Hq₃,Hq₂⟩ this -} + rw [one_div,padicValRat.inv,padicValRat.of_nat,ne_eq,neg_inj,Nat.cast_inj] + exact (fun Hnot => nat_log_not_dvd (n := n) q ⟨Hq₁,Hq₃,Hq₂⟩ + (Hnot ▸ pow_padicValNat_dvd (p := 2) (n := q))) lemma padicValRat_ge_neg_Nat_log_ge {p n : ℕ}[hp : Fact (Nat.Prime p)]: - ∀ q, q ≤ n → -Nat.log p n ≤ padicValRat p (1 / q) := by { + ∀ q, q ≤ n → -Nat.log p n ≤ padicValRat p (1 / q) := by intros q Hq rw [one_div,padicValRat.inv,padicValRat.of_nat, neg_le_neg_iff, Nat.cast_le] - apply padicValNat.le_nat_log_gen Hq -} + exact padicValNat.le_nat_log_gen Hq lemma padicValRat_ge_neg_Nat_log_lt (n : ℕ): -∀ q, 0 < q ∧ q ≤ n → q ≠ 2^Nat.log 2 n → padicValRat 2 (1 / 2^Nat.log 2 n) < padicValRat 2 (1 / q) := by { + ∀ q, 0 < q ∧ q ≤ n → q ≠ 2^Nat.log 2 n → + padicValRat 2 (1 / 2^Nat.log 2 n) < padicValRat 2 (1 / q) := by rintro q ⟨Hq₁,Hq₂⟩ Hq₃ - have H₁ := padicValRat_ge_neg_Nat_log_ne q ⟨Hq₁,Hq₂⟩ Hq₃ - have H₂ := padicValRat_ge_neg_Nat_log_ge (p := 2) q Hq₂ rw [padicValRat_2_pow] - exact lt_of_le_of_ne H₂ H₁.symm -} - -lemma pow2_log_in_finset {n : ℕ} (hn : 2 ≤ n) : 2^(Nat.log 2 n) - 1 ∈ Finset.range n := by { - simp only [ge_iff_le, Finset.mem_range] - have H := Nat.pow_log_le_add_one 2 n - rw [le_iff_lt_or_eq] at H - rcases H with H₁ | H₂ - { - simp only [ge_iff_le, gt_iff_lt] - refine' Nat.sub_lt_right_of_lt_add _ H₁ - calc 1 ≤ 2 := by norm_num - _ = 2^1 := by norm_num - _ <= 2 ^ Nat.log 2 n := by { - refine' Nat.pow_le_pow_of_le_right (by norm_num) _ - rw [←Nat.pow_le_iff_le_log (by norm_num) _] - simpa - linarith - } - } - have Habs : n + 1 ≤ n := by { - rw [← H₂] - apply Nat.pow_log_le_self; linarith - } - linarith -} - - -theorem harmonic_not_int : ∀ n, n ≥ 2 → ¬ (harmonic n).isInt := by { + exact (lt_of_le_of_ne + (padicValRat_ge_neg_Nat_log_ge (p := 2) q Hq₂) + (padicValRat_ge_neg_Nat_log_ne q ⟨Hq₁,Hq₂⟩ Hq₃).symm) + +lemma pow2_log_in_finset {n : ℕ} (hn : 2 ≤ n): + 2^(Nat.log 2 n) - 1 ∈ Finset.range n := Finset.mem_range.mpr <| lt_of_lt_of_le + (Nat.sub_lt (pow_pos zero_lt_two _) zero_lt_one) (Nat.pow_log_le_self 2 (by linarith)) + +/-- The 2-adic valuation of the n-th harmonic number is the negative of the logarithm + of n. -/ +theorem harmonic_padicValRat_2 {n : ℕ} (Hn : 2 ≤ n): + padicValRat 2 (harmonic n) = -Int.log 2 (n : ℚ) := by + rw [Int.log_natCast,harmonic_singleton (pow2_log_in_finset Hn)] + simp only [gt_iff_lt, pow_pos, Nat.cast_pred, Nat.cast_pow, Nat.cast_ofNat, sub_add_cancel] + rw [padicValRat.add_eq_of_lt (p := 2) _ (one_div_ne_zero <| pow_ne_zero _ two_ne_zero) + (harmonic_singleton_ne_zero (Hn)) _]; apply padicValRat_2_pow + · have := harmonic_ne_zero n + rw [harmonic_singleton (pow2_log_in_finset Hn)] at this + refine' ne_of_gt _ + simp only [ne_eq, ge_iff_le, gt_iff_lt, pow_pos, Nat.cast_pred, Nat.cast_pow, Nat.cast_ofNat, + sub_add_cancel, Finset.mem_range, not_lt, Finset.singleton_subset_iff] at this + exact (this (by linarith)) + · have := padicValRat.finset_gen_sum_lt_of_lt (p := 2) (j := 2 ^ Nat.log 2 n - 1) + (finset_range_sdiff_singleton_nonempty Hn) (F := fun n => 1 / ((n + 1) : ℚ)) + (S := Finset.range n \ {2^Nat.log 2 n - 1}) + simp [-one_div] at * + refine' this (fun i Hi₁ Hi₂ => _) (fun i => div_pos zero_lt_one (by norm_cast;linarith)) + · have := padicValRat_ge_neg_Nat_log_lt (n := n) (i+1) + simp only [Nat.cast_add, Nat.cast_one] at this + refine' this ⟨by linarith, by linarith⟩ (fun Hnot => _) + rw [← Hnot] at Hi₂ + refine' Hi₂ _ + simp only [add_tsub_cancel_right] + +theorem harmonic_not_int : ∀ n, n ≥ 2 → ¬ (harmonic n).isInt := by intro n Hn - apply not_int_of_not_padic_int - unfold padicNorm + refine' not_int_of_not_padic_int _ + dsimp [padicNorm] split_ifs with h - { - have := harmonic_ne_zero n - rw [h] at this - apply this - linarith - } - { - apply one_lt_zpow; try norm_num - apply lt_neg.mp - rw [neg_zero] - suffices Hlog : padicValRat 2 (harmonic n) = -Int.log 2 (n : ℚ) - rw [Hlog] - simp only [Int.log_natCast, Left.neg_neg_iff, Nat.cast_pos, Nat.log_pos_iff, and_true, Hn] - simp only [Int.log_natCast] - rw [harmonic_singleton (pow2_log_in_finset Hn)] - simp only [ge_iff_le, Finset.mem_range, not_lt, Finset.singleton_subset_iff, gt_iff_lt, pow_pos, - Nat.cast_pred, Nat.cast_pow, Nat.cast_ofNat, sub_add_cancel] - rw [padicValRat.add_eq_of_lt (p := 2) _ (one_div_ne_zero <| pow_ne_zero _ two_ne_zero) (harmonic_singleton_ne_zero (Hn)) _]; apply padicValRat_2_pow - { - rw [harmonic_singleton (pow2_log_in_finset Hn)] at h - simpa only [ge_iff_le, gt_iff_lt, pow_pos, Nat.cast_pred, Nat.cast_pow, Nat.cast_ofNat, sub_add_cancel, Finset.mem_range, not_lt, Finset.singleton_subset_iff] using h - } - { - have := @padicValRat.finset_gen_sum_lt_of_lt (p := 2) (j := 2 ^ Nat.log 2 n - 1) _ (F := fun n => 1 / ((n + 1) : ℚ)) (Finset.range n \ {2^Nat.log 2 n - 1}) - simp at * - apply this (finset_range_sdiff_singleton_nonempty Hn) - { - intros i Hi₁ Hi₂ - have := padicValRat_ge_neg_Nat_log_lt (n := n) (i+1) - simp at * - refine' this (by linarith) _ - intro Hnot - rw [←Hnot] at Hi₂ - simp only [ge_iff_le, add_le_iff_nonpos_left, nonpos_iff_eq_zero, add_tsub_cancel_right, not_true] at Hi₂ - } - { - intros i - norm_cast; linarith - } - } - } -} + · exfalso + exact ((ne_of_gt <| harmonic_ne_zero n (by linarith)) h) + · refine' one_lt_zpow one_lt_two _ (lt_neg.mp _) + rw [neg_zero, harmonic_padicValRat_2 Hn,Int.log_natCast, Left.neg_neg_iff, + Nat.cast_pos, Nat.log_pos_iff] + exact ⟨Hn,one_lt_two⟩ From 90cf3710e056deae76a64492b54d6b22246faec3 Mon Sep 17 00:00:00 2001 From: Koundinya Vajjha Date: Thu, 21 Sep 2023 22:16:10 -0700 Subject: [PATCH 28/74] factor out lemmas to go into padicval --- Counterexamples/HarmonicNotInteger.lean | 88 ++--------------------- Mathlib/NumberTheory/Padics/PadicVal.lean | 68 ++++++++++++++++++ 2 files changed, 75 insertions(+), 81 deletions(-) diff --git a/Counterexamples/HarmonicNotInteger.lean b/Counterexamples/HarmonicNotInteger.lean index d760dfe66a1cd..66e3f749621b8 100644 --- a/Counterexamples/HarmonicNotInteger.lean +++ b/Counterexamples/HarmonicNotInteger.lean @@ -18,88 +18,13 @@ https://kconrad.math.uconn.edu/blurbs/gradnumthy/padicharmonicsum.pdf namespace Counterexample -open Rat -open BigOperators - +open BigOperators -namespace padicValRat - - /-- Ultrametric property of a p-adic valuation. -/ - lemma min_eq_padicValRat_add {p : ℕ} [hp : Fact (Nat.Prime p)] {q r : ℚ} - (hqr : q + r ≠ 0) (hq : q ≠ 0) (hr : r ≠ 0) (hval : padicValRat p q ≠ padicValRat p r) : - padicValRat p (q + r) = min (padicValRat p q) (padicValRat p r) := by - have Hmin := padicValRat.min_le_padicValRat_add (p := p) (hp := hp) hqr - wlog h : padicValRat p q < padicValRat p r generalizing q r with Hgen - · push_neg at h; rw [add_comm, min_comm] - exact (Hgen (by rwa [add_comm]) hr hq hval.symm - (by rwa [min_comm,add_comm]) - (Ne.lt_of_le (Ne.symm hval) h)) - · rw [min_eq_left (le_of_lt h)] at Hmin ⊢ - suffices Hreq : padicValRat p q ≥ padicValRat p (q + r) by linarith - suffices Haux : padicValRat p q ≥ min (padicValRat p (q + r)) (padicValRat p r) by - rw [min_def] at Haux; split_ifs at Haux with Hspl; try assumption; linarith - calc padicValRat p q = padicValRat p ((q + r) - r) := by congr; simp - _ ≥ min (padicValRat p (q + r)) (padicValRat p (-r)) := - ge_iff_le.mp <| le_trans (padicValRat.min_le_padicValRat_add (q := q+r) (r := -r) (by simpa)) - (by rw [add_neg_cancel_right, add_sub_cancel]) - _ = min (padicValRat p (q + r)) (padicValRat p r) := by rw [padicValRat.neg] - - lemma add_eq_of_lt {p : ℕ} [hp : Fact (Nat.Prime p)] {q r : ℚ} (hqr : q + r ≠ 0) - (hq : q ≠ 0) (hr : r ≠ 0) (hval : padicValRat p q < padicValRat p r) : - padicValRat p (q + r) = padicValRat p q := - by rw [min_eq_padicValRat_add hqr hq hr (ne_of_lt hval),min_eq_left (le_of_lt hval)] - - lemma add_lt_of_lt {p : ℕ} [hp : Fact (Nat.Prime p)] {q r₁ r₂ : ℚ} (hqr : r₁ + r₂ ≠ 0) - (hval₁ : padicValRat p q < padicValRat p r₁) (hval₂ : padicValRat p q < padicValRat p r₂) : - padicValRat p q < padicValRat p (r₁ + r₂) := - lt_of_lt_of_le (lt_min hval₁ hval₂) (padicValRat.min_le_padicValRat_add (p := p) (hp := hp) hqr) - - /-- - If the p-adic valuation of a finite set of positive rationals is greater than a given rational - number, then the p-adic valuation of their sum is also greater than the same rational number. - -/ - theorem finset_gen_sum_lt_of_lt {p j : ℕ} [hp : Fact (Nat.Prime p)] - {F : ℕ → ℚ} {S : Finset ℕ} (hS : S.Nonempty) - (hF : ∀ i, i ∈ S → padicValRat p (F j) < padicValRat p (F i)) - (hn1 : ∀ i : ℕ, 0 < F i): - padicValRat p (F j) < padicValRat p (∑ i in S, F i) := by - induction' hS using Finset.Nonempty.cons_induction with k s S' Hnot Hne Hind - · rw [Finset.sum_singleton] - exact hF k (by simp) - · rw [Finset.cons_eq_insert, Finset.sum_insert Hnot] - exact add_lt_of_lt - (ne_of_gt (add_pos (hn1 s) (Finset.sum_pos (fun i _ => hn1 i) Hne))) - (hF _ (by simp [Finset.mem_insert, true_or])) - (Hind (fun i hi => hF _ (by rw [Finset.cons_eq_insert,Finset.mem_insert]; exact Or.inr hi))) - -end padicValRat - -namespace padicValNat - - lemma le_nat_log {p : ℕ} [hp : Fact (Nat.Prime p)] (n : ℕ): - padicValNat p n ≤ Nat.log p n := by - by_cases (n = 0) - · simp only [h, padicValNat.zero, Nat.log_zero_right, le_refl] - · refine' Nat.le_log_of_pow_le (Nat.Prime.one_lt hp.elim) _ - by_contra Hnot; push_neg at Hnot - exact h (Nat.eq_zero_of_dvd_of_lt (pow_padicValNat_dvd (p := p) (n := n)) Hnot) - - lemma le_nat_log_gen {p n₁ n₂ : ℕ} [Fact (Nat.Prime p)] (hn : n₁ ≤ n₂): - padicValNat p n₁ ≤ Nat.log p n₂ := le_trans (le_nat_log n₁) (Nat.log_mono_right hn) - - lemma nat_log_eq_padicvalnat_iff {p : ℕ} [hp : Fact (Nat.Prime p)] (n : ℕ)(hn : 0 < n): - Nat.log p n = padicValNat p n ↔ n < p^(padicValNat p n + 1) := by - · rw [Nat.log_eq_iff (Or.inr ⟨(Nat.Prime.one_lt' p).out, by linarith⟩)] - · rw [and_iff_right_iff_imp] - exact (fun _ => Nat.le_of_dvd hn pow_padicValNat_dvd) - -end padicValNat - -/-- The nth-harmonic number defined as a finset sum of reciprocals. -/ +/-- The nth-harmonic number defined as a finset sum of consecutive reciprocals. -/ def harmonic : ℕ → ℚ := fun n => ∑ i in Finset.range n, 1 / (i + 1) -lemma harmonic_ne_zero : ∀ n, n ≠ 0 → harmonic n > 0 := fun n Hn => +lemma harmonic_pos : ∀ n, n ≠ 0 → 0 < harmonic n := fun n Hn => Finset.sum_pos (fun _ _ => div_pos zero_lt_one (by norm_cast; linarith)) (by (rwa [Finset.nonempty_range_iff])) @@ -147,7 +72,7 @@ lemma padicValRat_2_pow (r : ℕ) : padicValRat 2 (1 / 2^r) = -r := by simp only [this, mul_one] rw [←padicValRat.self (p := 2) one_lt_two]; simp only [Nat.cast_ofNat] - +/-- For `i` less than `n`, `2^Nat.log 2 n` does not divide `i` unless `i` is equal to it. -/ lemma nat_log_not_dvd {n : ℕ} : ∀ i, 0 < i ∧ i ≠ 2^(Nat.log 2 n) ∧ i ≤ n → ¬ 2^(Nat.log 2 n) ∣ i := by rintro i ⟨Hi₁,Hi₂,Hi₃⟩ @@ -209,7 +134,7 @@ theorem harmonic_padicValRat_2 {n : ℕ} (Hn : 2 ≤ n): simp only [gt_iff_lt, pow_pos, Nat.cast_pred, Nat.cast_pow, Nat.cast_ofNat, sub_add_cancel] rw [padicValRat.add_eq_of_lt (p := 2) _ (one_div_ne_zero <| pow_ne_zero _ two_ne_zero) (harmonic_singleton_ne_zero (Hn)) _]; apply padicValRat_2_pow - · have := harmonic_ne_zero n + · have := harmonic_pos n rw [harmonic_singleton (pow2_log_in_finset Hn)] at this refine' ne_of_gt _ simp only [ne_eq, ge_iff_le, gt_iff_lt, pow_pos, Nat.cast_pred, Nat.cast_pow, Nat.cast_ofNat, @@ -227,13 +152,14 @@ theorem harmonic_padicValRat_2 {n : ℕ} (Hn : 2 ≤ n): refine' Hi₂ _ simp only [add_tsub_cancel_right] +/-- The n-th harmonic number is not an integer for n ≥ 2. -/ theorem harmonic_not_int : ∀ n, n ≥ 2 → ¬ (harmonic n).isInt := by intro n Hn refine' not_int_of_not_padic_int _ dsimp [padicNorm] split_ifs with h · exfalso - exact ((ne_of_gt <| harmonic_ne_zero n (by linarith)) h) + exact (ne_of_gt <| harmonic_pos n (by linarith)) h · refine' one_lt_zpow one_lt_two _ (lt_neg.mp _) rw [neg_zero, harmonic_padicValRat_2 Hn,Int.log_natCast, Left.neg_neg_iff, Nat.cast_pos, Nat.log_pos_iff] diff --git a/Mathlib/NumberTheory/Padics/PadicVal.lean b/Mathlib/NumberTheory/Padics/PadicVal.lean index e39110d8d95c4..daff8451efe90 100644 --- a/Mathlib/NumberTheory/Padics/PadicVal.lean +++ b/Mathlib/NumberTheory/Padics/PadicVal.lean @@ -431,6 +431,36 @@ theorem min_le_padicValRat_add {q r : ℚ} (hqr : q + r ≠ 0) : (fun h => by rw [min_eq_right h, add_comm]; exact le_padicValRat_add_of_le (by rwa [add_comm]) h) #align padic_val_rat.min_le_padic_val_rat_add padicValRat.min_le_padicValRat_add +/-- Ultrametric property of a p-adic valuation. -/ +lemma min_eq_padicValRat_add {q r : ℚ} (hqr : q + r ≠ 0) (hq : q ≠ 0) (hr : r ≠ 0) +(hval : padicValRat p q ≠ padicValRat p r) : +padicValRat p (q + r) = min (padicValRat p q) (padicValRat p r) := by + have Hmin := padicValRat.min_le_padicValRat_add (p := p) (hp := hp) hqr + wlog h : padicValRat p q < padicValRat p r generalizing q r with Hgen + · push_neg at h; rw [add_comm, min_comm] + exact (Hgen (by rwa [add_comm]) hr hq hval.symm + (by rwa [min_comm,add_comm]) + (Ne.lt_of_le (Ne.symm hval) h)) + · rw [min_eq_left (le_of_lt h)] at Hmin ⊢ + suffices Hreq : padicValRat p q ≥ padicValRat p (q + r) by linarith + suffices Haux : padicValRat p q ≥ min (padicValRat p (q + r)) (padicValRat p r) by + rw [min_def] at Haux; split_ifs at Haux with Hspl; try assumption; linarith + calc padicValRat p q = padicValRat p ((q + r) - r) := by congr; simp + _ ≥ min (padicValRat p (q + r)) (padicValRat p (-r)) := + ge_iff_le.mp <| le_trans (padicValRat.min_le_padicValRat_add (q := q+r) (r := -r) (by simpa)) + (by rw [add_neg_cancel_right, add_sub_cancel]) + _ = min (padicValRat p (q + r)) (padicValRat p r) := by rw [padicValRat.neg] + +lemma add_eq_of_lt {q r : ℚ} (hqr : q + r ≠ 0) +(hq : q ≠ 0) (hr : r ≠ 0) (hval : padicValRat p q < padicValRat p r) : +padicValRat p (q + r) = padicValRat p q := +by rw [min_eq_padicValRat_add hqr hq hr (ne_of_lt hval),min_eq_left (le_of_lt hval)] + +lemma add_lt_of_lt {q r₁ r₂ : ℚ} (hqr : r₁ + r₂ ≠ 0) +(hval₁ : padicValRat p q < padicValRat p r₁) (hval₂ : padicValRat p q < padicValRat p r₂) : + padicValRat p q < padicValRat p (r₁ + r₂) := + lt_of_lt_of_le (lt_min hval₁ hval₂) (padicValRat.min_le_padicValRat_add (p := p) (hp := hp) hqr) + open BigOperators /-- A finite sum of rationals with positive `p`-adic valuation has positive `p`-adic valuation @@ -448,6 +478,24 @@ theorem sum_pos_of_pos {n : ℕ} {F : ℕ → ℚ} (hF : ∀ i, i < n → 0 < pa exact hF _ (lt_trans hi (lt_add_one _)) #align padic_val_rat.sum_pos_of_pos padicValRat.sum_pos_of_pos +/-- + If the p-adic valuation of a finite set of positive rationals is greater than a given rational + number, then the p-adic valuation of their sum is also greater than the same rational number. +-/ +theorem finset_gen_sum_lt_of_lt {p j : ℕ} [hp : Fact (Nat.Prime p)] +{F : ℕ → ℚ} {S : Finset ℕ} (hS : S.Nonempty) +(hF : ∀ i, i ∈ S → padicValRat p (F j) < padicValRat p (F i)) +(hn1 : ∀ i : ℕ, 0 < F i): +padicValRat p (F j) < padicValRat p (∑ i in S, F i) := by + induction' hS using Finset.Nonempty.cons_induction with k s S' Hnot Hne Hind + · rw [Finset.sum_singleton] + exact hF k (by simp) + · rw [Finset.cons_eq_insert, Finset.sum_insert Hnot] + exact padicValRat.add_lt_of_lt + (ne_of_gt (add_pos (hn1 s) (Finset.sum_pos (fun i _ => hn1 i) Hne))) + (hF _ (by simp [Finset.mem_insert, true_or])) + (Hind (fun i hi => hF _ (by rw [Finset.cons_eq_insert,Finset.mem_insert]; exact Or.inr hi))) + end padicValRat namespace padicValNat @@ -535,6 +583,26 @@ theorem padicValNat_primes {q : ℕ} [hp : Fact p.Prime] [hq : Fact q.Prime] (ne (not_congr (Iff.symm (prime_dvd_prime_iff_eq hp.1 hq.1))).mp neq #align padic_val_nat_primes padicValNat_primes + /-- The p-adic valuation of `n` is less than or equal to its logarithm w.r.t `p`.-/ + lemma le_nat_log {p : ℕ} [hp : Fact (Nat.Prime p)] (n : ℕ): + padicValNat p n ≤ Nat.log p n := by + by_cases (n = 0) + · simp only [h, padicValNat.zero, Nat.log_zero_right, le_refl] + · refine' Nat.le_log_of_pow_le (Nat.Prime.one_lt hp.elim) _ + by_contra Hnot; push_neg at Hnot + exact h (Nat.eq_zero_of_dvd_of_lt (pow_padicValNat_dvd (p := p) (n := n)) Hnot) + + lemma le_nat_log_gen {p n₁ n₂ : ℕ} [Fact (Nat.Prime p)] (hn : n₁ ≤ n₂): + padicValNat p n₁ ≤ Nat.log p n₂ := le_trans (le_nat_log n₁) (Nat.log_mono_right hn) + + /-- The p-adic valuation of `n` is equal to the logarithm w.r.t `p` iff + `n` is less than `p` raised to one plus the p-adic valuation of `n`. -/ + lemma nat_log_eq_padicvalnat_iff {p : ℕ} [hp : Fact (Nat.Prime p)] (n : ℕ)(hn : 0 < n): + Nat.log p n = padicValNat p n ↔ n < p^(padicValNat p n + 1) := by + · rw [Nat.log_eq_iff (Or.inr ⟨(Nat.Prime.one_lt' p).out, by linarith⟩)] + · rw [and_iff_right_iff_imp] + exact (fun _ => Nat.le_of_dvd hn pow_padicValNat_dvd) + open BigOperators theorem range_pow_padicValNat_subset_divisors {n : ℕ} (hn : n ≠ 0) : From 90fd8dc5f96aa48efe36838856feaf9da422a649 Mon Sep 17 00:00:00 2001 From: Koundinya Vajjha Date: Thu, 21 Sep 2023 22:59:15 -0700 Subject: [PATCH 29/74] remove unneeded argument --- Mathlib/NumberTheory/Padics/PadicVal.lean | 6 +++--- 1 file changed, 3 insertions(+), 3 deletions(-) diff --git a/Mathlib/NumberTheory/Padics/PadicVal.lean b/Mathlib/NumberTheory/Padics/PadicVal.lean index daff8451efe90..1837fc7d07a29 100644 --- a/Mathlib/NumberTheory/Padics/PadicVal.lean +++ b/Mathlib/NumberTheory/Padics/PadicVal.lean @@ -584,7 +584,7 @@ theorem padicValNat_primes {q : ℕ} [hp : Fact p.Prime] [hq : Fact q.Prime] (ne #align padic_val_nat_primes padicValNat_primes /-- The p-adic valuation of `n` is less than or equal to its logarithm w.r.t `p`.-/ - lemma le_nat_log {p : ℕ} [hp : Fact (Nat.Prime p)] (n : ℕ): + lemma le_nat_log (n : ℕ) [hp : Fact (Nat.Prime p)]: padicValNat p n ≤ Nat.log p n := by by_cases (n = 0) · simp only [h, padicValNat.zero, Nat.log_zero_right, le_refl] @@ -592,12 +592,12 @@ theorem padicValNat_primes {q : ℕ} [hp : Fact p.Prime] [hq : Fact q.Prime] (ne by_contra Hnot; push_neg at Hnot exact h (Nat.eq_zero_of_dvd_of_lt (pow_padicValNat_dvd (p := p) (n := n)) Hnot) - lemma le_nat_log_gen {p n₁ n₂ : ℕ} [Fact (Nat.Prime p)] (hn : n₁ ≤ n₂): + lemma le_nat_log_gen {n₁ n₂ : ℕ} [Fact (Nat.Prime p)] (hn : n₁ ≤ n₂): padicValNat p n₁ ≤ Nat.log p n₂ := le_trans (le_nat_log n₁) (Nat.log_mono_right hn) /-- The p-adic valuation of `n` is equal to the logarithm w.r.t `p` iff `n` is less than `p` raised to one plus the p-adic valuation of `n`. -/ - lemma nat_log_eq_padicvalnat_iff {p : ℕ} [hp : Fact (Nat.Prime p)] (n : ℕ)(hn : 0 < n): + lemma nat_log_eq_padicvalnat_iff {n : ℕ} [hp : Fact (Nat.Prime p)] (hn : 0 < n): Nat.log p n = padicValNat p n ↔ n < p^(padicValNat p n + 1) := by · rw [Nat.log_eq_iff (Or.inr ⟨(Nat.Prime.one_lt' p).out, by linarith⟩)] · rw [and_iff_right_iff_imp] From f70cd3451b422ded521b4ae1fd908e012081b3c6 Mon Sep 17 00:00:00 2001 From: Koundinya Vajjha Date: Thu, 21 Sep 2023 23:01:22 -0700 Subject: [PATCH 30/74] counterexample in right place --- Counterexamples.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Counterexamples.lean b/Counterexamples.lean index b3b4e15cc9478..c1feedde8d404 100644 --- a/Counterexamples.lean +++ b/Counterexamples.lean @@ -3,6 +3,7 @@ import Counterexamples.CharPZeroNeCharZero import Counterexamples.Cyclotomic105 import Counterexamples.DirectSumIsInternal import Counterexamples.Girard +import Counterexamples.HarmonicNotInteger import Counterexamples.HomogeneousPrimeNotPrime import Counterexamples.LinearOrderWithPosMulPosEqZero import Counterexamples.MapFloor @@ -13,4 +14,3 @@ import Counterexamples.QuadraticForm import Counterexamples.SeminormLatticeNotDistrib import Counterexamples.SorgenfreyLine import Counterexamples.ZeroDivisorsInAddMonoidAlgebras -import Counterexamples.HarmonicNotInteger From d44e4e63f6a80e99e367fcadd13444026eb1cea3 Mon Sep 17 00:00:00 2001 From: Koundinya Vajjha Date: Fri, 22 Sep 2023 16:17:34 -0700 Subject: [PATCH 31/74] fix: build --- Counterexamples/HarmonicNotInteger.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Counterexamples/HarmonicNotInteger.lean b/Counterexamples/HarmonicNotInteger.lean index 66e3f749621b8..17c30afbaf243 100644 --- a/Counterexamples/HarmonicNotInteger.lean +++ b/Counterexamples/HarmonicNotInteger.lean @@ -111,7 +111,7 @@ lemma padicValRat_ge_neg_Nat_log_ge {p n : ℕ}[hp : Fact (Nat.Prime p)]: ∀ q, q ≤ n → -Nat.log p n ≤ padicValRat p (1 / q) := by intros q Hq rw [one_div,padicValRat.inv,padicValRat.of_nat, neg_le_neg_iff, Nat.cast_le] - exact padicValNat.le_nat_log_gen Hq + exact le_nat_log_gen Hq lemma padicValRat_ge_neg_Nat_log_lt (n : ℕ): ∀ q, 0 < q ∧ q ≤ n → q ≠ 2^Nat.log 2 n → From 7fa0088f366fa7f1c50a9b6ffd81b5f83ab86e56 Mon Sep 17 00:00:00 2001 From: Koundinya Vajjha Date: Fri, 22 Sep 2023 17:15:43 -0700 Subject: [PATCH 32/74] address comments --- Counterexamples/HarmonicNotInteger.lean | 113 +++++++++++----------- Mathlib/NumberTheory/Padics/PadicVal.lean | 58 ++++++----- 2 files changed, 82 insertions(+), 89 deletions(-) diff --git a/Counterexamples/HarmonicNotInteger.lean b/Counterexamples/HarmonicNotInteger.lean index 17c30afbaf243..cc22d37358673 100644 --- a/Counterexamples/HarmonicNotInteger.lean +++ b/Counterexamples/HarmonicNotInteger.lean @@ -9,7 +9,7 @@ import Mathlib.Data.Int.Log /-! The nth Harmonic number is not an integer. We formalize the proof using -2-adic valuations. This proof is due to K¨ursch´ak. +2-adic valuations. This proof is due to Kürschák. Reference: https://kconrad.math.uconn.edu/blurbs/gradnumthy/padicharmonicsum.pdf @@ -24,42 +24,43 @@ open BigOperators /-- The nth-harmonic number defined as a finset sum of consecutive reciprocals. -/ def harmonic : ℕ → ℚ := fun n => ∑ i in Finset.range n, 1 / (i + 1) -lemma harmonic_pos : ∀ n, n ≠ 0 → 0 < harmonic n := fun n Hn => +lemma harmonic_pos (n : ℕ) (Hn : n ≠ 0) : 0 < harmonic n := Finset.sum_pos (fun _ _ => div_pos zero_lt_one (by norm_cast; linarith)) - (by (rwa [Finset.nonempty_range_iff])) + (by rwa [Finset.nonempty_range_iff]) lemma harmonic_singleton {n c : ℕ} (hc : c ∈ Finset.range n): - harmonic n = 1 / (c + 1) + ∑ x in Finset.range n \ {c}, 1 / ((x : ℚ) + 1) := by - unfold harmonic - rwa [add_comm, Finset.sum_eq_sum_diff_singleton_add (i := c)] + harmonic n = 1 / (c + 1) + ∑ x in Finset.range n \ {c}, 1 / ((x : ℚ) + 1) := by + unfold harmonic + rwa [add_comm, Finset.sum_eq_sum_diff_singleton_add (i := c)] lemma finset_range_sdiff_singleton_nonempty {c n : ℕ} (hn : 2 ≤ n): - Finset.Nonempty (Finset.range n \ {c}) := by - rw [Finset.sdiff_nonempty,Finset.subset_singleton_iff, Finset.range_eq_empty_iff,not_or] - refine' ⟨by linarith,fun Hnot => _⟩ - suffices : n = 1; linarith + Finset.Nonempty (Finset.range n \ {c}) := by + rw [Finset.sdiff_nonempty, Finset.subset_singleton_iff, Finset.range_eq_empty_iff, not_or] + refine' ⟨by linarith, fun Hnot => _⟩ + suffices n = 1 by linarith rw [← Finset.card_range n, ← Finset.card_singleton c, Hnot] lemma harmonic_singleton_ne_zero {c n : ℕ} (hn : 2 ≤ n) : - (∑ x in Finset.range n \ {c}, 1 / ((x + 1 : ℚ)) ≠ 0) := by + ∑ x in Finset.range n \ {c}, 1 / (x + 1 : ℚ) ≠ 0 := by refine' ne_of_gt <| Finset.sum_pos (fun i _ => div_pos zero_lt_one _) (finset_range_sdiff_singleton_nonempty hn) norm_cast; simp only [add_pos_iff, or_true] /-- If a rational is not a 2-adic integer, it is not an integer. -/ theorem not_int_of_not_padic_int {a : ℚ} (H : 1 < padicNorm 2 a): - ¬ a.isInt := by + ¬ a.isInt := by suffices : a.den ≠ 1; simpa [Rat.isInt] - by_cases (a = 0) - · subst h; try contradiction + by_cases h : a = 0 + · subst h; contradiction · unfold padicNorm at H - split_ifs at H; try contradiction - suffices : padicValRat 2 a < 0; unfold padicValRat at this - intro Hden - rw [Hden] at this; simp only [padicValNat.one, sub_zero] at this - norm_cast + split_ifs at H; contradiction + suffices padicValRat 2 a < 0 by + unfold padicValRat at this + intro Hden + rw [Hden] at this; simp only [padicValNat.one, sub_zero] at this + norm_cast refine' neg_of_neg_pos _ have hx : (1 : ℚ) < 2 := by norm_num rw [← zpow_lt_iff_lt hx] @@ -67,21 +68,18 @@ theorem not_int_of_not_padic_int {a : ℚ} (H : 1 < padicNorm 2 a): lemma padicValRat_2_pow (r : ℕ) : padicValRat 2 (1 / 2^r) = -r := by - rw [one_div,padicValRat.inv,neg_inj,padicValRat.pow (by simp)] - suffices : padicValRat 2 2 = 1 - simp only [this, mul_one] + rw [one_div, padicValRat.inv, neg_inj, padicValRat.pow (by simp)] + suffices padicValRat 2 2 = 1 by simp only [this, mul_one] rw [←padicValRat.self (p := 2) one_lt_two]; simp only [Nat.cast_ofNat] /-- For `i` less than `n`, `2^Nat.log 2 n` does not divide `i` unless `i` is equal to it. -/ -lemma nat_log_not_dvd {n : ℕ} : - ∀ i, 0 < i ∧ i ≠ 2^(Nat.log 2 n) ∧ i ≤ n → ¬ 2^(Nat.log 2 n) ∣ i := by - rintro i ⟨Hi₁,Hi₂,Hi₃⟩ - simp only [Nat.instDvdNat] - push_neg; intros c Hc +lemma nat_log_not_dvd {n i : ℕ} (Hi₁ : 0 < i) (Hi₂ : i ≠ 2 ^ Nat.log 2 n) (Hi₃ : i ≤ n): + ¬ 2^(Nat.log 2 n) ∣ i := by + rintro ⟨c,Hc⟩ have Hle : n < 2 ^ (Nat.log 2 n + 1) := Nat.lt_pow_succ_log_self (by simp) n have Hpow : 2 ^ Nat.log 2 n * 2 = 2 ^(Nat.log 2 n + 1) := by exact rfl obtain ⟨k,Hk⟩ := Nat.even_or_odd' c - by_cases (k = 0) + by_cases h : k = 0 · subst h rw [mul_zero, zero_add] at Hk rcases Hk with rfl | rfl @@ -100,29 +98,25 @@ lemma nat_log_not_dvd {n : ℕ} : _ < 2 ^ (Nat.log 2 n + 1)*k + 2^ Nat.log 2 n := by linarith _ = i := Hc.symm -lemma padicValRat_ge_neg_Nat_log_ne {n : ℕ}: - ∀ q, 0 < q ∧ q ≤ n → q ≠ 2^Nat.log 2 n → padicValRat 2 (1 / q) ≠ -Nat.log 2 n := by - rintro q ⟨Hq₁,Hq₂⟩ Hq₃ - rw [one_div,padicValRat.inv,padicValRat.of_nat,ne_eq,neg_inj,Nat.cast_inj] - exact (fun Hnot => nat_log_not_dvd (n := n) q ⟨Hq₁,Hq₃,Hq₂⟩ +lemma padicValRat_ge_neg_nat_log_ne {n q : ℕ}(Hq₁ : 0 < q)(Hq₂ : q ≤ n)(Hq₃ : q ≠ 2^Nat.log 2 n): + padicValRat 2 (1 / q) ≠ -Nat.log 2 n := by + rw [one_div, padicValRat.inv, padicValRat.of_nat, ne_eq, neg_inj, Nat.cast_inj] + exact (fun Hnot => nat_log_not_dvd (n := n) Hq₁ Hq₃ Hq₂ (Hnot ▸ pow_padicValNat_dvd (p := 2) (n := q))) -lemma padicValRat_ge_neg_Nat_log_ge {p n : ℕ}[hp : Fact (Nat.Prime p)]: - ∀ q, q ≤ n → -Nat.log p n ≤ padicValRat p (1 / q) := by - intros q Hq - rw [one_div,padicValRat.inv,padicValRat.of_nat, neg_le_neg_iff, Nat.cast_le] +lemma padicValRat_ge_neg_Nat_log_ge {p n q : ℕ}[hp : Fact (Nat.Prime p)] (Hq : q ≤ n): + -Nat.log p n ≤ padicValRat p (1 / q) := by + rw [one_div, padicValRat.inv, padicValRat.of_nat, neg_le_neg_iff, Nat.cast_le] exact le_nat_log_gen Hq -lemma padicValRat_ge_neg_Nat_log_lt (n : ℕ): - ∀ q, 0 < q ∧ q ≤ n → q ≠ 2^Nat.log 2 n → - padicValRat 2 (1 / 2^Nat.log 2 n) < padicValRat 2 (1 / q) := by - rintro q ⟨Hq₁,Hq₂⟩ Hq₃ +lemma padicValRat_ge_neg_Nat_log_lt (n q : ℕ) (Hq₁ : 0 < q) (Hq₂ : q ≤ n) (Hq₃ : q ≠ 2^Nat.log 2 n): + padicValRat 2 (1 / 2^Nat.log 2 n) < padicValRat 2 (1 / q) := by rw [padicValRat_2_pow] exact (lt_of_le_of_ne - (padicValRat_ge_neg_Nat_log_ge (p := 2) q Hq₂) - (padicValRat_ge_neg_Nat_log_ne q ⟨Hq₁,Hq₂⟩ Hq₃).symm) + (padicValRat_ge_neg_Nat_log_ge (p := 2) Hq₂) + (padicValRat_ge_neg_nat_log_ne Hq₁ Hq₂ Hq₃).symm) -lemma pow2_log_in_finset {n : ℕ} (hn : 2 ≤ n): +lemma two_pow_log_mem_range {n : ℕ} (hn : 2 ≤ n): 2^(Nat.log 2 n) - 1 ∈ Finset.range n := Finset.mem_range.mpr <| lt_of_lt_of_le (Nat.sub_lt (pow_pos zero_lt_two _) zero_lt_one) (Nat.pow_log_le_self 2 (by linarith)) @@ -130,24 +124,27 @@ lemma pow2_log_in_finset {n : ℕ} (hn : 2 ≤ n): of n. -/ theorem harmonic_padicValRat_2 {n : ℕ} (Hn : 2 ≤ n): padicValRat 2 (harmonic n) = -Int.log 2 (n : ℚ) := by - rw [Int.log_natCast,harmonic_singleton (pow2_log_in_finset Hn)] + rw [Int.log_natCast,harmonic_singleton (two_pow_log_mem_range Hn)] simp only [gt_iff_lt, pow_pos, Nat.cast_pred, Nat.cast_pow, Nat.cast_ofNat, sub_add_cancel] - rw [padicValRat.add_eq_of_lt (p := 2) _ (one_div_ne_zero <| pow_ne_zero _ two_ne_zero) - (harmonic_singleton_ne_zero (Hn)) _]; apply padicValRat_2_pow + rw [padicValRat.add_eq_of_lt (p := 2) _ + (one_div_ne_zero <| pow_ne_zero _ two_ne_zero) (harmonic_singleton_ne_zero Hn) _] + apply padicValRat_2_pow · have := harmonic_pos n - rw [harmonic_singleton (pow2_log_in_finset Hn)] at this + rw [harmonic_singleton (two_pow_log_mem_range Hn)] at this refine' ne_of_gt _ simp only [ne_eq, ge_iff_le, gt_iff_lt, pow_pos, Nat.cast_pred, Nat.cast_pow, Nat.cast_ofNat, sub_add_cancel, Finset.mem_range, not_lt, Finset.singleton_subset_iff] at this - exact (this (by linarith)) + exact this (by linarith) · have := padicValRat.finset_gen_sum_lt_of_lt (p := 2) (j := 2 ^ Nat.log 2 n - 1) (finset_range_sdiff_singleton_nonempty Hn) (F := fun n => 1 / ((n + 1) : ℚ)) (S := Finset.range n \ {2^Nat.log 2 n - 1}) - simp [-one_div] at * - refine' this (fun i Hi₁ Hi₂ => _) (fun i => div_pos zero_lt_one (by norm_cast;linarith)) + simp only [Finset.mem_range, Finset.mem_sdiff, Finset.mem_singleton, + Finset.singleton_subset_iff, pow_pos, Nat.cast_pred, Nat.cast_pow, + sub_add_cancel, and_imp,gt_iff_lt] at this + refine' this (fun i Hi₁ Hi₂ => _) (fun i => div_pos zero_lt_one (by norm_cast; linarith)) · have := padicValRat_ge_neg_Nat_log_lt (n := n) (i+1) simp only [Nat.cast_add, Nat.cast_one] at this - refine' this ⟨by linarith, by linarith⟩ (fun Hnot => _) + refine' this (by linarith) (by linarith) (fun Hnot => _) rw [← Hnot] at Hi₂ refine' Hi₂ _ simp only [add_tsub_cancel_right] @@ -157,10 +154,8 @@ theorem harmonic_not_int : ∀ n, n ≥ 2 → ¬ (harmonic n).isInt := by intro n Hn refine' not_int_of_not_padic_int _ dsimp [padicNorm] - split_ifs with h - · exfalso - exact (ne_of_gt <| harmonic_pos n (by linarith)) h - · refine' one_lt_zpow one_lt_two _ (lt_neg.mp _) - rw [neg_zero, harmonic_padicValRat_2 Hn,Int.log_natCast, Left.neg_neg_iff, - Nat.cast_pos, Nat.log_pos_iff] - exact ⟨Hn,one_lt_two⟩ + rw [if_neg (ne_of_gt <| harmonic_pos n (by linarith))] + refine' one_lt_zpow one_lt_two _ (lt_neg.mp _) + rw [neg_zero, harmonic_padicValRat_2 Hn,Int.log_natCast, Left.neg_neg_iff, + Nat.cast_pos, Nat.log_pos_iff] + exact ⟨Hn,one_lt_two⟩ diff --git a/Mathlib/NumberTheory/Padics/PadicVal.lean b/Mathlib/NumberTheory/Padics/PadicVal.lean index 1837fc7d07a29..17f795e058536 100644 --- a/Mathlib/NumberTheory/Padics/PadicVal.lean +++ b/Mathlib/NumberTheory/Padics/PadicVal.lean @@ -433,18 +433,21 @@ theorem min_le_padicValRat_add {q r : ℚ} (hqr : q + r ≠ 0) : /-- Ultrametric property of a p-adic valuation. -/ lemma min_eq_padicValRat_add {q r : ℚ} (hqr : q + r ≠ 0) (hq : q ≠ 0) (hr : r ≠ 0) -(hval : padicValRat p q ≠ padicValRat p r) : -padicValRat p (q + r) = min (padicValRat p q) (padicValRat p r) := by + (hval : padicValRat p q ≠ padicValRat p r) : + padicValRat p (q + r) = min (padicValRat p q) (padicValRat p r) := by have Hmin := padicValRat.min_le_padicValRat_add (p := p) (hp := hp) hqr wlog h : padicValRat p q < padicValRat p r generalizing q r with Hgen · push_neg at h; rw [add_comm, min_comm] - exact (Hgen (by rwa [add_comm]) hr hq hval.symm - (by rwa [min_comm,add_comm]) - (Ne.lt_of_le (Ne.symm hval) h)) + exact Hgen (by rwa [add_comm]) hr hq hval.symm + (by rwa [min_comm, add_comm]) + (hval.symm.lt_of_le h) · rw [min_eq_left (le_of_lt h)] at Hmin ⊢ suffices Hreq : padicValRat p q ≥ padicValRat p (q + r) by linarith suffices Haux : padicValRat p q ≥ min (padicValRat p (q + r)) (padicValRat p r) by - rw [min_def] at Haux; split_ifs at Haux with Hspl; try assumption; linarith + rw [min_def] at Haux + split_ifs at Haux with Hspl + · assumption + · linarith calc padicValRat p q = padicValRat p ((q + r) - r) := by congr; simp _ ≥ min (padicValRat p (q + r)) (padicValRat p (-r)) := ge_iff_le.mp <| le_trans (padicValRat.min_le_padicValRat_add (q := q+r) (r := -r) (by simpa)) @@ -459,7 +462,7 @@ by rw [min_eq_padicValRat_add hqr hq hr (ne_of_lt hval),min_eq_left (le_of_lt hv lemma add_lt_of_lt {q r₁ r₂ : ℚ} (hqr : r₁ + r₂ ≠ 0) (hval₁ : padicValRat p q < padicValRat p r₁) (hval₂ : padicValRat p q < padicValRat p r₂) : padicValRat p q < padicValRat p (r₁ + r₂) := - lt_of_lt_of_le (lt_min hval₁ hval₂) (padicValRat.min_le_padicValRat_add (p := p) (hp := hp) hqr) + lt_of_lt_of_le (lt_min hval₁ hval₂) (padicValRat.min_le_padicValRat_add hqr) open BigOperators @@ -478,15 +481,12 @@ theorem sum_pos_of_pos {n : ℕ} {F : ℕ → ℚ} (hF : ∀ i, i < n → 0 < pa exact hF _ (lt_trans hi (lt_add_one _)) #align padic_val_rat.sum_pos_of_pos padicValRat.sum_pos_of_pos -/-- - If the p-adic valuation of a finite set of positive rationals is greater than a given rational - number, then the p-adic valuation of their sum is also greater than the same rational number. --/ +/-- If the p-adic valuation of a finite set of positive rationals is greater than a given rational +number, then the p-adic valuation of their sum is also greater than the same rational number. -/ theorem finset_gen_sum_lt_of_lt {p j : ℕ} [hp : Fact (Nat.Prime p)] {F : ℕ → ℚ} {S : Finset ℕ} (hS : S.Nonempty) -(hF : ∀ i, i ∈ S → padicValRat p (F j) < padicValRat p (F i)) -(hn1 : ∀ i : ℕ, 0 < F i): -padicValRat p (F j) < padicValRat p (∑ i in S, F i) := by +(hF : ∀ i, i ∈ S → padicValRat p (F j) < padicValRat p (F i)) (hn1 : ∀ i : ℕ, 0 < F i): + padicValRat p (F j) < padicValRat p (∑ i in S, F i) := by induction' hS using Finset.Nonempty.cons_induction with k s S' Hnot Hne Hind · rw [Finset.sum_singleton] exact hF k (by simp) @@ -583,25 +583,23 @@ theorem padicValNat_primes {q : ℕ} [hp : Fact p.Prime] [hq : Fact q.Prime] (ne (not_congr (Iff.symm (prime_dvd_prime_iff_eq hp.1 hq.1))).mp neq #align padic_val_nat_primes padicValNat_primes - /-- The p-adic valuation of `n` is less than or equal to its logarithm w.r.t `p`.-/ - lemma le_nat_log (n : ℕ) [hp : Fact (Nat.Prime p)]: - padicValNat p n ≤ Nat.log p n := by - by_cases (n = 0) - · simp only [h, padicValNat.zero, Nat.log_zero_right, le_refl] - · refine' Nat.le_log_of_pow_le (Nat.Prime.one_lt hp.elim) _ - by_contra Hnot; push_neg at Hnot - exact h (Nat.eq_zero_of_dvd_of_lt (pow_padicValNat_dvd (p := p) (n := n)) Hnot) +/-- The p-adic valuation of `n` is less than or equal to its logarithm w.r.t `p`.-/ +lemma le_nat_log (n : ℕ) [hp : Fact (Nat.Prime p)]: padicValNat p n ≤ Nat.log p n := by + by_cases (n = 0) + · simp only [h, padicValNat.zero, Nat.log_zero_right, le_refl] + · refine' Nat.le_log_of_pow_le (Nat.Prime.one_lt hp.elim) _ + by_contra Hnot; push_neg at Hnot + exact h (Nat.eq_zero_of_dvd_of_lt (pow_padicValNat_dvd (p := p) (n := n)) Hnot) - lemma le_nat_log_gen {n₁ n₂ : ℕ} [Fact (Nat.Prime p)] (hn : n₁ ≤ n₂): +lemma le_nat_log_gen {n₁ n₂ : ℕ} [Fact (Nat.Prime p)] (hn : n₁ ≤ n₂): padicValNat p n₁ ≤ Nat.log p n₂ := le_trans (le_nat_log n₁) (Nat.log_mono_right hn) - /-- The p-adic valuation of `n` is equal to the logarithm w.r.t `p` iff - `n` is less than `p` raised to one plus the p-adic valuation of `n`. -/ - lemma nat_log_eq_padicvalnat_iff {n : ℕ} [hp : Fact (Nat.Prime p)] (hn : 0 < n): - Nat.log p n = padicValNat p n ↔ n < p^(padicValNat p n + 1) := by - · rw [Nat.log_eq_iff (Or.inr ⟨(Nat.Prime.one_lt' p).out, by linarith⟩)] - · rw [and_iff_right_iff_imp] - exact (fun _ => Nat.le_of_dvd hn pow_padicValNat_dvd) +/-- The p-adic valuation of `n` is equal to the logarithm w.r.t `p` iff + `n` is less than `p` raised to one plus the p-adic valuation of `n`. -/ +lemma nat_log_eq_padicvalnat_iff {n : ℕ} [hp : Fact (Nat.Prime p)] (hn : 0 < n): + Nat.log p n = padicValNat p n ↔ n < p^(padicValNat p n + 1) := by + rw [Nat.log_eq_iff (Or.inr ⟨(Nat.Prime.one_lt' p).out, by linarith⟩), and_iff_right_iff_imp] + exact (fun _ => Nat.le_of_dvd hn pow_padicValNat_dvd) open BigOperators From d7f092f78bc846941cd97b2a3ff69b6387adbbf0 Mon Sep 17 00:00:00 2001 From: Koundinya Vajjha Date: Mon, 25 Sep 2023 15:43:48 -0700 Subject: [PATCH 33/74] various indentation suggestions + renaming --- Counterexamples/HarmonicNotInteger.lean | 34 +++++++++++------------ Mathlib/NumberTheory/Padics/PadicVal.lean | 29 ++++++++++--------- 2 files changed, 31 insertions(+), 32 deletions(-) diff --git a/Counterexamples/HarmonicNotInteger.lean b/Counterexamples/HarmonicNotInteger.lean index cc22d37358673..4bad4e7f3780c 100644 --- a/Counterexamples/HarmonicNotInteger.lean +++ b/Counterexamples/HarmonicNotInteger.lean @@ -25,16 +25,15 @@ open BigOperators def harmonic : ℕ → ℚ := fun n => ∑ i in Finset.range n, 1 / (i + 1) lemma harmonic_pos (n : ℕ) (Hn : n ≠ 0) : 0 < harmonic n := - Finset.sum_pos (fun _ _ => div_pos zero_lt_one - (by norm_cast; linarith)) + Finset.sum_pos (fun _ _ => div_pos zero_lt_one (by norm_cast; linarith)) (by rwa [Finset.nonempty_range_iff]) -lemma harmonic_singleton {n c : ℕ} (hc : c ∈ Finset.range n): +lemma harmonic_singleton {n c : ℕ} (hc : c ∈ Finset.range n) : harmonic n = 1 / (c + 1) + ∑ x in Finset.range n \ {c}, 1 / ((x : ℚ) + 1) := by unfold harmonic rwa [add_comm, Finset.sum_eq_sum_diff_singleton_add (i := c)] -lemma finset_range_sdiff_singleton_nonempty {c n : ℕ} (hn : 2 ≤ n): +lemma finset_range_sdiff_singleton_nonempty {c n : ℕ} (hn : 2 ≤ n) : Finset.Nonempty (Finset.range n \ {c}) := by rw [Finset.sdiff_nonempty, Finset.subset_singleton_iff, Finset.range_eq_empty_iff, not_or] refine' ⟨by linarith, fun Hnot => _⟩ @@ -49,7 +48,7 @@ lemma harmonic_singleton_ne_zero {c n : ℕ} (hn : 2 ≤ n) : norm_cast; simp only [add_pos_iff, or_true] /-- If a rational is not a 2-adic integer, it is not an integer. -/ -theorem not_int_of_not_padic_int {a : ℚ} (H : 1 < padicNorm 2 a): +theorem not_int_of_not_padic_int {a : ℚ} (H : 1 < padicNorm 2 a) : ¬ a.isInt := by suffices : a.den ≠ 1; simpa [Rat.isInt] by_cases h : a = 0 @@ -73,7 +72,7 @@ lemma padicValRat_2_pow (r : ℕ) : padicValRat 2 (1 / 2^r) = -r := by rw [←padicValRat.self (p := 2) one_lt_two]; simp only [Nat.cast_ofNat] /-- For `i` less than `n`, `2^Nat.log 2 n` does not divide `i` unless `i` is equal to it. -/ -lemma nat_log_not_dvd {n i : ℕ} (Hi₁ : 0 < i) (Hi₂ : i ≠ 2 ^ Nat.log 2 n) (Hi₃ : i ≤ n): +lemma nat_log_not_dvd {n i : ℕ} (Hi₁ : 0 < i) (Hi₂ : i ≠ 2 ^ Nat.log 2 n) (Hi₃ : i ≤ n) : ¬ 2^(Nat.log 2 n) ∣ i := by rintro ⟨c,Hc⟩ have Hle : n < 2 ^ (Nat.log 2 n + 1) := Nat.lt_pow_succ_log_self (by simp) n @@ -98,7 +97,7 @@ lemma nat_log_not_dvd {n i : ℕ} (Hi₁ : 0 < i) (Hi₂ : i ≠ 2 ^ Nat.log 2 n _ < 2 ^ (Nat.log 2 n + 1)*k + 2^ Nat.log 2 n := by linarith _ = i := Hc.symm -lemma padicValRat_ge_neg_nat_log_ne {n q : ℕ}(Hq₁ : 0 < q)(Hq₂ : q ≤ n)(Hq₃ : q ≠ 2^Nat.log 2 n): +lemma padicValRat_ge_neg_nat_log_ne {n q : ℕ}(Hq₁ : 0 < q)(Hq₂ : q ≤ n)(Hq₃ : q ≠ 2^Nat.log 2 n) : padicValRat 2 (1 / q) ≠ -Nat.log 2 n := by rw [one_div, padicValRat.inv, padicValRat.of_nat, ne_eq, neg_inj, Nat.cast_inj] exact (fun Hnot => nat_log_not_dvd (n := n) Hq₁ Hq₃ Hq₂ @@ -107,23 +106,24 @@ lemma padicValRat_ge_neg_nat_log_ne {n q : ℕ}(Hq₁ : 0 < q)(Hq₂ : q ≤ n)( lemma padicValRat_ge_neg_Nat_log_ge {p n q : ℕ}[hp : Fact (Nat.Prime p)] (Hq : q ≤ n): -Nat.log p n ≤ padicValRat p (1 / q) := by rw [one_div, padicValRat.inv, padicValRat.of_nat, neg_le_neg_iff, Nat.cast_le] - exact le_nat_log_gen Hq + exact le_nat_log_of_le Hq -lemma padicValRat_ge_neg_Nat_log_lt (n q : ℕ) (Hq₁ : 0 < q) (Hq₂ : q ≤ n) (Hq₃ : q ≠ 2^Nat.log 2 n): +lemma padicValRat_ge_neg_Nat_log_lt (n q : ℕ) (Hq₁ : 0 < q) (Hq₂ : q ≤ n)(Hq₃ : q ≠ 2^Nat.log 2 n) : padicValRat 2 (1 / 2^Nat.log 2 n) < padicValRat 2 (1 / q) := by rw [padicValRat_2_pow] exact (lt_of_le_of_ne (padicValRat_ge_neg_Nat_log_ge (p := 2) Hq₂) (padicValRat_ge_neg_nat_log_ne Hq₁ Hq₂ Hq₃).symm) -lemma two_pow_log_mem_range {n : ℕ} (hn : 2 ≤ n): - 2^(Nat.log 2 n) - 1 ∈ Finset.range n := Finset.mem_range.mpr <| lt_of_lt_of_le - (Nat.sub_lt (pow_pos zero_lt_two _) zero_lt_one) (Nat.pow_log_le_self 2 (by linarith)) +lemma two_pow_log_mem_range {n : ℕ} (hn : 2 ≤ n) : + 2^(Nat.log 2 n) - 1 ∈ Finset.range n := + Finset.mem_range.mpr <| lt_of_lt_of_le (Nat.sub_lt (pow_pos zero_lt_two _) zero_lt_one) + (Nat.pow_log_le_self 2 (by linarith)) /-- The 2-adic valuation of the n-th harmonic number is the negative of the logarithm of n. -/ -theorem harmonic_padicValRat_2 {n : ℕ} (Hn : 2 ≤ n): - padicValRat 2 (harmonic n) = -Int.log 2 (n : ℚ) := by +theorem harmonic_padicValRat_2 {n : ℕ} (Hn : 2 ≤ n) : + padicValRat 2 (harmonic n) = -Int.log 2 (n : ℚ) := by rw [Int.log_natCast,harmonic_singleton (two_pow_log_mem_range Hn)] simp only [gt_iff_lt, pow_pos, Nat.cast_pred, Nat.cast_pow, Nat.cast_ofNat, sub_add_cancel] rw [padicValRat.add_eq_of_lt (p := 2) _ @@ -135,7 +135,7 @@ theorem harmonic_padicValRat_2 {n : ℕ} (Hn : 2 ≤ n): simp only [ne_eq, ge_iff_le, gt_iff_lt, pow_pos, Nat.cast_pred, Nat.cast_pow, Nat.cast_ofNat, sub_add_cancel, Finset.mem_range, not_lt, Finset.singleton_subset_iff] at this exact this (by linarith) - · have := padicValRat.finset_gen_sum_lt_of_lt (p := 2) (j := 2 ^ Nat.log 2 n - 1) + · have := padicValRat.sum_gt_of_gt (p := 2) (j := 2 ^ Nat.log 2 n - 1) (finset_range_sdiff_singleton_nonempty Hn) (F := fun n => 1 / ((n + 1) : ℚ)) (S := Finset.range n \ {2^Nat.log 2 n - 1}) simp only [Finset.mem_range, Finset.mem_sdiff, Finset.mem_singleton, @@ -150,8 +150,8 @@ theorem harmonic_padicValRat_2 {n : ℕ} (Hn : 2 ≤ n): simp only [add_tsub_cancel_right] /-- The n-th harmonic number is not an integer for n ≥ 2. -/ -theorem harmonic_not_int : ∀ n, n ≥ 2 → ¬ (harmonic n).isInt := by - intro n Hn +theorem harmonic_not_int {n : ℕ} (Hn : 2 ≤ n) : + ¬(harmonic n).isInt := by refine' not_int_of_not_padic_int _ dsimp [padicNorm] rw [if_neg (ne_of_gt <| harmonic_pos n (by linarith))] diff --git a/Mathlib/NumberTheory/Padics/PadicVal.lean b/Mathlib/NumberTheory/Padics/PadicVal.lean index 17f795e058536..5447753b77e4a 100644 --- a/Mathlib/NumberTheory/Padics/PadicVal.lean +++ b/Mathlib/NumberTheory/Padics/PadicVal.lean @@ -442,8 +442,8 @@ lemma min_eq_padicValRat_add {q r : ℚ} (hqr : q + r ≠ 0) (hq : q ≠ 0) (hr (by rwa [min_comm, add_comm]) (hval.symm.lt_of_le h) · rw [min_eq_left (le_of_lt h)] at Hmin ⊢ - suffices Hreq : padicValRat p q ≥ padicValRat p (q + r) by linarith - suffices Haux : padicValRat p q ≥ min (padicValRat p (q + r)) (padicValRat p r) by + suffices Hreq : padicValRat p (q + r) ≤ padicValRat p q by linarith + suffices Haux : min (padicValRat p (q + r)) (padicValRat p r) ≤ padicValRat p q by rw [min_def] at Haux split_ifs at Haux with Hspl · assumption @@ -455,13 +455,13 @@ lemma min_eq_padicValRat_add {q r : ℚ} (hqr : q + r ≠ 0) (hq : q ≠ 0) (hr _ = min (padicValRat p (q + r)) (padicValRat p r) := by rw [padicValRat.neg] lemma add_eq_of_lt {q r : ℚ} (hqr : q + r ≠ 0) -(hq : q ≠ 0) (hr : r ≠ 0) (hval : padicValRat p q < padicValRat p r) : -padicValRat p (q + r) = padicValRat p q := -by rw [min_eq_padicValRat_add hqr hq hr (ne_of_lt hval),min_eq_left (le_of_lt hval)] + (hq : q ≠ 0) (hr : r ≠ 0) (hval : padicValRat p q < padicValRat p r) : + padicValRat p (q + r) = padicValRat p q := by + rw [min_eq_padicValRat_add hqr hq hr (ne_of_lt hval),min_eq_left (le_of_lt hval)] lemma add_lt_of_lt {q r₁ r₂ : ℚ} (hqr : r₁ + r₂ ≠ 0) -(hval₁ : padicValRat p q < padicValRat p r₁) (hval₂ : padicValRat p q < padicValRat p r₂) : - padicValRat p q < padicValRat p (r₁ + r₂) := + (hval₁ : padicValRat p q < padicValRat p r₁) (hval₂ : padicValRat p q < padicValRat p r₂) : + padicValRat p q < padicValRat p (r₁ + r₂) := lt_of_lt_of_le (lt_min hval₁ hval₂) (padicValRat.min_le_padicValRat_add hqr) open BigOperators @@ -483,10 +483,9 @@ theorem sum_pos_of_pos {n : ℕ} {F : ℕ → ℚ} (hF : ∀ i, i < n → 0 < pa /-- If the p-adic valuation of a finite set of positive rationals is greater than a given rational number, then the p-adic valuation of their sum is also greater than the same rational number. -/ -theorem finset_gen_sum_lt_of_lt {p j : ℕ} [hp : Fact (Nat.Prime p)] -{F : ℕ → ℚ} {S : Finset ℕ} (hS : S.Nonempty) -(hF : ∀ i, i ∈ S → padicValRat p (F j) < padicValRat p (F i)) (hn1 : ∀ i : ℕ, 0 < F i): - padicValRat p (F j) < padicValRat p (∑ i in S, F i) := by +theorem sum_gt_of_gt {p j : ℕ} [hp : Fact (Nat.Prime p)] {F : ℕ → ℚ} {S : Finset ℕ} + (hS : S.Nonempty) (hF : ∀ i, i ∈ S → padicValRat p (F j) < padicValRat p (F i)) + (hn1 : ∀ i : ℕ, 0 < F i) : padicValRat p (F j) < padicValRat p (∑ i in S, F i) := by induction' hS using Finset.Nonempty.cons_induction with k s S' Hnot Hne Hind · rw [Finset.sum_singleton] exact hF k (by simp) @@ -584,20 +583,20 @@ theorem padicValNat_primes {q : ℕ} [hp : Fact p.Prime] [hq : Fact q.Prime] (ne #align padic_val_nat_primes padicValNat_primes /-- The p-adic valuation of `n` is less than or equal to its logarithm w.r.t `p`.-/ -lemma le_nat_log (n : ℕ) [hp : Fact (Nat.Prime p)]: padicValNat p n ≤ Nat.log p n := by +lemma le_nat_log (n : ℕ) [hp : Fact (Nat.Prime p)] : padicValNat p n ≤ Nat.log p n := by by_cases (n = 0) · simp only [h, padicValNat.zero, Nat.log_zero_right, le_refl] · refine' Nat.le_log_of_pow_le (Nat.Prime.one_lt hp.elim) _ by_contra Hnot; push_neg at Hnot exact h (Nat.eq_zero_of_dvd_of_lt (pow_padicValNat_dvd (p := p) (n := n)) Hnot) -lemma le_nat_log_gen {n₁ n₂ : ℕ} [Fact (Nat.Prime p)] (hn : n₁ ≤ n₂): +lemma le_nat_log_of_le {n₁ n₂ : ℕ} [Fact (Nat.Prime p)] (hn : n₁ ≤ n₂) : padicValNat p n₁ ≤ Nat.log p n₂ := le_trans (le_nat_log n₁) (Nat.log_mono_right hn) /-- The p-adic valuation of `n` is equal to the logarithm w.r.t `p` iff `n` is less than `p` raised to one plus the p-adic valuation of `n`. -/ -lemma nat_log_eq_padicvalnat_iff {n : ℕ} [hp : Fact (Nat.Prime p)] (hn : 0 < n): - Nat.log p n = padicValNat p n ↔ n < p^(padicValNat p n + 1) := by +lemma nat_log_eq_padicValNat_iff {n : ℕ} [hp : Fact (Nat.Prime p)] (hn : 0 < n) : + Nat.log p n = padicValNat p n ↔ n < p ^ (padicValNat p n + 1) := by rw [Nat.log_eq_iff (Or.inr ⟨(Nat.Prime.one_lt' p).out, by linarith⟩), and_iff_right_iff_imp] exact (fun _ => Nat.le_of_dvd hn pow_padicValNat_dvd) From ff72121236d90fad2dc171b9cfa6c502596dfc37 Mon Sep 17 00:00:00 2001 From: Koundinya Vajjha Date: Mon, 25 Sep 2023 15:54:03 -0700 Subject: [PATCH 34/74] move to number theory + padics --- Counterexamples.lean | 1 - Mathlib.lean | 1 + .../NumberTheory/Padics/Harmonic.lean | 6 +----- 3 files changed, 2 insertions(+), 6 deletions(-) rename Counterexamples/HarmonicNotInteger.lean => Mathlib/NumberTheory/Padics/Harmonic.lean (99%) diff --git a/Counterexamples.lean b/Counterexamples.lean index c1feedde8d404..18cd76a5ceb5f 100644 --- a/Counterexamples.lean +++ b/Counterexamples.lean @@ -3,7 +3,6 @@ import Counterexamples.CharPZeroNeCharZero import Counterexamples.Cyclotomic105 import Counterexamples.DirectSumIsInternal import Counterexamples.Girard -import Counterexamples.HarmonicNotInteger import Counterexamples.HomogeneousPrimeNotPrime import Counterexamples.LinearOrderWithPosMulPosEqZero import Counterexamples.MapFloor diff --git a/Mathlib.lean b/Mathlib.lean index ceee3fd57cb97..3579dc013e7c6 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -2609,6 +2609,7 @@ import Mathlib.NumberTheory.NumberField.Discriminant import Mathlib.NumberTheory.NumberField.Embeddings import Mathlib.NumberTheory.NumberField.Norm import Mathlib.NumberTheory.NumberField.Units +import Mathlib.NumberTheory.Padics.Harmonic import Mathlib.NumberTheory.Padics.Hensel import Mathlib.NumberTheory.Padics.PadicIntegers import Mathlib.NumberTheory.Padics.PadicNorm diff --git a/Counterexamples/HarmonicNotInteger.lean b/Mathlib/NumberTheory/Padics/Harmonic.lean similarity index 99% rename from Counterexamples/HarmonicNotInteger.lean rename to Mathlib/NumberTheory/Padics/Harmonic.lean index 4bad4e7f3780c..52b8e08a64b11 100644 --- a/Counterexamples/HarmonicNotInteger.lean +++ b/Mathlib/NumberTheory/Padics/Harmonic.lean @@ -16,9 +16,6 @@ https://kconrad.math.uconn.edu/blurbs/gradnumthy/padicharmonicsum.pdf -/ -namespace Counterexample - - open BigOperators /-- The nth-harmonic number defined as a finset sum of consecutive reciprocals. -/ @@ -40,7 +37,6 @@ lemma finset_range_sdiff_singleton_nonempty {c n : ℕ} (hn : 2 ≤ n) : suffices n = 1 by linarith rw [← Finset.card_range n, ← Finset.card_singleton c, Hnot] - lemma harmonic_singleton_ne_zero {c n : ℕ} (hn : 2 ≤ n) : ∑ x in Finset.range n \ {c}, 1 / (x + 1 : ℚ) ≠ 0 := by refine' ne_of_gt <| Finset.sum_pos (fun i _ => div_pos zero_lt_one _) @@ -103,7 +99,7 @@ lemma padicValRat_ge_neg_nat_log_ne {n q : ℕ}(Hq₁ : 0 < q)(Hq₂ : q ≤ n)( exact (fun Hnot => nat_log_not_dvd (n := n) Hq₁ Hq₃ Hq₂ (Hnot ▸ pow_padicValNat_dvd (p := 2) (n := q))) -lemma padicValRat_ge_neg_Nat_log_ge {p n q : ℕ}[hp : Fact (Nat.Prime p)] (Hq : q ≤ n): +lemma padicValRat_ge_neg_Nat_log_ge {p n q : ℕ}[hp : Fact (Nat.Prime p)](Hq : q ≤ n) : -Nat.log p n ≤ padicValRat p (1 / q) := by rw [one_div, padicValRat.inv, padicValRat.of_nat, neg_le_neg_iff, Nat.cast_le] exact le_nat_log_of_le Hq From be2bc38db6f45c55ce35888a2230248219681025 Mon Sep 17 00:00:00 2001 From: Koundinya Vajjha Date: Mon, 25 Sep 2023 15:58:37 -0700 Subject: [PATCH 35/74] padicvalrat_2_pow -> padicvalrat_two_pow_div --- Mathlib/NumberTheory/Padics/Harmonic.lean | 6 +++--- 1 file changed, 3 insertions(+), 3 deletions(-) diff --git a/Mathlib/NumberTheory/Padics/Harmonic.lean b/Mathlib/NumberTheory/Padics/Harmonic.lean index 52b8e08a64b11..7f48b124e10a1 100644 --- a/Mathlib/NumberTheory/Padics/Harmonic.lean +++ b/Mathlib/NumberTheory/Padics/Harmonic.lean @@ -62,7 +62,7 @@ theorem not_int_of_not_padic_int {a : ℚ} (H : 1 < padicNorm 2 a) : simpa only [zpow_zero] -lemma padicValRat_2_pow (r : ℕ) : padicValRat 2 (1 / 2^r) = -r := by +lemma padicValRat_two_pow_div (r : ℕ) : padicValRat 2 (1 / 2^r) = -r := by rw [one_div, padicValRat.inv, neg_inj, padicValRat.pow (by simp)] suffices padicValRat 2 2 = 1 by simp only [this, mul_one] rw [←padicValRat.self (p := 2) one_lt_two]; simp only [Nat.cast_ofNat] @@ -106,7 +106,7 @@ lemma padicValRat_ge_neg_Nat_log_ge {p n q : ℕ}[hp : Fact (Nat.Prime p)](Hq : lemma padicValRat_ge_neg_Nat_log_lt (n q : ℕ) (Hq₁ : 0 < q) (Hq₂ : q ≤ n)(Hq₃ : q ≠ 2^Nat.log 2 n) : padicValRat 2 (1 / 2^Nat.log 2 n) < padicValRat 2 (1 / q) := by - rw [padicValRat_2_pow] + rw [padicValRat_two_pow_div] exact (lt_of_le_of_ne (padicValRat_ge_neg_Nat_log_ge (p := 2) Hq₂) (padicValRat_ge_neg_nat_log_ne Hq₁ Hq₂ Hq₃).symm) @@ -124,7 +124,7 @@ theorem harmonic_padicValRat_2 {n : ℕ} (Hn : 2 ≤ n) : simp only [gt_iff_lt, pow_pos, Nat.cast_pred, Nat.cast_pow, Nat.cast_ofNat, sub_add_cancel] rw [padicValRat.add_eq_of_lt (p := 2) _ (one_div_ne_zero <| pow_ne_zero _ two_ne_zero) (harmonic_singleton_ne_zero Hn) _] - apply padicValRat_2_pow + apply padicValRat_two_pow_div · have := harmonic_pos n rw [harmonic_singleton (two_pow_log_mem_range Hn)] at this refine' ne_of_gt _ From 4c3a0e88ad345fba137b77dd99605d19371e4ed8 Mon Sep 17 00:00:00 2001 From: Koundinya Vajjha Date: Mon, 25 Sep 2023 16:56:53 -0700 Subject: [PATCH 36/74] generalize to arbitrary prime p --- Mathlib/NumberTheory/Padics/Harmonic.lean | 6 ++---- Mathlib/NumberTheory/Padics/PadicVal.lean | 8 ++++++++ 2 files changed, 10 insertions(+), 4 deletions(-) diff --git a/Mathlib/NumberTheory/Padics/Harmonic.lean b/Mathlib/NumberTheory/Padics/Harmonic.lean index 7f48b124e10a1..cb16c97b9c6ff 100644 --- a/Mathlib/NumberTheory/Padics/Harmonic.lean +++ b/Mathlib/NumberTheory/Padics/Harmonic.lean @@ -61,11 +61,9 @@ theorem not_int_of_not_padic_int {a : ℚ} (H : 1 < padicNorm 2 a) : rw [← zpow_lt_iff_lt hx] simpa only [zpow_zero] - lemma padicValRat_two_pow_div (r : ℕ) : padicValRat 2 (1 / 2^r) = -r := by - rw [one_div, padicValRat.inv, neg_inj, padicValRat.pow (by simp)] - suffices padicValRat 2 2 = 1 by simp only [this, mul_one] - rw [←padicValRat.self (p := 2) one_lt_two]; simp only [Nat.cast_ofNat] + rw [← padicValRat.self_pow_div (p := 2) (r := r)] + simp only [one_div, Nat.cast_ofNat] /-- For `i` less than `n`, `2^Nat.log 2 n` does not divide `i` unless `i` is equal to it. -/ lemma nat_log_not_dvd {n i : ℕ} (Hi₁ : 0 < i) (Hi₂ : i ≠ 2 ^ Nat.log 2 n) (Hi₃ : i ≤ n) : diff --git a/Mathlib/NumberTheory/Padics/PadicVal.lean b/Mathlib/NumberTheory/Padics/PadicVal.lean index 5447753b77e4a..88ca75d9a7bc1 100644 --- a/Mathlib/NumberTheory/Padics/PadicVal.lean +++ b/Mathlib/NumberTheory/Padics/PadicVal.lean @@ -464,6 +464,14 @@ lemma add_lt_of_lt {q r₁ r₂ : ℚ} (hqr : r₁ + r₂ ≠ 0) padicValRat p q < padicValRat p (r₁ + r₂) := lt_of_lt_of_le (lt_min hval₁ hval₂) (padicValRat.min_le_padicValRat_add hqr) +lemma self_pow_div (r : ℕ) : + padicValRat p (1 / (p : ℚ) ^ r) = -r := by + norm_cast; simp only [Nat.cast_pow] + rw [one_div,padicValRat.inv, neg_inj, padicValRat.pow, + padicValRat.self (Nat.Prime.one_lt hp.elim), mul_one] + simp only [ne_eq, Nat.cast_eq_zero] + exact Nat.Prime.ne_zero hp.elim + open BigOperators /-- A finite sum of rationals with positive `p`-adic valuation has positive `p`-adic valuation From faacc317d2d897241b8cc9f62ee9cc4872d67501 Mon Sep 17 00:00:00 2001 From: Koundinya Vajjha Date: Mon, 25 Sep 2023 21:56:12 -0700 Subject: [PATCH 37/74] rename --- Mathlib/NumberTheory/Padics/Harmonic.lean | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/Mathlib/NumberTheory/Padics/Harmonic.lean b/Mathlib/NumberTheory/Padics/Harmonic.lean index cb16c97b9c6ff..f816e6699bcdc 100644 --- a/Mathlib/NumberTheory/Padics/Harmonic.lean +++ b/Mathlib/NumberTheory/Padics/Harmonic.lean @@ -116,7 +116,7 @@ lemma two_pow_log_mem_range {n : ℕ} (hn : 2 ≤ n) : /-- The 2-adic valuation of the n-th harmonic number is the negative of the logarithm of n. -/ -theorem harmonic_padicValRat_2 {n : ℕ} (Hn : 2 ≤ n) : +theorem harmonic_two_adicValRat {n : ℕ} (Hn : 2 ≤ n) : padicValRat 2 (harmonic n) = -Int.log 2 (n : ℚ) := by rw [Int.log_natCast,harmonic_singleton (two_pow_log_mem_range Hn)] simp only [gt_iff_lt, pow_pos, Nat.cast_pred, Nat.cast_pow, Nat.cast_ofNat, sub_add_cancel] @@ -150,6 +150,6 @@ theorem harmonic_not_int {n : ℕ} (Hn : 2 ≤ n) : dsimp [padicNorm] rw [if_neg (ne_of_gt <| harmonic_pos n (by linarith))] refine' one_lt_zpow one_lt_two _ (lt_neg.mp _) - rw [neg_zero, harmonic_padicValRat_2 Hn,Int.log_natCast, Left.neg_neg_iff, + rw [neg_zero, harmonic_two_adicValRat Hn,Int.log_natCast, Left.neg_neg_iff, Nat.cast_pos, Nat.log_pos_iff] exact ⟨Hn,one_lt_two⟩ From 4d46a3b98a05e4f3294e267fb4b7e79bd5d2622b Mon Sep 17 00:00:00 2001 From: Koundinya Vajjha Date: Tue, 26 Sep 2023 07:53:16 -0700 Subject: [PATCH 38/74] more updates --- Mathlib/NumberTheory/Padics/Harmonic.lean | 10 +++++----- Mathlib/NumberTheory/Padics/PadicVal.lean | 1 + 2 files changed, 6 insertions(+), 5 deletions(-) diff --git a/Mathlib/NumberTheory/Padics/Harmonic.lean b/Mathlib/NumberTheory/Padics/Harmonic.lean index f816e6699bcdc..0db267f28a87b 100644 --- a/Mathlib/NumberTheory/Padics/Harmonic.lean +++ b/Mathlib/NumberTheory/Padics/Harmonic.lean @@ -21,7 +21,7 @@ open BigOperators /-- The nth-harmonic number defined as a finset sum of consecutive reciprocals. -/ def harmonic : ℕ → ℚ := fun n => ∑ i in Finset.range n, 1 / (i + 1) -lemma harmonic_pos (n : ℕ) (Hn : n ≠ 0) : 0 < harmonic n := +lemma harmonic_pos {n : ℕ} (Hn : n ≠ 0) : 0 < harmonic n := Finset.sum_pos (fun _ _ => div_pos zero_lt_one (by norm_cast; linarith)) (by rwa [Finset.nonempty_range_iff]) @@ -70,7 +70,7 @@ lemma nat_log_not_dvd {n i : ℕ} (Hi₁ : 0 < i) (Hi₂ : i ≠ 2 ^ Nat.log 2 n ¬ 2^(Nat.log 2 n) ∣ i := by rintro ⟨c,Hc⟩ have Hle : n < 2 ^ (Nat.log 2 n + 1) := Nat.lt_pow_succ_log_self (by simp) n - have Hpow : 2 ^ Nat.log 2 n * 2 = 2 ^(Nat.log 2 n + 1) := by exact rfl + have Hpow : 2 ^ Nat.log 2 n * 2 = 2 ^ (Nat.log 2 n + 1) := by rfl obtain ⟨k,Hk⟩ := Nat.even_or_odd' c by_cases h : k = 0 · subst h @@ -118,12 +118,12 @@ lemma two_pow_log_mem_range {n : ℕ} (hn : 2 ≤ n) : of n. -/ theorem harmonic_two_adicValRat {n : ℕ} (Hn : 2 ≤ n) : padicValRat 2 (harmonic n) = -Int.log 2 (n : ℚ) := by - rw [Int.log_natCast,harmonic_singleton (two_pow_log_mem_range Hn)] + rw [Int.log_natCast, harmonic_singleton (two_pow_log_mem_range Hn)] simp only [gt_iff_lt, pow_pos, Nat.cast_pred, Nat.cast_pow, Nat.cast_ofNat, sub_add_cancel] rw [padicValRat.add_eq_of_lt (p := 2) _ (one_div_ne_zero <| pow_ne_zero _ two_ne_zero) (harmonic_singleton_ne_zero Hn) _] apply padicValRat_two_pow_div - · have := harmonic_pos n + · have := harmonic_pos (n := n) rw [harmonic_singleton (two_pow_log_mem_range Hn)] at this refine' ne_of_gt _ simp only [ne_eq, ge_iff_le, gt_iff_lt, pow_pos, Nat.cast_pred, Nat.cast_pow, Nat.cast_ofNat, @@ -148,7 +148,7 @@ theorem harmonic_not_int {n : ℕ} (Hn : 2 ≤ n) : ¬(harmonic n).isInt := by refine' not_int_of_not_padic_int _ dsimp [padicNorm] - rw [if_neg (ne_of_gt <| harmonic_pos n (by linarith))] + rw [if_neg (ne_of_gt <| harmonic_pos (by linarith))] refine' one_lt_zpow one_lt_two _ (lt_neg.mp _) rw [neg_zero, harmonic_two_adicValRat Hn,Int.log_natCast, Left.neg_neg_iff, Nat.cast_pos, Nat.log_pos_iff] diff --git a/Mathlib/NumberTheory/Padics/PadicVal.lean b/Mathlib/NumberTheory/Padics/PadicVal.lean index 88ca75d9a7bc1..9d00b885f0a30 100644 --- a/Mathlib/NumberTheory/Padics/PadicVal.lean +++ b/Mathlib/NumberTheory/Padics/PadicVal.lean @@ -464,6 +464,7 @@ lemma add_lt_of_lt {q r₁ r₂ : ℚ} (hqr : r₁ + r₂ ≠ 0) padicValRat p q < padicValRat p (r₁ + r₂) := lt_of_lt_of_le (lt_min hval₁ hval₂) (padicValRat.min_le_padicValRat_add hqr) +@[simp] lemma self_pow_div (r : ℕ) : padicValRat p (1 / (p : ℚ) ^ r) = -r := by norm_cast; simp only [Nat.cast_pow] From e3d29255a6f27d463363d596a6c56e723ef327ad Mon Sep 17 00:00:00 2001 From: Koundinya Vajjha Date: Tue, 26 Sep 2023 08:17:28 -0700 Subject: [PATCH 39/74] address review comments --- Mathlib/NumberTheory/Padics/Harmonic.lean | 42 +++++++--------------- Mathlib/NumberTheory/Padics/PadicNorm.lean | 18 ++++++++++ 2 files changed, 31 insertions(+), 29 deletions(-) diff --git a/Mathlib/NumberTheory/Padics/Harmonic.lean b/Mathlib/NumberTheory/Padics/Harmonic.lean index 0db267f28a87b..6bc2731043345 100644 --- a/Mathlib/NumberTheory/Padics/Harmonic.lean +++ b/Mathlib/NumberTheory/Padics/Harmonic.lean @@ -27,7 +27,7 @@ lemma harmonic_pos {n : ℕ} (Hn : n ≠ 0) : 0 < harmonic n := lemma harmonic_singleton {n c : ℕ} (hc : c ∈ Finset.range n) : harmonic n = 1 / (c + 1) + ∑ x in Finset.range n \ {c}, 1 / ((x : ℚ) + 1) := by - unfold harmonic + dsimp [harmonic] rwa [add_comm, Finset.sum_eq_sum_diff_singleton_add (i := c)] lemma finset_range_sdiff_singleton_nonempty {c n : ℕ} (hn : 2 ≤ n) : @@ -43,29 +43,11 @@ lemma harmonic_singleton_ne_zero {c n : ℕ} (hn : 2 ≤ n) : (finset_range_sdiff_singleton_nonempty hn) norm_cast; simp only [add_pos_iff, or_true] -/-- If a rational is not a 2-adic integer, it is not an integer. -/ -theorem not_int_of_not_padic_int {a : ℚ} (H : 1 < padicNorm 2 a) : - ¬ a.isInt := by - suffices : a.den ≠ 1; simpa [Rat.isInt] - by_cases h : a = 0 - · subst h; contradiction - · unfold padicNorm at H - split_ifs at H; contradiction - suffices padicValRat 2 a < 0 by - unfold padicValRat at this - intro Hden - rw [Hden] at this; simp only [padicValNat.one, sub_zero] at this - norm_cast - refine' neg_of_neg_pos _ - have hx : (1 : ℚ) < 2 := by norm_num - rw [← zpow_lt_iff_lt hx] - simpa only [zpow_zero] - -lemma padicValRat_two_pow_div (r : ℕ) : padicValRat 2 (1 / 2^r) = -r := by +lemma padicValRat_two_pow_div (r : ℕ) : padicValRat 2 (1 / 2 ^ r) = -r := by rw [← padicValRat.self_pow_div (p := 2) (r := r)] simp only [one_div, Nat.cast_ofNat] -/-- For `i` less than `n`, `2^Nat.log 2 n` does not divide `i` unless `i` is equal to it. -/ +/-- For `i` less than `n`, `2 ^ Nat.log 2 n` does not divide `i` unless `i` is equal to it. -/ lemma nat_log_not_dvd {n i : ℕ} (Hi₁ : 0 < i) (Hi₂ : i ≠ 2 ^ Nat.log 2 n) (Hi₃ : i ≤ n) : ¬ 2^(Nat.log 2 n) ∣ i := by rintro ⟨c,Hc⟩ @@ -77,7 +59,7 @@ lemma nat_log_not_dvd {n i : ℕ} (Hi₁ : 0 < i) (Hi₂ : i ≠ 2 ^ Nat.log 2 n rw [mul_zero, zero_add] at Hk rcases Hk with rfl | rfl · linarith - · exact Hi₂ (Hc ▸ mul_one (2^Nat.log 2 n)) + · exact Hi₂ (Hc ▸ mul_one (2 ^ Nat.log 2 n)) rcases Hk with rfl | rfl · rw [← mul_assoc, Hpow] at Hc suffices Hnk : n < i by linarith @@ -88,10 +70,11 @@ lemma nat_log_not_dvd {n i : ℕ} (Hi₁ : 0 < i) (Hi₂ : i ≠ 2 ^ Nat.log 2 n suffices Hnk : n < i by linarith calc n < 2 ^ (Nat.log 2 n + 1) := Hle _ ≤ 2 ^ (Nat.log 2 n + 1)*k := Nat.le_mul_of_pos_right (Nat.pos_of_ne_zero h) - _ < 2 ^ (Nat.log 2 n + 1)*k + 2^ Nat.log 2 n := by linarith + _ < 2 ^ (Nat.log 2 n + 1)*k + 2 ^ Nat.log 2 n := by linarith _ = i := Hc.symm -lemma padicValRat_ge_neg_nat_log_ne {n q : ℕ}(Hq₁ : 0 < q)(Hq₂ : q ≤ n)(Hq₃ : q ≠ 2^Nat.log 2 n) : +lemma padicValRat_ge_neg_nat_log_ne {n q : ℕ}(Hq₁ : 0 < q) + (Hq₂ : q ≤ n)(Hq₃ : q ≠ 2 ^ Nat.log 2 n) : padicValRat 2 (1 / q) ≠ -Nat.log 2 n := by rw [one_div, padicValRat.inv, padicValRat.of_nat, ne_eq, neg_inj, Nat.cast_inj] exact (fun Hnot => nat_log_not_dvd (n := n) Hq₁ Hq₃ Hq₂ @@ -102,15 +85,16 @@ lemma padicValRat_ge_neg_Nat_log_ge {p n q : ℕ}[hp : Fact (Nat.Prime p)](Hq : rw [one_div, padicValRat.inv, padicValRat.of_nat, neg_le_neg_iff, Nat.cast_le] exact le_nat_log_of_le Hq -lemma padicValRat_ge_neg_Nat_log_lt (n q : ℕ) (Hq₁ : 0 < q) (Hq₂ : q ≤ n)(Hq₃ : q ≠ 2^Nat.log 2 n) : - padicValRat 2 (1 / 2^Nat.log 2 n) < padicValRat 2 (1 / q) := by +lemma padicValRat_ge_neg_Nat_log_lt (n q : ℕ) (Hq₁ : 0 < q) + (Hq₂ : q ≤ n) (Hq₃ : q ≠ 2 ^ Nat.log 2 n) : + padicValRat 2 (1 / 2 ^ Nat.log 2 n) < padicValRat 2 (1 / q) := by rw [padicValRat_two_pow_div] exact (lt_of_le_of_ne (padicValRat_ge_neg_Nat_log_ge (p := 2) Hq₂) (padicValRat_ge_neg_nat_log_ne Hq₁ Hq₂ Hq₃).symm) lemma two_pow_log_mem_range {n : ℕ} (hn : 2 ≤ n) : - 2^(Nat.log 2 n) - 1 ∈ Finset.range n := + 2 ^ (Nat.log 2 n) - 1 ∈ Finset.range n := Finset.mem_range.mpr <| lt_of_lt_of_le (Nat.sub_lt (pow_pos zero_lt_two _) zero_lt_one) (Nat.pow_log_le_self 2 (by linarith)) @@ -131,7 +115,7 @@ theorem harmonic_two_adicValRat {n : ℕ} (Hn : 2 ≤ n) : exact this (by linarith) · have := padicValRat.sum_gt_of_gt (p := 2) (j := 2 ^ Nat.log 2 n - 1) (finset_range_sdiff_singleton_nonempty Hn) (F := fun n => 1 / ((n + 1) : ℚ)) - (S := Finset.range n \ {2^Nat.log 2 n - 1}) + (S := Finset.range n \ {2 ^ Nat.log 2 n - 1}) simp only [Finset.mem_range, Finset.mem_sdiff, Finset.mem_singleton, Finset.singleton_subset_iff, pow_pos, Nat.cast_pred, Nat.cast_pow, sub_add_cancel, and_imp,gt_iff_lt] at this @@ -146,7 +130,7 @@ theorem harmonic_two_adicValRat {n : ℕ} (Hn : 2 ≤ n) : /-- The n-th harmonic number is not an integer for n ≥ 2. -/ theorem harmonic_not_int {n : ℕ} (Hn : 2 ≤ n) : ¬(harmonic n).isInt := by - refine' not_int_of_not_padic_int _ + refine' padicNorm.not_int_of_not_padic_int 2 _ dsimp [padicNorm] rw [if_neg (ne_of_gt <| harmonic_pos (by linarith))] refine' one_lt_zpow one_lt_two _ (lt_neg.mp _) diff --git a/Mathlib/NumberTheory/Padics/PadicNorm.lean b/Mathlib/NumberTheory/Padics/PadicNorm.lean index 262c19a892cf4..c5c9184efab9e 100644 --- a/Mathlib/NumberTheory/Padics/PadicNorm.lean +++ b/Mathlib/NumberTheory/Padics/PadicNorm.lean @@ -305,6 +305,24 @@ theorem nat_lt_one_iff (m : ℕ) : padicNorm p m < 1 ↔ p ∣ m := by rw [← Int.coe_nat_dvd, ← int_lt_one_iff, Int.cast_ofNat] #align padic_norm.nat_lt_one_iff padicNorm.nat_lt_one_iff +/-- If a rational is not a p-adic integer, it is not an integer. -/ +theorem not_int_of_not_padic_int (p : ℕ) {a : ℚ} [hp : Fact (Nat.Prime p)] + (H : 1 < padicNorm p a) : ¬ a.isInt := by + suffices : a.den ≠ 1; simpa [Rat.isInt] + by_cases h : a = 0 + · subst h; contradiction + · unfold padicNorm at H + split_ifs at H; contradiction + suffices padicValRat p a < 0 by + simp only [padicValRat] at this + intro Hden + rw [Hden] at this; simp only [padicValNat.one, sub_zero] at this + norm_cast + refine' neg_of_neg_pos _ + have hx : (1 : ℚ) < p := by norm_cast; exact Nat.Prime.one_lt hp.elim + rw [← zpow_lt_iff_lt hx] + simpa only [zpow_zero] + open BigOperators theorem sum_lt {α : Type*} {F : α → ℚ} {t : ℚ} {s : Finset α} : From e484f4f1aad5a0a9b994095699fb2c2a299adaba Mon Sep 17 00:00:00 2001 From: Koundinya Vajjha Date: Tue, 26 Sep 2023 08:25:32 -0700 Subject: [PATCH 40/74] Update Mathlib/NumberTheory/Padics/PadicNorm.lean Co-authored-by: Alex J Best --- Mathlib/NumberTheory/Padics/PadicNorm.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/NumberTheory/Padics/PadicNorm.lean b/Mathlib/NumberTheory/Padics/PadicNorm.lean index c5c9184efab9e..580df2839fda8 100644 --- a/Mathlib/NumberTheory/Padics/PadicNorm.lean +++ b/Mathlib/NumberTheory/Padics/PadicNorm.lean @@ -307,7 +307,7 @@ theorem nat_lt_one_iff (m : ℕ) : padicNorm p m < 1 ↔ p ∣ m := by /-- If a rational is not a p-adic integer, it is not an integer. -/ theorem not_int_of_not_padic_int (p : ℕ) {a : ℚ} [hp : Fact (Nat.Prime p)] - (H : 1 < padicNorm p a) : ¬ a.isInt := by + (H : 1 < padicNorm p a) : ¬ a.isInt := by suffices : a.den ≠ 1; simpa [Rat.isInt] by_cases h : a = 0 · subst h; contradiction From 8905db727ad67fa98dbf274157f81ce4f0bc3e3a Mon Sep 17 00:00:00 2001 From: Koundinya Vajjha Date: Tue, 26 Sep 2023 09:01:39 -0700 Subject: [PATCH 41/74] add padicValRat_def --- Mathlib/NumberTheory/Padics/PadicNorm.lean | 2 +- Mathlib/NumberTheory/Padics/PadicVal.lean | 4 ++++ 2 files changed, 5 insertions(+), 1 deletion(-) diff --git a/Mathlib/NumberTheory/Padics/PadicNorm.lean b/Mathlib/NumberTheory/Padics/PadicNorm.lean index 580df2839fda8..7a73e26975d47 100644 --- a/Mathlib/NumberTheory/Padics/PadicNorm.lean +++ b/Mathlib/NumberTheory/Padics/PadicNorm.lean @@ -314,7 +314,7 @@ theorem not_int_of_not_padic_int (p : ℕ) {a : ℚ} [hp : Fact (Nat.Prime p)] · unfold padicNorm at H split_ifs at H; contradiction suffices padicValRat p a < 0 by - simp only [padicValRat] at this + rw [padicValRat_def] at this intro Hden rw [Hden] at this; simp only [padicValNat.one, sub_zero] at this norm_cast diff --git a/Mathlib/NumberTheory/Padics/PadicVal.lean b/Mathlib/NumberTheory/Padics/PadicVal.lean index 9d00b885f0a30..ba301f12a995e 100644 --- a/Mathlib/NumberTheory/Padics/PadicVal.lean +++ b/Mathlib/NumberTheory/Padics/PadicVal.lean @@ -202,6 +202,10 @@ def padicValRat (p : ℕ) (q : ℚ) : ℤ := padicValInt p q.num - padicValNat p q.den #align padic_val_rat padicValRat +lemma padicValRat_def (p : ℕ) (q : ℚ) : + padicValRat p q = padicValInt p q.num - padicValNat p q.den + := rfl + namespace padicValRat open multiplicity From ec511de7762b99ff9e795683b26019d3b52c05c8 Mon Sep 17 00:00:00 2001 From: Koundinya Vajjha Date: Tue, 26 Sep 2023 09:07:35 -0700 Subject: [PATCH 42/74] appease linter --- Mathlib/NumberTheory/Padics/Harmonic.lean | 2 +- Mathlib/NumberTheory/Padics/PadicVal.lean | 4 ++-- 2 files changed, 3 insertions(+), 3 deletions(-) diff --git a/Mathlib/NumberTheory/Padics/Harmonic.lean b/Mathlib/NumberTheory/Padics/Harmonic.lean index 6bc2731043345..f7bd212202871 100644 --- a/Mathlib/NumberTheory/Padics/Harmonic.lean +++ b/Mathlib/NumberTheory/Padics/Harmonic.lean @@ -44,7 +44,7 @@ lemma harmonic_singleton_ne_zero {c n : ℕ} (hn : 2 ≤ n) : norm_cast; simp only [add_pos_iff, or_true] lemma padicValRat_two_pow_div (r : ℕ) : padicValRat 2 (1 / 2 ^ r) = -r := by - rw [← padicValRat.self_pow_div (p := 2) (r := r)] + rw [one_div, ← padicValRat.self_pow_div (p := 2) (r := r)] simp only [one_div, Nat.cast_ofNat] /-- For `i` less than `n`, `2 ^ Nat.log 2 n` does not divide `i` unless `i` is equal to it. -/ diff --git a/Mathlib/NumberTheory/Padics/PadicVal.lean b/Mathlib/NumberTheory/Padics/PadicVal.lean index ba301f12a995e..4792821322d0b 100644 --- a/Mathlib/NumberTheory/Padics/PadicVal.lean +++ b/Mathlib/NumberTheory/Padics/PadicVal.lean @@ -470,9 +470,9 @@ lemma add_lt_of_lt {q r₁ r₂ : ℚ} (hqr : r₁ + r₂ ≠ 0) @[simp] lemma self_pow_div (r : ℕ) : - padicValRat p (1 / (p : ℚ) ^ r) = -r := by + padicValRat p ((p : ℚ) ^ r)⁻¹ = -r := by norm_cast; simp only [Nat.cast_pow] - rw [one_div,padicValRat.inv, neg_inj, padicValRat.pow, + rw [padicValRat.inv, neg_inj, padicValRat.pow, padicValRat.self (Nat.Prime.one_lt hp.elim), mul_one] simp only [ne_eq, Nat.cast_eq_zero] exact Nat.Prime.ne_zero hp.elim From c7d4c9aefe4ff214ca29832e2cc3a2f5d476b496 Mon Sep 17 00:00:00 2001 From: Koundinya Vajjha Date: Tue, 26 Sep 2023 15:00:09 -0700 Subject: [PATCH 43/74] alex's comments --- Mathlib/NumberTheory/Padics/Harmonic.lean | 20 ++++++++++---------- Mathlib/NumberTheory/Padics/PadicVal.lean | 2 +- 2 files changed, 11 insertions(+), 11 deletions(-) diff --git a/Mathlib/NumberTheory/Padics/Harmonic.lean b/Mathlib/NumberTheory/Padics/Harmonic.lean index f7bd212202871..68a8e7a6259c0 100644 --- a/Mathlib/NumberTheory/Padics/Harmonic.lean +++ b/Mathlib/NumberTheory/Padics/Harmonic.lean @@ -73,14 +73,14 @@ lemma nat_log_not_dvd {n i : ℕ} (Hi₁ : 0 < i) (Hi₂ : i ≠ 2 ^ Nat.log 2 n _ < 2 ^ (Nat.log 2 n + 1)*k + 2 ^ Nat.log 2 n := by linarith _ = i := Hc.symm -lemma padicValRat_ge_neg_nat_log_ne {n q : ℕ}(Hq₁ : 0 < q) +lemma padicValRat_le_neg_nat_log_ne {n q : ℕ} (Hq₁ : 0 < q) (Hq₂ : q ≤ n)(Hq₃ : q ≠ 2 ^ Nat.log 2 n) : padicValRat 2 (1 / q) ≠ -Nat.log 2 n := by rw [one_div, padicValRat.inv, padicValRat.of_nat, ne_eq, neg_inj, Nat.cast_inj] exact (fun Hnot => nat_log_not_dvd (n := n) Hq₁ Hq₃ Hq₂ (Hnot ▸ pow_padicValNat_dvd (p := 2) (n := q))) -lemma padicValRat_ge_neg_Nat_log_ge {p n q : ℕ}[hp : Fact (Nat.Prime p)](Hq : q ≤ n) : +lemma padicValRat_le_neg_nat_log_le {p n q : ℕ} [hp : Fact (Nat.Prime p)] (Hq : q ≤ n) : -Nat.log p n ≤ padicValRat p (1 / q) := by rw [one_div, padicValRat.inv, padicValRat.of_nat, neg_le_neg_iff, Nat.cast_le] exact le_nat_log_of_le Hq @@ -90,30 +90,30 @@ lemma padicValRat_ge_neg_Nat_log_lt (n q : ℕ) (Hq₁ : 0 < q) padicValRat 2 (1 / 2 ^ Nat.log 2 n) < padicValRat 2 (1 / q) := by rw [padicValRat_two_pow_div] exact (lt_of_le_of_ne - (padicValRat_ge_neg_Nat_log_ge (p := 2) Hq₂) - (padicValRat_ge_neg_nat_log_ne Hq₁ Hq₂ Hq₃).symm) + (padicValRat_le_neg_nat_log_le (p := 2) Hq₂) + (padicValRat_le_neg_nat_log_ne Hq₁ Hq₂ Hq₃).symm) -lemma two_pow_log_mem_range {n : ℕ} (hn : 2 ≤ n) : +lemma two_pow_log_sub_one_mem_range {n : ℕ} (hn : 2 ≤ n) : 2 ^ (Nat.log 2 n) - 1 ∈ Finset.range n := Finset.mem_range.mpr <| lt_of_lt_of_le (Nat.sub_lt (pow_pos zero_lt_two _) zero_lt_one) (Nat.pow_log_le_self 2 (by linarith)) /-- The 2-adic valuation of the n-th harmonic number is the negative of the logarithm of n. -/ -theorem harmonic_two_adicValRat {n : ℕ} (Hn : 2 ≤ n) : +theorem padicValRat_two_harmonic {n : ℕ} (Hn : 2 ≤ n) : padicValRat 2 (harmonic n) = -Int.log 2 (n : ℚ) := by - rw [Int.log_natCast, harmonic_singleton (two_pow_log_mem_range Hn)] + rw [Int.log_natCast, harmonic_singleton (two_pow_log_sub_one_mem_range Hn)] simp only [gt_iff_lt, pow_pos, Nat.cast_pred, Nat.cast_pow, Nat.cast_ofNat, sub_add_cancel] rw [padicValRat.add_eq_of_lt (p := 2) _ (one_div_ne_zero <| pow_ne_zero _ two_ne_zero) (harmonic_singleton_ne_zero Hn) _] apply padicValRat_two_pow_div · have := harmonic_pos (n := n) - rw [harmonic_singleton (two_pow_log_mem_range Hn)] at this + rw [harmonic_singleton (two_pow_log_sub_one_mem_range Hn)] at this refine' ne_of_gt _ simp only [ne_eq, ge_iff_le, gt_iff_lt, pow_pos, Nat.cast_pred, Nat.cast_pow, Nat.cast_ofNat, sub_add_cancel, Finset.mem_range, not_lt, Finset.singleton_subset_iff] at this exact this (by linarith) - · have := padicValRat.sum_gt_of_gt (p := 2) (j := 2 ^ Nat.log 2 n - 1) + · have := padicValRat.lt_sum_of_lt (p := 2) (j := 2 ^ Nat.log 2 n - 1) (finset_range_sdiff_singleton_nonempty Hn) (F := fun n => 1 / ((n + 1) : ℚ)) (S := Finset.range n \ {2 ^ Nat.log 2 n - 1}) simp only [Finset.mem_range, Finset.mem_sdiff, Finset.mem_singleton, @@ -134,6 +134,6 @@ theorem harmonic_not_int {n : ℕ} (Hn : 2 ≤ n) : dsimp [padicNorm] rw [if_neg (ne_of_gt <| harmonic_pos (by linarith))] refine' one_lt_zpow one_lt_two _ (lt_neg.mp _) - rw [neg_zero, harmonic_two_adicValRat Hn,Int.log_natCast, Left.neg_neg_iff, + rw [neg_zero, padicValRat_two_harmonic Hn,Int.log_natCast, Left.neg_neg_iff, Nat.cast_pos, Nat.log_pos_iff] exact ⟨Hn,one_lt_two⟩ diff --git a/Mathlib/NumberTheory/Padics/PadicVal.lean b/Mathlib/NumberTheory/Padics/PadicVal.lean index 4792821322d0b..a448a30946154 100644 --- a/Mathlib/NumberTheory/Padics/PadicVal.lean +++ b/Mathlib/NumberTheory/Padics/PadicVal.lean @@ -496,7 +496,7 @@ theorem sum_pos_of_pos {n : ℕ} {F : ℕ → ℚ} (hF : ∀ i, i < n → 0 < pa /-- If the p-adic valuation of a finite set of positive rationals is greater than a given rational number, then the p-adic valuation of their sum is also greater than the same rational number. -/ -theorem sum_gt_of_gt {p j : ℕ} [hp : Fact (Nat.Prime p)] {F : ℕ → ℚ} {S : Finset ℕ} +theorem lt_sum_of_lt {p j : ℕ} [hp : Fact (Nat.Prime p)] {F : ℕ → ℚ} {S : Finset ℕ} (hS : S.Nonempty) (hF : ∀ i, i ∈ S → padicValRat p (F j) < padicValRat p (F i)) (hn1 : ∀ i : ℕ, 0 < F i) : padicValRat p (F j) < padicValRat p (∑ i in S, F i) := by induction' hS using Finset.Nonempty.cons_induction with k s S' Hnot Hne Hind From f9e5dd3502404a418967146d5e92438c68bbffee Mon Sep 17 00:00:00 2001 From: Koundinya Vajjha Date: Tue, 26 Sep 2023 15:54:47 -0700 Subject: [PATCH 44/74] simplify not_int_if_not_padic_int --- Mathlib/NumberTheory/Padics/Harmonic.lean | 6 +++--- 1 file changed, 3 insertions(+), 3 deletions(-) diff --git a/Mathlib/NumberTheory/Padics/Harmonic.lean b/Mathlib/NumberTheory/Padics/Harmonic.lean index 68a8e7a6259c0..9e3c7fd7ce354 100644 --- a/Mathlib/NumberTheory/Padics/Harmonic.lean +++ b/Mathlib/NumberTheory/Padics/Harmonic.lean @@ -33,9 +33,9 @@ lemma harmonic_singleton {n c : ℕ} (hc : c ∈ Finset.range n) : lemma finset_range_sdiff_singleton_nonempty {c n : ℕ} (hn : 2 ≤ n) : Finset.Nonempty (Finset.range n \ {c}) := by rw [Finset.sdiff_nonempty, Finset.subset_singleton_iff, Finset.range_eq_empty_iff, not_or] - refine' ⟨by linarith, fun Hnot => _⟩ - suffices n = 1 by linarith - rw [← Finset.card_range n, ← Finset.card_singleton c, Hnot] + refine' ⟨ne_of_gt <| lt_of_lt_of_le zero_lt_two hn, fun Hnot => _⟩ + have : n = 1 := by rw [← Finset.card_range n, ← Finset.card_singleton c, Hnot] + simp only [this] at hn lemma harmonic_singleton_ne_zero {c n : ℕ} (hn : 2 ≤ n) : ∑ x in Finset.range n \ {c}, 1 / (x + 1 : ℚ) ≠ 0 := by From ef7a81e000a603f9722268e00142fe03b260b488 Mon Sep 17 00:00:00 2001 From: Koundinya Vajjha Date: Tue, 26 Sep 2023 15:55:16 -0700 Subject: [PATCH 45/74] simplify not_int_of_not_padic_int --- Mathlib/NumberTheory/Padics/PadicNorm.lean | 20 ++++++-------------- 1 file changed, 6 insertions(+), 14 deletions(-) diff --git a/Mathlib/NumberTheory/Padics/PadicNorm.lean b/Mathlib/NumberTheory/Padics/PadicNorm.lean index 7a73e26975d47..4b2065f907c23 100644 --- a/Mathlib/NumberTheory/Padics/PadicNorm.lean +++ b/Mathlib/NumberTheory/Padics/PadicNorm.lean @@ -305,23 +305,15 @@ theorem nat_lt_one_iff (m : ℕ) : padicNorm p m < 1 ↔ p ∣ m := by rw [← Int.coe_nat_dvd, ← int_lt_one_iff, Int.cast_ofNat] #align padic_norm.nat_lt_one_iff padicNorm.nat_lt_one_iff +lemma _root_.Rat.isInt.eq_num {q : ℚ} (h : q.isInt) : q = q.num := by + rw [Rat.isInt, Nat.beq_eq_true_eq] at h + exact (Rat.coe_int_num_of_den_eq_one h).symm /-- If a rational is not a p-adic integer, it is not an integer. -/ theorem not_int_of_not_padic_int (p : ℕ) {a : ℚ} [hp : Fact (Nat.Prime p)] (H : 1 < padicNorm p a) : ¬ a.isInt := by - suffices : a.den ≠ 1; simpa [Rat.isInt] - by_cases h : a = 0 - · subst h; contradiction - · unfold padicNorm at H - split_ifs at H; contradiction - suffices padicValRat p a < 0 by - rw [padicValRat_def] at this - intro Hden - rw [Hden] at this; simp only [padicValNat.one, sub_zero] at this - norm_cast - refine' neg_of_neg_pos _ - have hx : (1 : ℚ) < p := by norm_cast; exact Nat.Prime.one_lt hp.elim - rw [← zpow_lt_iff_lt hx] - simpa only [zpow_zero] + contrapose! H + rw [Rat.isInt.eq_num H] + apply padicNorm.of_int open BigOperators From bf025529bb0187390ee2fbba0326f3e01753871d Mon Sep 17 00:00:00 2001 From: Koundinya Vajjha Date: Tue, 26 Sep 2023 15:56:14 -0700 Subject: [PATCH 46/74] Update Mathlib/NumberTheory/Padics/PadicVal.lean Co-authored-by: Thomas Browning --- Mathlib/NumberTheory/Padics/PadicVal.lean | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/Mathlib/NumberTheory/Padics/PadicVal.lean b/Mathlib/NumberTheory/Padics/PadicVal.lean index a448a30946154..f203c6a782bb4 100644 --- a/Mathlib/NumberTheory/Padics/PadicVal.lean +++ b/Mathlib/NumberTheory/Padics/PadicVal.lean @@ -203,8 +203,8 @@ def padicValRat (p : ℕ) (q : ℚ) : ℤ := #align padic_val_rat padicValRat lemma padicValRat_def (p : ℕ) (q : ℚ) : - padicValRat p q = padicValInt p q.num - padicValNat p q.den - := rfl + padicValRat p q = padicValInt p q.num - padicValNat p q.den := + rfl namespace padicValRat From edd54dc7ad23cbea8818718e7561b16c5bc70c70 Mon Sep 17 00:00:00 2001 From: Koundinya Vajjha Date: Tue, 26 Sep 2023 15:57:45 -0700 Subject: [PATCH 47/74] Update Mathlib/NumberTheory/Padics/PadicVal.lean Co-authored-by: Thomas Browning --- Mathlib/NumberTheory/Padics/PadicVal.lean | 30 +++++++++-------------- 1 file changed, 12 insertions(+), 18 deletions(-) diff --git a/Mathlib/NumberTheory/Padics/PadicVal.lean b/Mathlib/NumberTheory/Padics/PadicVal.lean index f203c6a782bb4..1154f2a2550e5 100644 --- a/Mathlib/NumberTheory/Padics/PadicVal.lean +++ b/Mathlib/NumberTheory/Padics/PadicVal.lean @@ -439,24 +439,18 @@ theorem min_le_padicValRat_add {q r : ℚ} (hqr : q + r ≠ 0) : lemma min_eq_padicValRat_add {q r : ℚ} (hqr : q + r ≠ 0) (hq : q ≠ 0) (hr : r ≠ 0) (hval : padicValRat p q ≠ padicValRat p r) : padicValRat p (q + r) = min (padicValRat p q) (padicValRat p r) := by - have Hmin := padicValRat.min_le_padicValRat_add (p := p) (hp := hp) hqr - wlog h : padicValRat p q < padicValRat p r generalizing q r with Hgen - · push_neg at h; rw [add_comm, min_comm] - exact Hgen (by rwa [add_comm]) hr hq hval.symm - (by rwa [min_comm, add_comm]) - (hval.symm.lt_of_le h) - · rw [min_eq_left (le_of_lt h)] at Hmin ⊢ - suffices Hreq : padicValRat p (q + r) ≤ padicValRat p q by linarith - suffices Haux : min (padicValRat p (q + r)) (padicValRat p r) ≤ padicValRat p q by - rw [min_def] at Haux - split_ifs at Haux with Hspl - · assumption - · linarith - calc padicValRat p q = padicValRat p ((q + r) - r) := by congr; simp - _ ≥ min (padicValRat p (q + r)) (padicValRat p (-r)) := - ge_iff_le.mp <| le_trans (padicValRat.min_le_padicValRat_add (q := q+r) (r := -r) (by simpa)) - (by rw [add_neg_cancel_right, add_sub_cancel]) - _ = min (padicValRat p (q + r)) (padicValRat p r) := by rw [padicValRat.neg] + have h1 := min_le_padicValRat_add (p := p) hqr + have h2 := min_le_padicValRat_add (p := p) (ne_of_eq_of_ne (add_neg_cancel_right q r) hq) + have h3 := min_le_padicValRat_add (p := p) (ne_of_eq_of_ne (add_neg_cancel_right r q) hr) + rw [add_neg_cancel_right, padicValRat.neg] at h2 h3 + rw [add_comm] at h3 + refine' le_antisymm (le_min _ _) h1 + · contrapose! h2 + rw [min_eq_right h2.le] at h3 + exact lt_min h2 (lt_of_le_of_ne h3 hval) + · contrapose! h3 + rw [min_eq_right h3.le] at h2 + exact lt_min h3 (lt_of_le_of_ne h2 hval.symm) lemma add_eq_of_lt {q r : ℚ} (hqr : q + r ≠ 0) (hq : q ≠ 0) (hr : r ≠ 0) (hval : padicValRat p q < padicValRat p r) : From db51df416706d45dd7d9add7726f8f084252532e Mon Sep 17 00:00:00 2001 From: Koundinya Vajjha Date: Tue, 26 Sep 2023 15:59:38 -0700 Subject: [PATCH 48/74] add_eq_min --- Mathlib/NumberTheory/Padics/PadicVal.lean | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/Mathlib/NumberTheory/Padics/PadicVal.lean b/Mathlib/NumberTheory/Padics/PadicVal.lean index 1154f2a2550e5..3a79aa2b9238c 100644 --- a/Mathlib/NumberTheory/Padics/PadicVal.lean +++ b/Mathlib/NumberTheory/Padics/PadicVal.lean @@ -436,7 +436,7 @@ theorem min_le_padicValRat_add {q r : ℚ} (hqr : q + r ≠ 0) : #align padic_val_rat.min_le_padic_val_rat_add padicValRat.min_le_padicValRat_add /-- Ultrametric property of a p-adic valuation. -/ -lemma min_eq_padicValRat_add {q r : ℚ} (hqr : q + r ≠ 0) (hq : q ≠ 0) (hr : r ≠ 0) +lemma add_eq_min {q r : ℚ} (hqr : q + r ≠ 0) (hq : q ≠ 0) (hr : r ≠ 0) (hval : padicValRat p q ≠ padicValRat p r) : padicValRat p (q + r) = min (padicValRat p q) (padicValRat p r) := by have h1 := min_le_padicValRat_add (p := p) hqr @@ -455,7 +455,7 @@ lemma min_eq_padicValRat_add {q r : ℚ} (hqr : q + r ≠ 0) (hq : q ≠ 0) (hr lemma add_eq_of_lt {q r : ℚ} (hqr : q + r ≠ 0) (hq : q ≠ 0) (hr : r ≠ 0) (hval : padicValRat p q < padicValRat p r) : padicValRat p (q + r) = padicValRat p q := by - rw [min_eq_padicValRat_add hqr hq hr (ne_of_lt hval),min_eq_left (le_of_lt hval)] + rw [add_eq_min hqr hq hr (ne_of_lt hval),min_eq_left (le_of_lt hval)] lemma add_lt_of_lt {q r₁ r₂ : ℚ} (hqr : r₁ + r₂ ≠ 0) (hval₁ : padicValRat p q < padicValRat p r₁) (hval₂ : padicValRat p q < padicValRat p r₂) : From bc58ab8a4b929d504043cbd49cba0f0ed6e710a2 Mon Sep 17 00:00:00 2001 From: Koundinya Vajjha Date: Tue, 26 Sep 2023 16:00:36 -0700 Subject: [PATCH 49/74] add_lt_of_lt -> lt_add_of_lt --- Mathlib/NumberTheory/Padics/PadicVal.lean | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/Mathlib/NumberTheory/Padics/PadicVal.lean b/Mathlib/NumberTheory/Padics/PadicVal.lean index 3a79aa2b9238c..71f240deadb30 100644 --- a/Mathlib/NumberTheory/Padics/PadicVal.lean +++ b/Mathlib/NumberTheory/Padics/PadicVal.lean @@ -457,7 +457,7 @@ lemma add_eq_of_lt {q r : ℚ} (hqr : q + r ≠ 0) padicValRat p (q + r) = padicValRat p q := by rw [add_eq_min hqr hq hr (ne_of_lt hval),min_eq_left (le_of_lt hval)] -lemma add_lt_of_lt {q r₁ r₂ : ℚ} (hqr : r₁ + r₂ ≠ 0) +lemma lt_add_of_lt {q r₁ r₂ : ℚ} (hqr : r₁ + r₂ ≠ 0) (hval₁ : padicValRat p q < padicValRat p r₁) (hval₂ : padicValRat p q < padicValRat p r₂) : padicValRat p q < padicValRat p (r₁ + r₂) := lt_of_lt_of_le (lt_min hval₁ hval₂) (padicValRat.min_le_padicValRat_add hqr) @@ -497,7 +497,7 @@ theorem lt_sum_of_lt {p j : ℕ} [hp : Fact (Nat.Prime p)] {F : ℕ → ℚ} {S · rw [Finset.sum_singleton] exact hF k (by simp) · rw [Finset.cons_eq_insert, Finset.sum_insert Hnot] - exact padicValRat.add_lt_of_lt + exact padicValRat.lt_add_of_lt (ne_of_gt (add_pos (hn1 s) (Finset.sum_pos (fun i _ => hn1 i) Hne))) (hF _ (by simp [Finset.mem_insert, true_or])) (Hind (fun i hi => hF _ (by rw [Finset.cons_eq_insert,Finset.mem_insert]; exact Or.inr hi))) From 51f2acc81223e1b8d91d27c393f9e0c58fc7b951 Mon Sep 17 00:00:00 2001 From: Koundinya Vajjha Date: Tue, 26 Sep 2023 16:01:38 -0700 Subject: [PATCH 50/74] simplify self_pow_div --- Mathlib/NumberTheory/Padics/PadicVal.lean | 10 +++------- 1 file changed, 3 insertions(+), 7 deletions(-) diff --git a/Mathlib/NumberTheory/Padics/PadicVal.lean b/Mathlib/NumberTheory/Padics/PadicVal.lean index 71f240deadb30..625778a9a1f9d 100644 --- a/Mathlib/NumberTheory/Padics/PadicVal.lean +++ b/Mathlib/NumberTheory/Padics/PadicVal.lean @@ -463,13 +463,9 @@ lemma lt_add_of_lt {q r₁ r₂ : ℚ} (hqr : r₁ + r₂ ≠ 0) lt_of_lt_of_le (lt_min hval₁ hval₂) (padicValRat.min_le_padicValRat_add hqr) @[simp] -lemma self_pow_div (r : ℕ) : - padicValRat p ((p : ℚ) ^ r)⁻¹ = -r := by - norm_cast; simp only [Nat.cast_pow] - rw [padicValRat.inv, neg_inj, padicValRat.pow, - padicValRat.self (Nat.Prime.one_lt hp.elim), mul_one] - simp only [ne_eq, Nat.cast_eq_zero] - exact Nat.Prime.ne_zero hp.elim +lemma self_pow_div (r : ℕ) : padicValRat p ((p : ℚ) ^ r)⁻¹ = -r := by + rw [padicValRat.inv, neg_inj, padicValRat.pow (Nat.cast_ne_zero.mpr hp.elim.ne_zero), + padicValRat.self hp.elim.one_lt, mul_one] open BigOperators From 39c6fb000a3a57e9fe4d0a037bfa1ec9314a2d4d Mon Sep 17 00:00:00 2001 From: Koundinya Vajjha Date: Tue, 26 Sep 2023 16:03:33 -0700 Subject: [PATCH 51/74] simplify harmonic not int --- Mathlib/NumberTheory/Padics/Harmonic.lean | 3 +-- 1 file changed, 1 insertion(+), 2 deletions(-) diff --git a/Mathlib/NumberTheory/Padics/Harmonic.lean b/Mathlib/NumberTheory/Padics/Harmonic.lean index 9e3c7fd7ce354..3c472b9498785 100644 --- a/Mathlib/NumberTheory/Padics/Harmonic.lean +++ b/Mathlib/NumberTheory/Padics/Harmonic.lean @@ -131,8 +131,7 @@ theorem padicValRat_two_harmonic {n : ℕ} (Hn : 2 ≤ n) : theorem harmonic_not_int {n : ℕ} (Hn : 2 ≤ n) : ¬(harmonic n).isInt := by refine' padicNorm.not_int_of_not_padic_int 2 _ - dsimp [padicNorm] - rw [if_neg (ne_of_gt <| harmonic_pos (by linarith))] + rw [padicNorm.eq_zpow_of_nonzero (harmonic_pos (by linarith)).ne'] refine' one_lt_zpow one_lt_two _ (lt_neg.mp _) rw [neg_zero, padicValRat_two_harmonic Hn,Int.log_natCast, Left.neg_neg_iff, Nat.cast_pos, Nat.log_pos_iff] From 797e82f894573cec661c4c97c07ad921c018ef05 Mon Sep 17 00:00:00 2001 From: Koundinya Vajjha Date: Tue, 26 Sep 2023 16:08:41 -0700 Subject: [PATCH 52/74] Update Mathlib/NumberTheory/Padics/PadicVal.lean Co-authored-by: Thomas Browning --- Mathlib/NumberTheory/Padics/PadicVal.lean | 9 ++++----- 1 file changed, 4 insertions(+), 5 deletions(-) diff --git a/Mathlib/NumberTheory/Padics/PadicVal.lean b/Mathlib/NumberTheory/Padics/PadicVal.lean index 625778a9a1f9d..bf237f3739a69 100644 --- a/Mathlib/NumberTheory/Padics/PadicVal.lean +++ b/Mathlib/NumberTheory/Padics/PadicVal.lean @@ -587,11 +587,10 @@ theorem padicValNat_primes {q : ℕ} [hp : Fact p.Prime] [hq : Fact q.Prime] (ne /-- The p-adic valuation of `n` is less than or equal to its logarithm w.r.t `p`.-/ lemma le_nat_log (n : ℕ) [hp : Fact (Nat.Prime p)] : padicValNat p n ≤ Nat.log p n := by - by_cases (n = 0) - · simp only [h, padicValNat.zero, Nat.log_zero_right, le_refl] - · refine' Nat.le_log_of_pow_le (Nat.Prime.one_lt hp.elim) _ - by_contra Hnot; push_neg at Hnot - exact h (Nat.eq_zero_of_dvd_of_lt (pow_padicValNat_dvd (p := p) (n := n)) Hnot) + rcases eq_or_ne n 0 with rfl | h + · rw [padicValNat.zero, Nat.log_zero_right] + · refine' Nat.le_log_of_pow_le hp.elim.one_lt _ + exact le_of_dvd (Nat.pos_of_ne_zero h) pow_padicValNat_dvd lemma le_nat_log_of_le {n₁ n₂ : ℕ} [Fact (Nat.Prime p)] (hn : n₁ ≤ n₂) : padicValNat p n₁ ≤ Nat.log p n₂ := le_trans (le_nat_log n₁) (Nat.log_mono_right hn) From 6338778d53fbc18e738d3aa359bcfadad45a506d Mon Sep 17 00:00:00 2001 From: Koundinya Vajjha Date: Tue, 26 Sep 2023 17:53:31 -0700 Subject: [PATCH 53/74] move finset lemma to basic file --- Mathlib/Data/Finset/Basic.lean | 12 ++++++++++++ Mathlib/NumberTheory/Padics/Harmonic.lean | 9 +-------- 2 files changed, 13 insertions(+), 8 deletions(-) diff --git a/Mathlib/Data/Finset/Basic.lean b/Mathlib/Data/Finset/Basic.lean index 9a18d146db9c6..cd25ee709daeb 100644 --- a/Mathlib/Data/Finset/Basic.lean +++ b/Mathlib/Data/Finset/Basic.lean @@ -2325,6 +2325,12 @@ theorem sdiff_singleton_eq_self (ha : a ∉ s) : s \ {a} = s := sdiff_eq_self_iff_disjoint.2 <| by simp [ha] #align finset.sdiff_singleton_eq_self Finset.sdiff_singleton_eq_self +theorem sdiff_singleton_nonempty_of_nontrivial {c : α} {s : Finset α} (hS : s.Nontrivial) : + (s \ {c}).Nonempty := by + rw [Finset.sdiff_nonempty, Finset.subset_singleton_iff] + push_neg + exact ⟨by rintro rfl; exact Finset.not_nontrivial_empty hS, Finset.Nontrivial.ne_singleton hS⟩ + theorem sdiff_sdiff_left' (s t u : Finset α) : (s \ t) \ u = s \ t ∩ (s \ u) := _root_.sdiff_sdiff_left' #align finset.sdiff_sdiff_left' Finset.sdiff_sdiff_left' @@ -3103,6 +3109,12 @@ theorem range_filter_eq {n m : ℕ} : (range n).filter (· = m) = if m < n then · simp #align finset.range_filter_eq Finset.range_filter_eq +lemma range_nontrivial {n : ℕ} (hn : 2 ≤ n) : (Finset.range n).Nontrivial := by + refine' Set.nontrivial_of_exists_lt ⟨0,_,1,_,zero_lt_one⟩ <;> + simp only [Finset.coe_range, Set.mem_Iio] + exact lt_of_lt_of_le zero_lt_two hn + exact lt_of_lt_of_le one_lt_two hn + end Range -- useful rules for calculations with quantifiers diff --git a/Mathlib/NumberTheory/Padics/Harmonic.lean b/Mathlib/NumberTheory/Padics/Harmonic.lean index 3c472b9498785..8144cef67c4d9 100644 --- a/Mathlib/NumberTheory/Padics/Harmonic.lean +++ b/Mathlib/NumberTheory/Padics/Harmonic.lean @@ -30,17 +30,10 @@ lemma harmonic_singleton {n c : ℕ} (hc : c ∈ Finset.range n) : dsimp [harmonic] rwa [add_comm, Finset.sum_eq_sum_diff_singleton_add (i := c)] -lemma finset_range_sdiff_singleton_nonempty {c n : ℕ} (hn : 2 ≤ n) : - Finset.Nonempty (Finset.range n \ {c}) := by - rw [Finset.sdiff_nonempty, Finset.subset_singleton_iff, Finset.range_eq_empty_iff, not_or] - refine' ⟨ne_of_gt <| lt_of_lt_of_le zero_lt_two hn, fun Hnot => _⟩ - have : n = 1 := by rw [← Finset.card_range n, ← Finset.card_singleton c, Hnot] - simp only [this] at hn - lemma harmonic_singleton_ne_zero {c n : ℕ} (hn : 2 ≤ n) : ∑ x in Finset.range n \ {c}, 1 / (x + 1 : ℚ) ≠ 0 := by refine' ne_of_gt <| Finset.sum_pos (fun i _ => div_pos zero_lt_one _) - (finset_range_sdiff_singleton_nonempty hn) + (Finset.sdiff_singleton_nonempty_of_nontrivial <| Finset.range_nontrivial hn) norm_cast; simp only [add_pos_iff, or_true] lemma padicValRat_two_pow_div (r : ℕ) : padicValRat 2 (1 / 2 ^ r) = -r := by From 5c63894a9740e9af590c7f71bdef513efafab7c5 Mon Sep 17 00:00:00 2001 From: Koundinya Vajjha Date: Tue, 26 Sep 2023 18:26:52 -0700 Subject: [PATCH 54/74] fix proof --- Mathlib/NumberTheory/Padics/Harmonic.lean | 3 ++- 1 file changed, 2 insertions(+), 1 deletion(-) diff --git a/Mathlib/NumberTheory/Padics/Harmonic.lean b/Mathlib/NumberTheory/Padics/Harmonic.lean index 8144cef67c4d9..e3d377a501d87 100644 --- a/Mathlib/NumberTheory/Padics/Harmonic.lean +++ b/Mathlib/NumberTheory/Padics/Harmonic.lean @@ -107,7 +107,8 @@ theorem padicValRat_two_harmonic {n : ℕ} (Hn : 2 ≤ n) : sub_add_cancel, Finset.mem_range, not_lt, Finset.singleton_subset_iff] at this exact this (by linarith) · have := padicValRat.lt_sum_of_lt (p := 2) (j := 2 ^ Nat.log 2 n - 1) - (finset_range_sdiff_singleton_nonempty Hn) (F := fun n => 1 / ((n + 1) : ℚ)) + (Finset.sdiff_singleton_nonempty_of_nontrivial <| Finset.range_nontrivial Hn) + (F := fun n => 1 / ((n + 1) : ℚ)) (S := Finset.range n \ {2 ^ Nat.log 2 n - 1}) simp only [Finset.mem_range, Finset.mem_sdiff, Finset.mem_singleton, Finset.singleton_subset_iff, pow_pos, Nat.cast_pred, Nat.cast_pow, From a1cce8d9fb39b1f50e8bd00953db4bdd5745bf04 Mon Sep 17 00:00:00 2001 From: Koundinya Vajjha Date: Tue, 26 Sep 2023 22:10:20 -0700 Subject: [PATCH 55/74] rename to allow dot notation --- Mathlib/Data/Finset/Basic.lean | 2 +- Mathlib/NumberTheory/Padics/Harmonic.lean | 4 ++-- 2 files changed, 3 insertions(+), 3 deletions(-) diff --git a/Mathlib/Data/Finset/Basic.lean b/Mathlib/Data/Finset/Basic.lean index cd25ee709daeb..e477814ac71fc 100644 --- a/Mathlib/Data/Finset/Basic.lean +++ b/Mathlib/Data/Finset/Basic.lean @@ -2325,7 +2325,7 @@ theorem sdiff_singleton_eq_self (ha : a ∉ s) : s \ {a} = s := sdiff_eq_self_iff_disjoint.2 <| by simp [ha] #align finset.sdiff_singleton_eq_self Finset.sdiff_singleton_eq_self -theorem sdiff_singleton_nonempty_of_nontrivial {c : α} {s : Finset α} (hS : s.Nontrivial) : +theorem Nontrivial.sdiff_singleton_nonempty {c : α} {s : Finset α} (hS : s.Nontrivial) : (s \ {c}).Nonempty := by rw [Finset.sdiff_nonempty, Finset.subset_singleton_iff] push_neg diff --git a/Mathlib/NumberTheory/Padics/Harmonic.lean b/Mathlib/NumberTheory/Padics/Harmonic.lean index e3d377a501d87..60578e1c11838 100644 --- a/Mathlib/NumberTheory/Padics/Harmonic.lean +++ b/Mathlib/NumberTheory/Padics/Harmonic.lean @@ -33,7 +33,7 @@ lemma harmonic_singleton {n c : ℕ} (hc : c ∈ Finset.range n) : lemma harmonic_singleton_ne_zero {c n : ℕ} (hn : 2 ≤ n) : ∑ x in Finset.range n \ {c}, 1 / (x + 1 : ℚ) ≠ 0 := by refine' ne_of_gt <| Finset.sum_pos (fun i _ => div_pos zero_lt_one _) - (Finset.sdiff_singleton_nonempty_of_nontrivial <| Finset.range_nontrivial hn) + (Finset.Nontrivial.sdiff_singleton_nonempty <| Finset.range_nontrivial hn) norm_cast; simp only [add_pos_iff, or_true] lemma padicValRat_two_pow_div (r : ℕ) : padicValRat 2 (1 / 2 ^ r) = -r := by @@ -107,7 +107,7 @@ theorem padicValRat_two_harmonic {n : ℕ} (Hn : 2 ≤ n) : sub_add_cancel, Finset.mem_range, not_lt, Finset.singleton_subset_iff] at this exact this (by linarith) · have := padicValRat.lt_sum_of_lt (p := 2) (j := 2 ^ Nat.log 2 n - 1) - (Finset.sdiff_singleton_nonempty_of_nontrivial <| Finset.range_nontrivial Hn) + (Finset.Nontrivial.sdiff_singleton_nonempty <| Finset.range_nontrivial Hn) (F := fun n => 1 / ((n + 1) : ℚ)) (S := Finset.range n \ {2 ^ Nat.log 2 n - 1}) simp only [Finset.mem_range, Finset.mem_sdiff, Finset.mem_singleton, From 8741aa3016ba612088924ef340a01a49060e6de6 Mon Sep 17 00:00:00 2001 From: Koundinya Vajjha Date: Tue, 26 Sep 2023 22:15:30 -0700 Subject: [PATCH 56/74] =?UTF-8?q?=E2=81=BB=C2=B9=20instead=20of=20/?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit --- Mathlib/NumberTheory/Padics/Harmonic.lean | 6 ++++-- 1 file changed, 4 insertions(+), 2 deletions(-) diff --git a/Mathlib/NumberTheory/Padics/Harmonic.lean b/Mathlib/NumberTheory/Padics/Harmonic.lean index 60578e1c11838..8e2ed50cbb899 100644 --- a/Mathlib/NumberTheory/Padics/Harmonic.lean +++ b/Mathlib/NumberTheory/Padics/Harmonic.lean @@ -19,15 +19,17 @@ https://kconrad.math.uconn.edu/blurbs/gradnumthy/padicharmonicsum.pdf open BigOperators /-- The nth-harmonic number defined as a finset sum of consecutive reciprocals. -/ -def harmonic : ℕ → ℚ := fun n => ∑ i in Finset.range n, 1 / (i + 1) +def harmonic : ℕ → ℚ := fun n => ∑ i in Finset.range n, (↑(i + 1))⁻¹ lemma harmonic_pos {n : ℕ} (Hn : n ≠ 0) : 0 < harmonic n := - Finset.sum_pos (fun _ _ => div_pos zero_lt_one (by norm_cast; linarith)) + Finset.sum_pos (fun _ _ => inv_pos.mpr (by norm_cast; linarith)) (by rwa [Finset.nonempty_range_iff]) lemma harmonic_singleton {n c : ℕ} (hc : c ∈ Finset.range n) : harmonic n = 1 / (c + 1) + ∑ x in Finset.range n \ {c}, 1 / ((x : ℚ) + 1) := by dsimp [harmonic] + simp only [Nat.cast_add, Nat.cast_one, one_div, Finset.mem_range, not_lt, + Finset.singleton_subset_iff] rwa [add_comm, Finset.sum_eq_sum_diff_singleton_add (i := c)] lemma harmonic_singleton_ne_zero {c n : ℕ} (hn : 2 ≤ n) : From 2b3ac1a2799158a20f98c34801e6e6291aafe3fb Mon Sep 17 00:00:00 2001 From: Koundinya Vajjha Date: Wed, 27 Sep 2023 15:12:05 -0700 Subject: [PATCH 57/74] use dot notation --- Mathlib/NumberTheory/Padics/Harmonic.lean | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/Mathlib/NumberTheory/Padics/Harmonic.lean b/Mathlib/NumberTheory/Padics/Harmonic.lean index 8e2ed50cbb899..2f56735771a6e 100644 --- a/Mathlib/NumberTheory/Padics/Harmonic.lean +++ b/Mathlib/NumberTheory/Padics/Harmonic.lean @@ -35,7 +35,7 @@ lemma harmonic_singleton {n c : ℕ} (hc : c ∈ Finset.range n) : lemma harmonic_singleton_ne_zero {c n : ℕ} (hn : 2 ≤ n) : ∑ x in Finset.range n \ {c}, 1 / (x + 1 : ℚ) ≠ 0 := by refine' ne_of_gt <| Finset.sum_pos (fun i _ => div_pos zero_lt_one _) - (Finset.Nontrivial.sdiff_singleton_nonempty <| Finset.range_nontrivial hn) + (Finset.range_nontrivial hn).sdiff_singleton_nonempty norm_cast; simp only [add_pos_iff, or_true] lemma padicValRat_two_pow_div (r : ℕ) : padicValRat 2 (1 / 2 ^ r) = -r := by @@ -109,7 +109,7 @@ theorem padicValRat_two_harmonic {n : ℕ} (Hn : 2 ≤ n) : sub_add_cancel, Finset.mem_range, not_lt, Finset.singleton_subset_iff] at this exact this (by linarith) · have := padicValRat.lt_sum_of_lt (p := 2) (j := 2 ^ Nat.log 2 n - 1) - (Finset.Nontrivial.sdiff_singleton_nonempty <| Finset.range_nontrivial Hn) + (Finset.range_nontrivial Hn).sdiff_singleton_nonempty (F := fun n => 1 / ((n + 1) : ℚ)) (S := Finset.range n \ {2 ^ Nat.log 2 n - 1}) simp only [Finset.mem_range, Finset.mem_sdiff, Finset.mem_singleton, From e249282cdb1d4fed4aa60e87ce7e9b9c65571836 Mon Sep 17 00:00:00 2001 From: Koundinya Vajjha Date: Wed, 27 Sep 2023 15:13:00 -0700 Subject: [PATCH 58/74] Nat.log instead of Int.log --- Mathlib/NumberTheory/Padics/Harmonic.lean | 6 +++--- 1 file changed, 3 insertions(+), 3 deletions(-) diff --git a/Mathlib/NumberTheory/Padics/Harmonic.lean b/Mathlib/NumberTheory/Padics/Harmonic.lean index 2f56735771a6e..af97fa8f3e773 100644 --- a/Mathlib/NumberTheory/Padics/Harmonic.lean +++ b/Mathlib/NumberTheory/Padics/Harmonic.lean @@ -96,8 +96,8 @@ lemma two_pow_log_sub_one_mem_range {n : ℕ} (hn : 2 ≤ n) : /-- The 2-adic valuation of the n-th harmonic number is the negative of the logarithm of n. -/ theorem padicValRat_two_harmonic {n : ℕ} (Hn : 2 ≤ n) : - padicValRat 2 (harmonic n) = -Int.log 2 (n : ℚ) := by - rw [Int.log_natCast, harmonic_singleton (two_pow_log_sub_one_mem_range Hn)] + padicValRat 2 (harmonic n) = -Nat.log 2 n := by + rw [harmonic_singleton (two_pow_log_sub_one_mem_range Hn)] simp only [gt_iff_lt, pow_pos, Nat.cast_pred, Nat.cast_pow, Nat.cast_ofNat, sub_add_cancel] rw [padicValRat.add_eq_of_lt (p := 2) _ (one_div_ne_zero <| pow_ne_zero _ two_ne_zero) (harmonic_singleton_ne_zero Hn) _] @@ -129,6 +129,6 @@ theorem harmonic_not_int {n : ℕ} (Hn : 2 ≤ n) : refine' padicNorm.not_int_of_not_padic_int 2 _ rw [padicNorm.eq_zpow_of_nonzero (harmonic_pos (by linarith)).ne'] refine' one_lt_zpow one_lt_two _ (lt_neg.mp _) - rw [neg_zero, padicValRat_two_harmonic Hn,Int.log_natCast, Left.neg_neg_iff, + rw [neg_zero, padicValRat_two_harmonic Hn, Left.neg_neg_iff, Nat.cast_pos, Nat.log_pos_iff] exact ⟨Hn,one_lt_two⟩ From f61a5c5328ed24b07d57992cc0374d9e88a73625 Mon Sep 17 00:00:00 2001 From: Koundinya Vajjha Date: Wed, 27 Sep 2023 15:13:27 -0700 Subject: [PATCH 59/74] simpler harmonic not int proof --- Mathlib/NumberTheory/Padics/Harmonic.lean | 13 +++++-------- 1 file changed, 5 insertions(+), 8 deletions(-) diff --git a/Mathlib/NumberTheory/Padics/Harmonic.lean b/Mathlib/NumberTheory/Padics/Harmonic.lean index af97fa8f3e773..abfbfa999d68f 100644 --- a/Mathlib/NumberTheory/Padics/Harmonic.lean +++ b/Mathlib/NumberTheory/Padics/Harmonic.lean @@ -124,11 +124,8 @@ theorem padicValRat_two_harmonic {n : ℕ} (Hn : 2 ≤ n) : simp only [add_tsub_cancel_right] /-- The n-th harmonic number is not an integer for n ≥ 2. -/ -theorem harmonic_not_int {n : ℕ} (Hn : 2 ≤ n) : - ¬(harmonic n).isInt := by - refine' padicNorm.not_int_of_not_padic_int 2 _ - rw [padicNorm.eq_zpow_of_nonzero (harmonic_pos (by linarith)).ne'] - refine' one_lt_zpow one_lt_two _ (lt_neg.mp _) - rw [neg_zero, padicValRat_two_harmonic Hn, Left.neg_neg_iff, - Nat.cast_pos, Nat.log_pos_iff] - exact ⟨Hn,one_lt_two⟩ +theorem harmonic_not_int {n : ℕ} (h : 2 ≤ n) : ¬ (harmonic n).isInt := by + apply padicNorm.not_int_of_not_padic_int 2 + rw [padicNorm.eq_zpow_of_nonzero (harmonic_pos (ne_zero_of_lt h)).ne', + padicValRat_two_harmonic h, neg_neg, zpow_coe_nat] + exact one_lt_pow one_lt_two (Nat.log_pos one_lt_two h).ne' From 7ce7924aca740aeb9bf572afc91e03161b0b198d Mon Sep 17 00:00:00 2001 From: Koundinya Vajjha Date: Wed, 27 Sep 2023 15:15:46 -0700 Subject: [PATCH 60/74] simplify proof --- Mathlib/NumberTheory/Padics/Harmonic.lean | 140 ++++++---------------- 1 file changed, 38 insertions(+), 102 deletions(-) diff --git a/Mathlib/NumberTheory/Padics/Harmonic.lean b/Mathlib/NumberTheory/Padics/Harmonic.lean index abfbfa999d68f..d6dfbb1d7f03f 100644 --- a/Mathlib/NumberTheory/Padics/Harmonic.lean +++ b/Mathlib/NumberTheory/Padics/Harmonic.lean @@ -1,7 +1,7 @@ /- Copyright (c) 2023 Koundinya Vajjha. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Koundinya Vajjha +Authors: Koundinya Vajjha, Thomas Browning -/ import Mathlib.NumberTheory.Padics.PadicIntegers @@ -21,111 +21,47 @@ open BigOperators /-- The nth-harmonic number defined as a finset sum of consecutive reciprocals. -/ def harmonic : ℕ → ℚ := fun n => ∑ i in Finset.range n, (↑(i + 1))⁻¹ +@[simp] +lemma harmonic_zero : harmonic 0 = 0 := + rfl + +@[simp] +lemma harmonic_succ (n : ℕ) : harmonic (n + 1) = harmonic n + (↑(n + 1))⁻¹ := by + apply Finset.sum_range_succ + lemma harmonic_pos {n : ℕ} (Hn : n ≠ 0) : 0 < harmonic n := Finset.sum_pos (fun _ _ => inv_pos.mpr (by norm_cast; linarith)) (by rwa [Finset.nonempty_range_iff]) -lemma harmonic_singleton {n c : ℕ} (hc : c ∈ Finset.range n) : - harmonic n = 1 / (c + 1) + ∑ x in Finset.range n \ {c}, 1 / ((x : ℚ) + 1) := by - dsimp [harmonic] - simp only [Nat.cast_add, Nat.cast_one, one_div, Finset.mem_range, not_lt, - Finset.singleton_subset_iff] - rwa [add_comm, Finset.sum_eq_sum_diff_singleton_add (i := c)] - -lemma harmonic_singleton_ne_zero {c n : ℕ} (hn : 2 ≤ n) : - ∑ x in Finset.range n \ {c}, 1 / (x + 1 : ℚ) ≠ 0 := by - refine' ne_of_gt <| Finset.sum_pos (fun i _ => div_pos zero_lt_one _) - (Finset.range_nontrivial hn).sdiff_singleton_nonempty - norm_cast; simp only [add_pos_iff, or_true] - -lemma padicValRat_two_pow_div (r : ℕ) : padicValRat 2 (1 / 2 ^ r) = -r := by - rw [one_div, ← padicValRat.self_pow_div (p := 2) (r := r)] - simp only [one_div, Nat.cast_ofNat] - -/-- For `i` less than `n`, `2 ^ Nat.log 2 n` does not divide `i` unless `i` is equal to it. -/ -lemma nat_log_not_dvd {n i : ℕ} (Hi₁ : 0 < i) (Hi₂ : i ≠ 2 ^ Nat.log 2 n) (Hi₃ : i ≤ n) : - ¬ 2^(Nat.log 2 n) ∣ i := by - rintro ⟨c,Hc⟩ - have Hle : n < 2 ^ (Nat.log 2 n + 1) := Nat.lt_pow_succ_log_self (by simp) n - have Hpow : 2 ^ Nat.log 2 n * 2 = 2 ^ (Nat.log 2 n + 1) := by rfl - obtain ⟨k,Hk⟩ := Nat.even_or_odd' c - by_cases h : k = 0 - · subst h - rw [mul_zero, zero_add] at Hk - rcases Hk with rfl | rfl - · linarith - · exact Hi₂ (Hc ▸ mul_one (2 ^ Nat.log 2 n)) - rcases Hk with rfl | rfl - · rw [← mul_assoc, Hpow] at Hc - suffices Hnk : n < i by linarith - calc n ≤ n * k := Nat.le_mul_of_pos_right (Nat.pos_of_ne_zero h) - _ < 2^(Nat.log 2 n + 1) * k := Nat.mul_lt_mul Hle (le_refl _) (Nat.pos_of_ne_zero h) - _ = i := Hc.symm - · rw [mul_add, ← mul_assoc, Hpow, mul_one] at Hc - suffices Hnk : n < i by linarith - calc n < 2 ^ (Nat.log 2 n + 1) := Hle - _ ≤ 2 ^ (Nat.log 2 n + 1)*k := Nat.le_mul_of_pos_right (Nat.pos_of_ne_zero h) - _ < 2 ^ (Nat.log 2 n + 1)*k + 2 ^ Nat.log 2 n := by linarith - _ = i := Hc.symm - -lemma padicValRat_le_neg_nat_log_ne {n q : ℕ} (Hq₁ : 0 < q) - (Hq₂ : q ≤ n)(Hq₃ : q ≠ 2 ^ Nat.log 2 n) : - padicValRat 2 (1 / q) ≠ -Nat.log 2 n := by - rw [one_div, padicValRat.inv, padicValRat.of_nat, ne_eq, neg_inj, Nat.cast_inj] - exact (fun Hnot => nat_log_not_dvd (n := n) Hq₁ Hq₃ Hq₂ - (Hnot ▸ pow_padicValNat_dvd (p := 2) (n := q))) - -lemma padicValRat_le_neg_nat_log_le {p n q : ℕ} [hp : Fact (Nat.Prime p)] (Hq : q ≤ n) : - -Nat.log p n ≤ padicValRat p (1 / q) := by - rw [one_div, padicValRat.inv, padicValRat.of_nat, neg_le_neg_iff, Nat.cast_le] - exact le_nat_log_of_le Hq - -lemma padicValRat_ge_neg_Nat_log_lt (n q : ℕ) (Hq₁ : 0 < q) - (Hq₂ : q ≤ n) (Hq₃ : q ≠ 2 ^ Nat.log 2 n) : - padicValRat 2 (1 / 2 ^ Nat.log 2 n) < padicValRat 2 (1 / q) := by - rw [padicValRat_two_pow_div] - exact (lt_of_le_of_ne - (padicValRat_le_neg_nat_log_le (p := 2) Hq₂) - (padicValRat_le_neg_nat_log_ne Hq₁ Hq₂ Hq₃).symm) - -lemma two_pow_log_sub_one_mem_range {n : ℕ} (hn : 2 ≤ n) : - 2 ^ (Nat.log 2 n) - 1 ∈ Finset.range n := - Finset.mem_range.mpr <| lt_of_lt_of_le (Nat.sub_lt (pow_pos zero_lt_two _) zero_lt_one) - (Nat.pow_log_le_self 2 (by linarith)) +lemma Nat.log_ne_padicValNat_succ {n : ℕ} (hn : n ≠ 0) : log 2 n ≠ padicValNat 2 (n + 1) := by + rw [Ne, log_eq_iff (by simp [hn])] + rintro ⟨h1, h2⟩ + rw [← lt_add_one_iff, ← mul_one (2 ^ _)] at h1 + rw [← add_one_le_iff, pow_succ] at h2 + refine' not_dvd_of_between_consec_multiples h1 (lt_of_le_of_ne' h2 _) pow_padicValNat_dvd + exact pow_succ_padicValNat_not_dvd n.succ_ne_zero ∘ dvd_of_eq +lemma Nat.max_log_padicValNat_succ_eq_log_succ (n : ℕ) : + max (log 2 n) (padicValNat 2 (n + 1)) = log 2 (n + 1) := by + apply le_antisymm (max_le (le_log_of_pow_le one_lt_two (pow_log_le_add_one 2 n)) + (le_log_of_pow_le one_lt_two (le_of_dvd n.succ_pos pow_padicValNat_dvd))) + rw [le_max_iff, or_iff_not_imp_left, not_le] + intro h + replace h := le_antisymm (add_one_le_iff.mpr (lt_pow_of_log_lt one_lt_two h)) + (pow_log_le_self 2 n.succ_ne_zero) + rw [h, padicValNat.prime_pow, ← h] /-- The 2-adic valuation of the n-th harmonic number is the negative of the logarithm of n. -/ -theorem padicValRat_two_harmonic {n : ℕ} (Hn : 2 ≤ n) : - padicValRat 2 (harmonic n) = -Nat.log 2 n := by - rw [harmonic_singleton (two_pow_log_sub_one_mem_range Hn)] - simp only [gt_iff_lt, pow_pos, Nat.cast_pred, Nat.cast_pow, Nat.cast_ofNat, sub_add_cancel] - rw [padicValRat.add_eq_of_lt (p := 2) _ - (one_div_ne_zero <| pow_ne_zero _ two_ne_zero) (harmonic_singleton_ne_zero Hn) _] - apply padicValRat_two_pow_div - · have := harmonic_pos (n := n) - rw [harmonic_singleton (two_pow_log_sub_one_mem_range Hn)] at this - refine' ne_of_gt _ - simp only [ne_eq, ge_iff_le, gt_iff_lt, pow_pos, Nat.cast_pred, Nat.cast_pow, Nat.cast_ofNat, - sub_add_cancel, Finset.mem_range, not_lt, Finset.singleton_subset_iff] at this - exact this (by linarith) - · have := padicValRat.lt_sum_of_lt (p := 2) (j := 2 ^ Nat.log 2 n - 1) - (Finset.range_nontrivial Hn).sdiff_singleton_nonempty - (F := fun n => 1 / ((n + 1) : ℚ)) - (S := Finset.range n \ {2 ^ Nat.log 2 n - 1}) - simp only [Finset.mem_range, Finset.mem_sdiff, Finset.mem_singleton, - Finset.singleton_subset_iff, pow_pos, Nat.cast_pred, Nat.cast_pow, - sub_add_cancel, and_imp,gt_iff_lt] at this - refine' this (fun i Hi₁ Hi₂ => _) (fun i => div_pos zero_lt_one (by norm_cast; linarith)) - · have := padicValRat_ge_neg_Nat_log_lt (n := n) (i+1) - simp only [Nat.cast_add, Nat.cast_one] at this - refine' this (by linarith) (by linarith) (fun Hnot => _) - rw [← Hnot] at Hi₂ - refine' Hi₂ _ - simp only [add_tsub_cancel_right] - -/-- The n-th harmonic number is not an integer for n ≥ 2. -/ -theorem harmonic_not_int {n : ℕ} (h : 2 ≤ n) : ¬ (harmonic n).isInt := by - apply padicNorm.not_int_of_not_padic_int 2 - rw [padicNorm.eq_zpow_of_nonzero (harmonic_pos (ne_zero_of_lt h)).ne', - padicValRat_two_harmonic h, neg_neg, zpow_coe_nat] - exact one_lt_pow one_lt_two (Nat.log_pos one_lt_two h).ne' +theorem padicValRat_two_harmonic {n : ℕ} : padicValRat 2 (harmonic n) = -Nat.log 2 n := by + induction' n with n ih + · simp + · rcases eq_or_ne n 0 with rfl | hn + · simp + rw [harmonic_succ] + have key : padicValRat 2 (harmonic n) ≠ padicValRat 2 (↑(n + 1))⁻¹ + · rw [ih, padicValRat.inv, padicValRat.of_nat, Ne, neg_inj, Nat.cast_inj] + exact Nat.log_ne_padicValNat_succ hn + rw [padicValRat.add_eq_min (harmonic_succ n ▸ (harmonic_pos n.succ_ne_zero).ne') + (harmonic_pos hn).ne' (inv_ne_zero (Nat.cast_ne_zero.mpr n.succ_ne_zero)) key, ih, + padicValRat.inv, padicValRat.of_nat, min_neg_neg, neg_inj, ← Nat.cast_max, Nat.cast_inj] + exact Nat.max_log_padicValNat_succ_eq_log_succ n From 7212811e10716366a3ea5ac66a1e45f55c8b12f2 Mon Sep 17 00:00:00 2001 From: Koundinya Vajjha Date: Wed, 27 Sep 2023 15:37:09 -0700 Subject: [PATCH 61/74] add main result --- Mathlib/NumberTheory/Padics/Harmonic.lean | 7 +++++++ 1 file changed, 7 insertions(+) diff --git a/Mathlib/NumberTheory/Padics/Harmonic.lean b/Mathlib/NumberTheory/Padics/Harmonic.lean index d6dfbb1d7f03f..50575c8ec8d83 100644 --- a/Mathlib/NumberTheory/Padics/Harmonic.lean +++ b/Mathlib/NumberTheory/Padics/Harmonic.lean @@ -65,3 +65,10 @@ theorem padicValRat_two_harmonic {n : ℕ} : padicValRat 2 (harmonic n) = -Nat.l (harmonic_pos hn).ne' (inv_ne_zero (Nat.cast_ne_zero.mpr n.succ_ne_zero)) key, ih, padicValRat.inv, padicValRat.of_nat, min_neg_neg, neg_inj, ← Nat.cast_max, Nat.cast_inj] exact Nat.max_log_padicValNat_succ_eq_log_succ n + +/-- The n-th harmonic number is not an integer for n ≥ 2. -/ +theorem harmonic_not_int {n : ℕ} (h : 2 ≤ n) : ¬ (harmonic n).isInt := by + apply padicNorm.not_int_of_not_padic_int 2 + rw [padicNorm.eq_zpow_of_nonzero (harmonic_pos (ne_zero_of_lt h)).ne', + padicValRat_two_harmonic, neg_neg, zpow_coe_nat] + exact one_lt_pow one_lt_two (Nat.log_pos one_lt_two h).ne' From 00a352072633ade4dcec27f9a6432c198898f061 Mon Sep 17 00:00:00 2001 From: Koundinya Vajjha Date: Wed, 27 Sep 2023 15:37:50 -0700 Subject: [PATCH 62/74] spacing --- Mathlib/NumberTheory/Padics/Harmonic.lean | 1 + 1 file changed, 1 insertion(+) diff --git a/Mathlib/NumberTheory/Padics/Harmonic.lean b/Mathlib/NumberTheory/Padics/Harmonic.lean index 50575c8ec8d83..e7c5707465079 100644 --- a/Mathlib/NumberTheory/Padics/Harmonic.lean +++ b/Mathlib/NumberTheory/Padics/Harmonic.lean @@ -40,6 +40,7 @@ lemma Nat.log_ne_padicValNat_succ {n : ℕ} (hn : n ≠ 0) : log 2 n ≠ padicVa rw [← add_one_le_iff, pow_succ] at h2 refine' not_dvd_of_between_consec_multiples h1 (lt_of_le_of_ne' h2 _) pow_padicValNat_dvd exact pow_succ_padicValNat_not_dvd n.succ_ne_zero ∘ dvd_of_eq + lemma Nat.max_log_padicValNat_succ_eq_log_succ (n : ℕ) : max (log 2 n) (padicValNat 2 (n + 1)) = log 2 (n + 1) := by apply le_antisymm (max_le (le_log_of_pow_le one_lt_two (pow_log_le_add_one 2 n)) From 1a2cb2bdad1012ae270435c4ddec603fb3d88d97 Mon Sep 17 00:00:00 2001 From: Koundinya Vajjha Date: Wed, 27 Sep 2023 22:21:18 -0700 Subject: [PATCH 63/74] Update Mathlib/NumberTheory/Padics/PadicNorm.lean Co-authored-by: Thomas Browning --- Mathlib/NumberTheory/Padics/PadicNorm.lean | 1 + 1 file changed, 1 insertion(+) diff --git a/Mathlib/NumberTheory/Padics/PadicNorm.lean b/Mathlib/NumberTheory/Padics/PadicNorm.lean index 4b2065f907c23..fcb7b06dddb9d 100644 --- a/Mathlib/NumberTheory/Padics/PadicNorm.lean +++ b/Mathlib/NumberTheory/Padics/PadicNorm.lean @@ -308,6 +308,7 @@ theorem nat_lt_one_iff (m : ℕ) : padicNorm p m < 1 ↔ p ∣ m := by lemma _root_.Rat.isInt.eq_num {q : ℚ} (h : q.isInt) : q = q.num := by rw [Rat.isInt, Nat.beq_eq_true_eq] at h exact (Rat.coe_int_num_of_den_eq_one h).symm + /-- If a rational is not a p-adic integer, it is not an integer. -/ theorem not_int_of_not_padic_int (p : ℕ) {a : ℚ} [hp : Fact (Nat.Prime p)] (H : 1 < padicNorm p a) : ¬ a.isInt := by From d8b593a5adecef1311f36cea0707cd1514075082 Mon Sep 17 00:00:00 2001 From: Koundinya Vajjha Date: Wed, 27 Sep 2023 22:21:32 -0700 Subject: [PATCH 64/74] Update Mathlib/NumberTheory/Padics/PadicVal.lean Co-authored-by: Thomas Browning --- Mathlib/NumberTheory/Padics/PadicVal.lean | 12 +++++++----- 1 file changed, 7 insertions(+), 5 deletions(-) diff --git a/Mathlib/NumberTheory/Padics/PadicVal.lean b/Mathlib/NumberTheory/Padics/PadicVal.lean index bf237f3739a69..e23a2a8e67d20 100644 --- a/Mathlib/NumberTheory/Padics/PadicVal.lean +++ b/Mathlib/NumberTheory/Padics/PadicVal.lean @@ -586,11 +586,13 @@ theorem padicValNat_primes {q : ℕ} [hp : Fact p.Prime] [hq : Fact q.Prime] (ne #align padic_val_nat_primes padicValNat_primes /-- The p-adic valuation of `n` is less than or equal to its logarithm w.r.t `p`.-/ -lemma le_nat_log (n : ℕ) [hp : Fact (Nat.Prime p)] : padicValNat p n ≤ Nat.log p n := by - rcases eq_or_ne n 0 with rfl | h - · rw [padicValNat.zero, Nat.log_zero_right] - · refine' Nat.le_log_of_pow_le hp.elim.one_lt _ - exact le_of_dvd (Nat.pos_of_ne_zero h) pow_padicValNat_dvd +lemma padicValNat_le_nat_log (n : ℕ) : padicValNat p n ≤ Nat.log p n := by + rcases n with _ | n + · simp + rcases p with _ | _ | p + · simp + · simp + exact Nat.le_log_of_pow_le p.one_lt_succ_succ (le_of_dvd n.succ_pos pow_padicValNat_dvd) lemma le_nat_log_of_le {n₁ n₂ : ℕ} [Fact (Nat.Prime p)] (hn : n₁ ≤ n₂) : padicValNat p n₁ ≤ Nat.log p n₂ := le_trans (le_nat_log n₁) (Nat.log_mono_right hn) From e4a946ddbda86877242c1a812220dce99eab87d7 Mon Sep 17 00:00:00 2001 From: Koundinya Vajjha Date: Wed, 27 Sep 2023 22:21:48 -0700 Subject: [PATCH 65/74] Update Mathlib/NumberTheory/Padics/Harmonic.lean Co-authored-by: Thomas Browning --- Mathlib/NumberTheory/Padics/Harmonic.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/NumberTheory/Padics/Harmonic.lean b/Mathlib/NumberTheory/Padics/Harmonic.lean index e7c5707465079..232c456a31422 100644 --- a/Mathlib/NumberTheory/Padics/Harmonic.lean +++ b/Mathlib/NumberTheory/Padics/Harmonic.lean @@ -44,7 +44,7 @@ lemma Nat.log_ne_padicValNat_succ {n : ℕ} (hn : n ≠ 0) : log 2 n ≠ padicVa lemma Nat.max_log_padicValNat_succ_eq_log_succ (n : ℕ) : max (log 2 n) (padicValNat 2 (n + 1)) = log 2 (n + 1) := by apply le_antisymm (max_le (le_log_of_pow_le one_lt_two (pow_log_le_add_one 2 n)) - (le_log_of_pow_le one_lt_two (le_of_dvd n.succ_pos pow_padicValNat_dvd))) + (padicValNat_le_nat_log (n + 1))) rw [le_max_iff, or_iff_not_imp_left, not_le] intro h replace h := le_antisymm (add_one_le_iff.mpr (lt_pow_of_log_lt one_lt_two h)) From 40ef8dc212b60c6a39bfec11463476d91649f6ea Mon Sep 17 00:00:00 2001 From: Koundinya Vajjha Date: Wed, 27 Sep 2023 22:34:18 -0700 Subject: [PATCH 66/74] fix build --- Mathlib/NumberTheory/Padics/PadicVal.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/NumberTheory/Padics/PadicVal.lean b/Mathlib/NumberTheory/Padics/PadicVal.lean index e23a2a8e67d20..c67b640a3cadb 100644 --- a/Mathlib/NumberTheory/Padics/PadicVal.lean +++ b/Mathlib/NumberTheory/Padics/PadicVal.lean @@ -595,7 +595,7 @@ lemma padicValNat_le_nat_log (n : ℕ) : padicValNat p n ≤ Nat.log p n := by exact Nat.le_log_of_pow_le p.one_lt_succ_succ (le_of_dvd n.succ_pos pow_padicValNat_dvd) lemma le_nat_log_of_le {n₁ n₂ : ℕ} [Fact (Nat.Prime p)] (hn : n₁ ≤ n₂) : - padicValNat p n₁ ≤ Nat.log p n₂ := le_trans (le_nat_log n₁) (Nat.log_mono_right hn) + padicValNat p n₁ ≤ Nat.log p n₂ := le_trans (padicValNat_le_nat_log n₁) (Nat.log_mono_right hn) /-- The p-adic valuation of `n` is equal to the logarithm w.r.t `p` iff `n` is less than `p` raised to one plus the p-adic valuation of `n`. -/ From 2365d537da3a23805b7ee9622d422ef7eeb7bb96 Mon Sep 17 00:00:00 2001 From: Koundinya Vajjha Date: Thu, 28 Sep 2023 15:20:08 -0700 Subject: [PATCH 67/74] appease linter --- Mathlib/NumberTheory/Padics/PadicVal.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/NumberTheory/Padics/PadicVal.lean b/Mathlib/NumberTheory/Padics/PadicVal.lean index c67b640a3cadb..d40cfaaa5cd28 100644 --- a/Mathlib/NumberTheory/Padics/PadicVal.lean +++ b/Mathlib/NumberTheory/Padics/PadicVal.lean @@ -594,7 +594,7 @@ lemma padicValNat_le_nat_log (n : ℕ) : padicValNat p n ≤ Nat.log p n := by · simp exact Nat.le_log_of_pow_le p.one_lt_succ_succ (le_of_dvd n.succ_pos pow_padicValNat_dvd) -lemma le_nat_log_of_le {n₁ n₂ : ℕ} [Fact (Nat.Prime p)] (hn : n₁ ≤ n₂) : +lemma le_nat_log_of_le {n₁ n₂ : ℕ} (hn : n₁ ≤ n₂) : padicValNat p n₁ ≤ Nat.log p n₂ := le_trans (padicValNat_le_nat_log n₁) (Nat.log_mono_right hn) /-- The p-adic valuation of `n` is equal to the logarithm w.r.t `p` iff From aa60496c0f055540dda47982f22ccb2b95e82cd4 Mon Sep 17 00:00:00 2001 From: Koundinya Vajjha Date: Thu, 28 Sep 2023 15:23:02 -0700 Subject: [PATCH 68/74] move lemmas around --- Mathlib/NumberTheory/Padics/Harmonic.lean | 18 ------------------ Mathlib/NumberTheory/Padics/PadicVal.lean | 18 ++++++++++++++++++ 2 files changed, 18 insertions(+), 18 deletions(-) diff --git a/Mathlib/NumberTheory/Padics/Harmonic.lean b/Mathlib/NumberTheory/Padics/Harmonic.lean index 232c456a31422..da13777253efc 100644 --- a/Mathlib/NumberTheory/Padics/Harmonic.lean +++ b/Mathlib/NumberTheory/Padics/Harmonic.lean @@ -33,24 +33,6 @@ lemma harmonic_pos {n : ℕ} (Hn : n ≠ 0) : 0 < harmonic n := Finset.sum_pos (fun _ _ => inv_pos.mpr (by norm_cast; linarith)) (by rwa [Finset.nonempty_range_iff]) -lemma Nat.log_ne_padicValNat_succ {n : ℕ} (hn : n ≠ 0) : log 2 n ≠ padicValNat 2 (n + 1) := by - rw [Ne, log_eq_iff (by simp [hn])] - rintro ⟨h1, h2⟩ - rw [← lt_add_one_iff, ← mul_one (2 ^ _)] at h1 - rw [← add_one_le_iff, pow_succ] at h2 - refine' not_dvd_of_between_consec_multiples h1 (lt_of_le_of_ne' h2 _) pow_padicValNat_dvd - exact pow_succ_padicValNat_not_dvd n.succ_ne_zero ∘ dvd_of_eq - -lemma Nat.max_log_padicValNat_succ_eq_log_succ (n : ℕ) : - max (log 2 n) (padicValNat 2 (n + 1)) = log 2 (n + 1) := by - apply le_antisymm (max_le (le_log_of_pow_le one_lt_two (pow_log_le_add_one 2 n)) - (padicValNat_le_nat_log (n + 1))) - rw [le_max_iff, or_iff_not_imp_left, not_le] - intro h - replace h := le_antisymm (add_one_le_iff.mpr (lt_pow_of_log_lt one_lt_two h)) - (pow_log_le_self 2 n.succ_ne_zero) - rw [h, padicValNat.prime_pow, ← h] - /-- The 2-adic valuation of the n-th harmonic number is the negative of the logarithm of n. -/ theorem padicValRat_two_harmonic {n : ℕ} : padicValRat 2 (harmonic n) = -Nat.log 2 n := by diff --git a/Mathlib/NumberTheory/Padics/PadicVal.lean b/Mathlib/NumberTheory/Padics/PadicVal.lean index d40cfaaa5cd28..0c8ed43269e8c 100644 --- a/Mathlib/NumberTheory/Padics/PadicVal.lean +++ b/Mathlib/NumberTheory/Padics/PadicVal.lean @@ -604,6 +604,24 @@ lemma nat_log_eq_padicValNat_iff {n : ℕ} [hp : Fact (Nat.Prime p)] (hn : 0 < n rw [Nat.log_eq_iff (Or.inr ⟨(Nat.Prime.one_lt' p).out, by linarith⟩), and_iff_right_iff_imp] exact (fun _ => Nat.le_of_dvd hn pow_padicValNat_dvd) +lemma Nat.log_ne_padicValNat_succ {n : ℕ} (hn : n ≠ 0) : log 2 n ≠ padicValNat 2 (n + 1) := by + rw [Ne, log_eq_iff (by simp [hn])] + rintro ⟨h1, h2⟩ + rw [← lt_add_one_iff, ← mul_one (2 ^ _)] at h1 + rw [← add_one_le_iff, pow_succ] at h2 + refine' not_dvd_of_between_consec_multiples h1 (lt_of_le_of_ne' h2 _) pow_padicValNat_dvd + exact pow_succ_padicValNat_not_dvd n.succ_ne_zero ∘ dvd_of_eq + +lemma Nat.max_log_padicValNat_succ_eq_log_succ (n : ℕ) : + max (log 2 n) (padicValNat 2 (n + 1)) = log 2 (n + 1) := by + apply le_antisymm (max_le (le_log_of_pow_le one_lt_two (pow_log_le_add_one 2 n)) + (padicValNat_le_nat_log (n + 1))) + rw [le_max_iff, or_iff_not_imp_left, not_le] + intro h + replace h := le_antisymm (add_one_le_iff.mpr (lt_pow_of_log_lt one_lt_two h)) + (pow_log_le_self 2 n.succ_ne_zero) + rw [h, padicValNat.prime_pow, ← h] + open BigOperators theorem range_pow_padicValNat_subset_divisors {n : ℕ} (hn : n ≠ 0) : From 3653d5b3f3f6263be5ae4b8fa1a08f8a089faaf8 Mon Sep 17 00:00:00 2001 From: Koundinya Vajjha Date: Thu, 28 Sep 2023 15:31:31 -0700 Subject: [PATCH 69/74] delete le_nat_log_of_le --- Mathlib/NumberTheory/Padics/PadicVal.lean | 3 --- 1 file changed, 3 deletions(-) diff --git a/Mathlib/NumberTheory/Padics/PadicVal.lean b/Mathlib/NumberTheory/Padics/PadicVal.lean index 0c8ed43269e8c..8056843156916 100644 --- a/Mathlib/NumberTheory/Padics/PadicVal.lean +++ b/Mathlib/NumberTheory/Padics/PadicVal.lean @@ -594,9 +594,6 @@ lemma padicValNat_le_nat_log (n : ℕ) : padicValNat p n ≤ Nat.log p n := by · simp exact Nat.le_log_of_pow_le p.one_lt_succ_succ (le_of_dvd n.succ_pos pow_padicValNat_dvd) -lemma le_nat_log_of_le {n₁ n₂ : ℕ} (hn : n₁ ≤ n₂) : - padicValNat p n₁ ≤ Nat.log p n₂ := le_trans (padicValNat_le_nat_log n₁) (Nat.log_mono_right hn) - /-- The p-adic valuation of `n` is equal to the logarithm w.r.t `p` iff `n` is less than `p` raised to one plus the p-adic valuation of `n`. -/ lemma nat_log_eq_padicValNat_iff {n : ℕ} [hp : Fact (Nat.Prime p)] (hn : 0 < n) : From f4cbf9c2526ddf98c9d47382097f471dee3f2361 Mon Sep 17 00:00:00 2001 From: Koundinya Vajjha Date: Thu, 28 Sep 2023 15:49:35 -0700 Subject: [PATCH 70/74] Update Mathlib/Data/Finset/Basic.lean Co-authored-by: Thomas Browning --- Mathlib/Data/Finset/Basic.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/Data/Finset/Basic.lean b/Mathlib/Data/Finset/Basic.lean index e477814ac71fc..f32216505d9dc 100644 --- a/Mathlib/Data/Finset/Basic.lean +++ b/Mathlib/Data/Finset/Basic.lean @@ -2329,7 +2329,7 @@ theorem Nontrivial.sdiff_singleton_nonempty {c : α} {s : Finset α} (hS : s.Non (s \ {c}).Nonempty := by rw [Finset.sdiff_nonempty, Finset.subset_singleton_iff] push_neg - exact ⟨by rintro rfl; exact Finset.not_nontrivial_empty hS, Finset.Nontrivial.ne_singleton hS⟩ + exact ⟨by rintro rfl; exact Finset.not_nontrivial_empty hS, hS.ne_singleton⟩ theorem sdiff_sdiff_left' (s t u : Finset α) : (s \ t) \ u = s \ t ∩ (s \ u) := _root_.sdiff_sdiff_left' From aeb9acbf330a4e6e8943936531d3746e3d493743 Mon Sep 17 00:00:00 2001 From: Koundinya Vajjha Date: Thu, 28 Sep 2023 15:50:27 -0700 Subject: [PATCH 71/74] Update Mathlib/Data/Finset/Basic.lean Co-authored-by: Thomas Browning --- Mathlib/Data/Finset/Basic.lean | 8 +++----- 1 file changed, 3 insertions(+), 5 deletions(-) diff --git a/Mathlib/Data/Finset/Basic.lean b/Mathlib/Data/Finset/Basic.lean index f32216505d9dc..6ad48bfe42588 100644 --- a/Mathlib/Data/Finset/Basic.lean +++ b/Mathlib/Data/Finset/Basic.lean @@ -3109,11 +3109,9 @@ theorem range_filter_eq {n m : ℕ} : (range n).filter (· = m) = if m < n then · simp #align finset.range_filter_eq Finset.range_filter_eq -lemma range_nontrivial {n : ℕ} (hn : 2 ≤ n) : (Finset.range n).Nontrivial := by - refine' Set.nontrivial_of_exists_lt ⟨0,_,1,_,zero_lt_one⟩ <;> - simp only [Finset.coe_range, Set.mem_Iio] - exact lt_of_lt_of_le zero_lt_two hn - exact lt_of_lt_of_le one_lt_two hn +lemma range_nontrivial {n : ℕ} (hn : 1 < n) : (Finset.range n).Nontrivial := by + rw [Finset.Nontrivial, Finset.coe_range] + exact ⟨0, zero_lt_one.trans hn, 1, hn, zero_ne_one⟩ end Range From 12f33fb490e7d75cd858b1c1227c029c64258fc0 Mon Sep 17 00:00:00 2001 From: Koundinya Vajjha Date: Thu, 28 Sep 2023 15:50:46 -0700 Subject: [PATCH 72/74] Update Mathlib/NumberTheory/Padics/PadicVal.lean Co-authored-by: Thomas Browning --- Mathlib/NumberTheory/Padics/PadicVal.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/NumberTheory/Padics/PadicVal.lean b/Mathlib/NumberTheory/Padics/PadicVal.lean index 8056843156916..fc1b68b5ad00b 100644 --- a/Mathlib/NumberTheory/Padics/PadicVal.lean +++ b/Mathlib/NumberTheory/Padics/PadicVal.lean @@ -455,7 +455,7 @@ lemma add_eq_min {q r : ℚ} (hqr : q + r ≠ 0) (hq : q ≠ 0) (hr : r ≠ 0) lemma add_eq_of_lt {q r : ℚ} (hqr : q + r ≠ 0) (hq : q ≠ 0) (hr : r ≠ 0) (hval : padicValRat p q < padicValRat p r) : padicValRat p (q + r) = padicValRat p q := by - rw [add_eq_min hqr hq hr (ne_of_lt hval),min_eq_left (le_of_lt hval)] + rw [add_eq_min hqr hq hr (ne_of_lt hval), min_eq_left (le_of_lt hval)] lemma lt_add_of_lt {q r₁ r₂ : ℚ} (hqr : r₁ + r₂ ≠ 0) (hval₁ : padicValRat p q < padicValRat p r₁) (hval₂ : padicValRat p q < padicValRat p r₂) : From 26fa1932f933f289e8ef8d9a1ffc2ee290dfbe8f Mon Sep 17 00:00:00 2001 From: Koundinya Vajjha Date: Thu, 28 Sep 2023 17:48:48 -0700 Subject: [PATCH 73/74] Update Mathlib/NumberTheory/Padics/PadicVal.lean Co-authored-by: Thomas Browning --- Mathlib/NumberTheory/Padics/PadicVal.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/NumberTheory/Padics/PadicVal.lean b/Mathlib/NumberTheory/Padics/PadicVal.lean index fc1b68b5ad00b..5af746959e6e5 100644 --- a/Mathlib/NumberTheory/Padics/PadicVal.lean +++ b/Mathlib/NumberTheory/Padics/PadicVal.lean @@ -463,7 +463,7 @@ lemma lt_add_of_lt {q r₁ r₂ : ℚ} (hqr : r₁ + r₂ ≠ 0) lt_of_lt_of_le (lt_min hval₁ hval₂) (padicValRat.min_le_padicValRat_add hqr) @[simp] -lemma self_pow_div (r : ℕ) : padicValRat p ((p : ℚ) ^ r)⁻¹ = -r := by +lemma self_pow_inv (r : ℕ) : padicValRat p ((p : ℚ) ^ r)⁻¹ = -r := by rw [padicValRat.inv, neg_inj, padicValRat.pow (Nat.cast_ne_zero.mpr hp.elim.ne_zero), padicValRat.self hp.elim.one_lt, mul_one] From c7549949a43cbac1ee1b2ce81c2ce3dd3dcb731d Mon Sep 17 00:00:00 2001 From: Koundinya Vajjha Date: Thu, 28 Sep 2023 20:15:29 -0700 Subject: [PATCH 74/74] move and rename --- Mathlib/Data/Rat/Defs.lean | 4 ++++ Mathlib/NumberTheory/Padics/PadicNorm.lean | 6 +----- 2 files changed, 5 insertions(+), 5 deletions(-) diff --git a/Mathlib/Data/Rat/Defs.lean b/Mathlib/Data/Rat/Defs.lean index 61fa0800e2f4f..1512f4c6965df 100644 --- a/Mathlib/Data/Rat/Defs.lean +++ b/Mathlib/Data/Rat/Defs.lean @@ -511,6 +511,10 @@ theorem coe_int_num_of_den_eq_one {q : ℚ} (hq : q.den = 1) : (q.num : ℚ) = q rfl #align rat.coe_int_num_of_denom_eq_one Rat.coe_int_num_of_den_eq_one +lemma eq_num_of_isInt {q : ℚ} (h : q.isInt) : q = q.num := by + rw [Rat.isInt, Nat.beq_eq_true_eq] at h + exact (Rat.coe_int_num_of_den_eq_one h).symm + theorem den_eq_one_iff (r : ℚ) : r.den = 1 ↔ ↑r.num = r := ⟨Rat.coe_int_num_of_den_eq_one, fun h => h ▸ Rat.coe_int_den r.num⟩ #align rat.denom_eq_one_iff Rat.den_eq_one_iff diff --git a/Mathlib/NumberTheory/Padics/PadicNorm.lean b/Mathlib/NumberTheory/Padics/PadicNorm.lean index fcb7b06dddb9d..acdb5ffa9f063 100644 --- a/Mathlib/NumberTheory/Padics/PadicNorm.lean +++ b/Mathlib/NumberTheory/Padics/PadicNorm.lean @@ -305,15 +305,11 @@ theorem nat_lt_one_iff (m : ℕ) : padicNorm p m < 1 ↔ p ∣ m := by rw [← Int.coe_nat_dvd, ← int_lt_one_iff, Int.cast_ofNat] #align padic_norm.nat_lt_one_iff padicNorm.nat_lt_one_iff -lemma _root_.Rat.isInt.eq_num {q : ℚ} (h : q.isInt) : q = q.num := by - rw [Rat.isInt, Nat.beq_eq_true_eq] at h - exact (Rat.coe_int_num_of_den_eq_one h).symm - /-- If a rational is not a p-adic integer, it is not an integer. -/ theorem not_int_of_not_padic_int (p : ℕ) {a : ℚ} [hp : Fact (Nat.Prime p)] (H : 1 < padicNorm p a) : ¬ a.isInt := by contrapose! H - rw [Rat.isInt.eq_num H] + rw [Rat.eq_num_of_isInt H] apply padicNorm.of_int open BigOperators