Update specialFunctions_Integrals.lean
This commit is contained in:
parent
0cb1914b18
commit
bc8fed96b0
|
@ -5,6 +5,8 @@ import Mathlib.MeasureTheory.Integral.CircleIntegral
|
||||||
open scoped Interval Topology
|
open scoped Interval Topology
|
||||||
open Real Filter MeasureTheory intervalIntegral
|
open Real Filter MeasureTheory intervalIntegral
|
||||||
|
|
||||||
|
-- The following theorem was suggested by Gareth Ma on Zulip
|
||||||
|
|
||||||
theorem logInt
|
theorem logInt
|
||||||
{t : ℝ}
|
{t : ℝ}
|
||||||
(ht : 0 < t) :
|
(ht : 0 < t) :
|
||||||
|
@ -63,6 +65,50 @@ theorem logInt
|
||||||
|
|
||||||
lemma int₁ :
|
lemma int₁ :
|
||||||
∫ x in (0)..(2 * π), log ‖circleMap 0 1 x - 1‖ = 0 := by
|
∫ x in (0)..(2 * π), log ‖circleMap 0 1 x - 1‖ = 0 := by
|
||||||
dsimp [circleMap]
|
|
||||||
|
|
||||||
|
|
||||||
|
have {x : ℝ} : Complex.normSq (circleMap 0 1 x - 1) = 2 - 2 * cos x := by
|
||||||
|
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
|
||||||
|
|
||||||
|
|
||||||
|
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
|
||||||
|
_ = 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
|
||||||
|
|
||||||
|
|
||||||
|
|
||||||
|
|
||||||
sorry
|
sorry
|
||||||
|
|
Loading…
Reference in New Issue