diff --git a/Nevanlinna/specialFunctions_CircleIntegral_affine.lean b/Nevanlinna/specialFunctions_CircleIntegral_affine.lean index 03603be..bfe992a 100644 --- a/Nevanlinna/specialFunctions_CircleIntegral_affine.lean +++ b/Nevanlinna/specialFunctions_CircleIntegral_affine.lean @@ -130,6 +130,7 @@ lemma int₀ -- Integral of log ‖circleMap 0 1 x - 1‖ +-- integral 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 @@ -248,6 +249,7 @@ lemma int₁ : -- Integral of log ‖circleMap 0 1 x - a‖, for ‖a‖ = 1 +-- integrability lemma int'₂ {a : ℂ} (ha : ‖a‖ = 1) : @@ -310,7 +312,7 @@ lemma int'₂ simp_rw [this] at A exact A - +-- integral lemma int₂ {a : ℂ} (ha : ‖a‖ = 1) : @@ -370,7 +372,7 @@ lemma int₂ simp_rw [this] exact int₁ - +-- integral lemma int₃ {a : ℂ} (ha : a ∈ Metric.closedBall 0 1) : @@ -383,7 +385,7 @@ lemma int₃ simp linarith - +-- integral lemma int₄ {a : ℂ} {R : ℝ} @@ -485,3 +487,33 @@ lemma int₄ exact Ne.symm (ne_of_lt hR) rw [div_self this] at h₂a tauto + + +lemma intervalIntegrable_logAbs_circleMap_sub_const + {a c : ℂ} + {r : ℝ} + (hr : r ≠ 0) : + IntervalIntegrable (fun x ↦ log ‖circleMap c r x - a‖) volume 0 (2 * π) := by + + have {x : ℝ} : log ‖circleMap c r x - a‖ = log ‖r * (circleMap 0 1 x - r⁻¹ * (a - c))‖ := by + unfold circleMap + congr 2 + simp + rw [mul_sub] + rw [← mul_assoc] + simp [hr] + ring + simp_rw [this] + + have {x : ℝ} : log ‖r * (circleMap 0 1 x - r⁻¹ * (a - c))‖ = log r + log ‖(circleMap 0 1 x - r⁻¹ * (a - c))‖ := by + rw [norm_mul] + rw [log_mul] + simp + -- + simp [hr] + -- + + + + sorry + sorry