Update specialFunctions_Integrals.lean

This commit is contained in:
Stefan Kebekus 2024-08-13 10:25:03 +02:00
parent bc8fed96b0
commit 0cc0c81508
1 changed files with 26 additions and 25 deletions

View File

@ -66,9 +66,10 @@ theorem logInt
lemma int₁ :
∫ x in (0)..(2 * π), log ‖circleMap 0 1 x - 1‖ = 0 := by
have {x : } : Complex.normSq (circleMap 0 1 x - 1) = 2 - 2 * cos x := 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]
@ -79,35 +80,35 @@ lemma int₁ :
_ = 2 - 2 * cos x := by
rw [sin_sq_add_cos_sq]
norm_num
have {x : } : 2 - 2 * cos x = 4 * sin (x / 2) ^ 2 := by
calc 2 - 2 * cos x
_ = 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]
norm_num
ring
_ = 4 * sin (x / 2) ^ 2 := by
nth_rw 1 [← mul_one 4]
nth_rw 1 [← sin_sq_add_cos_sq (x / 2)]
rw [mul_add]
abel
dsimp [Complex.abs]
sorry
sorry
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