Skip to content

Commit

Permalink
Merge pull request #210 from pitmonticone/chore-unused-variable
Browse files Browse the repository at this point in the history
chore: fix unused variable warnings
  • Loading branch information
AlexKontorovich authored Nov 13, 2024
2 parents d652430 + c367aca commit b1b1d8d
Showing 1 changed file with 10 additions and 10 deletions.
20 changes: 10 additions & 10 deletions PrimeNumberTheoremAnd/ZetaBounds.lean
Original file line number Diff line number Diff line change
Expand Up @@ -1343,7 +1343,7 @@ $$
\end{lemma}
%%-/
lemma ZetaUpperBnd :
∃ (A : ℝ) (hA : A ∈ Ioc 0 (1 / 2)) (C : ℝ) (Cpos : 0 < C), ∀ (σ : ℝ) (t : ℝ) (t_ge : 3 < |t|)
∃ (A : ℝ) (_ : A ∈ Ioc 0 (1 / 2)) (C : ℝ) (_ : 0 < C), ∀ (σ : ℝ) (t : ℝ) (_ : 3 < |t|)
(_ : σ ∈ Icc (1 - A / Real.log |t|) 2), ‖ζ (σ + t * I)‖ ≤ C * Real.log |t| := by
let A := (1 / 2 : ℝ)
let C := Real.exp A * (5 + 8 * 2) -- the 2 comes from ZetaBnd_aux1
Expand Down Expand Up @@ -1675,8 +1675,8 @@ $$
\end{lemma}
%%-/
lemma ZetaDerivUpperBnd :
∃ (A : ℝ) (hA : A ∈ Ioc 0 (1 / 2)) (C : ℝ) (Cpos : 0 < C), ∀ (σ : ℝ) (t : ℝ) (t_gt : 3 < |t|)
( : σ ∈ Icc (1 - A / Real.log |t|) 2),
∃ (A : ℝ) (_ : A ∈ Ioc 0 (1 / 2)) (C : ℝ) (_ : 0 < C), ∀ (σ : ℝ) (t : ℝ) (_ : 3 < |t|)
(_ : σ ∈ Icc (1 - A / Real.log |t|) 2),
‖deriv ζ (σ + t * I)‖ ≤ C * Real.log |t| ^ 2 := by
obtain ⟨A, hA, _, _, _⟩ := ZetaUpperBnd
let C := Real.exp A * 27
Expand Down Expand Up @@ -1789,7 +1789,7 @@ $$
\end{lemma}
%%-/
lemma ZetaNear1BndExact:
∃ (c : ℝ) (cpos : 0 < c), ∀ (σ : ℝ) (_ : σ ∈ Ioc 1 2), ‖ζ σ‖ ≤ c / (σ - 1) := by
∃ (c : ℝ) (_ : 0 < c), ∀ (σ : ℝ) (_ : σ ∈ Ioc 1 2), ‖ζ σ‖ ≤ c / (σ - 1) := by
have := ZetaNear1BndFilter
rw [Asymptotics.isBigO_iff] at this
obtain ⟨c, U, hU, V, hV, h⟩ := this
Expand Down Expand Up @@ -2063,8 +2063,8 @@ $$
\end{lemma}
%%-/
lemma Zeta_diff_Bnd :
∃ (A : ℝ) (hA : A ∈ Ioc 0 (1 / 2)) (C : ℝ) (Cpos : 0 < C), ∀ (σ₁ σ₂ : ℝ) (t : ℝ) (t_gt : 3 < |t|)
(σ₁_ge : 1 - A / Real.log |t| ≤ σ₁) (σ₂_le : σ₂ ≤ 2) (σ₁_lt_σ₂ : σ₁ < σ₂),
∃ (A : ℝ) (_ : A ∈ Ioc 0 (1 / 2)) (C : ℝ) (_ : 0 < C), ∀ (σ₁ σ₂ : ℝ) (t : ℝ) (_ : 3 < |t|)
(_ : 1 - A / Real.log |t| ≤ σ₁) (_ : σ₂ ≤ 2) (_ : σ₁ < σ₂),
‖ζ (σ₂ + t * I) - ζ (σ₁ + t * I)‖ ≤ C * Real.log |t| ^ 2 * (σ₂ - σ₁) := by
obtain ⟨A, hA, C, Cpos, hC⟩ := ZetaDerivUpperBnd
refine ⟨A, hA, C, Cpos, ?_⟩
Expand Down Expand Up @@ -2118,7 +2118,7 @@ $$
\end{lemma}
%%-/
lemma ZetaInvBnd :
∃ (A : ℝ) (hA : A ∈ Ioc 0 (1 / 2)) (C : ℝ) (Cpos : 0 < C), ∀ (σ : ℝ) (t : ℝ) (t_gt : 3 < |t|)
∃ (A : ℝ) (_ : A ∈ Ioc 0 (1 / 2)) (C : ℝ) (_ : 0 < C), ∀ (σ : ℝ) (t : ℝ) (_ : 3 < |t|)
(_ : σ ∈ Ico (1 - A / (Real.log |t|) ^ 9) 1),
1 / ‖ζ (σ + t * I)‖ ≤ C * (Real.log |t|) ^ (7 : ℝ) := by
obtain ⟨C', C'pos, hC₁⟩ := ZetaInvBound2
Expand Down Expand Up @@ -2226,8 +2226,8 @@ $$
\end{lemma}
%%-/
lemma LogDerivZetaBnd :
∃ (A : ℝ) (hA : A ∈ Ioc 0 (1 / 2)) (C : ℝ) (Cpos : 0 < C), ∀ (σ : ℝ) (t : ℝ) (t_gt : 3 < |t|)
( : σ ∈ Ico (1 - A / Real.log |t| ^ 9) 1),
∃ (A : ℝ) (_ : A ∈ Ioc 0 (1 / 2)) (C : ℝ) (_ : 0 < C), ∀ (σ : ℝ) (t : ℝ) (_ : 3 < |t|)
(_ : σ ∈ Ico (1 - A / Real.log |t| ^ 9) 1),
‖deriv ζ (σ + t * I) / ζ (σ + t * I)‖ ≤
C * Real.log |t| ^ 9 := by
obtain ⟨A, hA, C, hC, h⟩ := ZetaInvBnd
Expand Down Expand Up @@ -2271,7 +2271,7 @@ $$
\end{lemma}
%%-/
lemma LogDerivZetaBndAlt :
∃ A > 0, ∀ (σ) ( : σ ∈ Ico ((1 : ℝ) / 2) (1 : ℝ)),
∃ A > 0, ∀ (σ) (_ : σ ∈ Ico ((1 : ℝ) / 2) (1 : ℝ)),
(fun (t : ℝ) ↦ deriv ζ (σ + t * I) / ζ (σ + t * I)) =O[cocompact ℝ ⊓
Filter.principal {t | 1 - A / Real.log |t| ^ 9 < σ}]
fun t ↦ Real.log |t| ^ 9 := by
Expand Down

0 comments on commit b1b1d8d

Please sign in to comment.