94 lines
2.8 KiB
Plaintext
94 lines
2.8 KiB
Plaintext
import Mathlib.Analysis.SpecialFunctions.Integrals
|
||
import Mathlib.Analysis.SpecialFunctions.Log.NegMulLog
|
||
import Mathlib.MeasureTheory.Integral.CircleIntegral
|
||
import Mathlib.MeasureTheory.Measure.Restrict
|
||
import Nevanlinna.specialFunctions_Integral_log_sin
|
||
|
||
open scoped Interval Topology
|
||
open Real Filter MeasureTheory intervalIntegral
|
||
|
||
|
||
|
||
|
||
lemma int₁₁ : ∫ (x : ℝ) in (0)..π, log (4 * sin x ^ 2) = 0 := by
|
||
|
||
have t₁ {x : ℝ} : x ∈ Set.Ioo 0 π → log (4 * sin x ^ 2) = log 4 + 2 * log (sin x) := by
|
||
intro hx
|
||
rw [log_mul, log_pow]
|
||
rfl
|
||
exact Ne.symm (NeZero.ne' 4)
|
||
apply pow_ne_zero 2
|
||
apply (fun a => Ne.symm (ne_of_lt a))
|
||
exact sin_pos_of_mem_Ioo hx
|
||
|
||
|
||
have t₂ : Set.EqOn (fun y ↦ log (4 * sin y ^ 2)) (fun y ↦ log 4 + 2 * log (sin y)) (Set.Ioo 0 π) := by
|
||
intro x hx
|
||
simp
|
||
rw [t₁ hx]
|
||
|
||
rw [intervalIntegral.integral_congr_volume pi_pos t₂]
|
||
rw [intervalIntegral.integral_add]
|
||
rw [intervalIntegral.integral_const_mul]
|
||
simp
|
||
rw [integral_log_sin₂]
|
||
have : (4 : ℝ) = 2 * 2 := by norm_num
|
||
rw [this, log_mul]
|
||
ring
|
||
norm_num
|
||
norm_num
|
||
-- IntervalIntegrable (fun x => log 4) volume 0 π
|
||
simp
|
||
-- IntervalIntegrable (fun x => 2 * log (sin x)) volume 0 π
|
||
apply IntervalIntegrable.const_mul
|
||
exact intervalIntegrable_log_sin
|
||
|
||
|
||
lemma int₁ :
|
||
∫ x in (0)..(2 * π), log ‖circleMap 0 1 x - 1‖ = 0 := by
|
||
|
||
have {x : ℝ} : log ‖circleMap 0 1 x - 1‖ = log (4 * sin (x / 2) ^ 2) / 2 := by
|
||
dsimp [Complex.abs]
|
||
rw [log_sqrt (Complex.normSq_nonneg (circleMap 0 1 x - 1))]
|
||
congr
|
||
calc Complex.normSq (circleMap 0 1 x - 1)
|
||
_ = (cos x - 1) * (cos x - 1) + sin x * sin x := by
|
||
dsimp [circleMap]
|
||
rw [Complex.normSq_apply]
|
||
simp
|
||
_ = sin x ^ 2 + cos x ^ 2 + 1 - 2 * cos x := by
|
||
ring
|
||
_ = 2 - 2 * cos x := by
|
||
rw [sin_sq_add_cos_sq]
|
||
norm_num
|
||
_ = 2 - 2 * cos (2 * (x / 2)) := by
|
||
rw [← mul_div_assoc]
|
||
congr; norm_num
|
||
_ = 4 - 4 * Real.cos (x / 2) ^ 2 := by
|
||
rw [cos_two_mul]
|
||
ring
|
||
_ = 4 * sin (x / 2) ^ 2 := by
|
||
nth_rw 1 [← mul_one 4, ← sin_sq_add_cos_sq (x / 2)]
|
||
ring
|
||
simp_rw [this]
|
||
simp
|
||
|
||
have : ∫ (x : ℝ) in (0)..2 * π, log (4 * sin (x / 2) ^ 2) = 2 * ∫ (x : ℝ) in (0)..π, log (4 * sin x ^ 2) := by
|
||
have : 1 = 2 * (2 : ℝ)⁻¹ := by exact Eq.symm (mul_inv_cancel_of_invertible 2)
|
||
nth_rw 1 [← one_mul (∫ (x : ℝ) in (0)..2 * π, log (4 * sin (x / 2) ^ 2))]
|
||
rw [← mul_inv_cancel_of_invertible 2, mul_assoc]
|
||
let f := fun y ↦ log (4 * sin y ^ 2)
|
||
have {x : ℝ} : log (4 * sin (x / 2) ^ 2) = f (x / 2) := by simp
|
||
conv =>
|
||
left
|
||
right
|
||
right
|
||
arg 1
|
||
intro x
|
||
rw [this]
|
||
rw [intervalIntegral.inv_mul_integral_comp_div 2]
|
||
simp
|
||
rw [this]
|
||
simp
|
||
exact int₁₁
|