working
This commit is contained in:
		| @@ -36,10 +36,13 @@ lemma l₂ {x : ℝ} : ‖(circleMap 0 1 x) - a‖ = ‖1 - (circleMap 0 1 (-x)) | ||||
|     dsimp [circleMap] | ||||
|     simp | ||||
|  | ||||
|  | ||||
| -- Integral of log ‖circleMap 0 1 x - a‖, for ‖a‖ < 1. | ||||
|  | ||||
| lemma int₀ | ||||
|   {a : ℂ} | ||||
|   (ha : a ∈ Metric.ball 0 1) : | ||||
|   ∫ (x : ℝ) in (0)..2 * Real.pi, Real.log ‖circleMap 0 1 x - a‖ = 0 := by | ||||
|   ∫ (x : ℝ) in (0)..2 * π, log ‖circleMap 0 1 x - a‖ = 0 := by | ||||
|  | ||||
|   by_cases h₁a : a = 0 | ||||
|   · -- case: a = 0 | ||||
| @@ -108,6 +111,8 @@ lemma int₀ | ||||
|   exact A | ||||
|  | ||||
|  | ||||
| -- Integral of log ‖circleMap 0 1 x - 1‖ | ||||
|  | ||||
| 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 | ||||
| @@ -141,11 +146,7 @@ lemma int₁₁ : ∫ (x : ℝ) in (0)..π, log (4 * sin x ^ 2) = 0 := by | ||||
|   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 | ||||
| lemma logAffineHelper {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 | ||||
| @@ -168,7 +169,20 @@ lemma int₁ : | ||||
|   _ = 4 * sin (x / 2) ^ 2 := by | ||||
|     nth_rw 1 [← mul_one 4, ← sin_sq_add_cos_sq (x / 2)] | ||||
|     ring | ||||
|   simp_rw [this] | ||||
|  | ||||
| lemma int'₁ : -- Integrability of log ‖circleMap 0 1 x - 1‖ | ||||
|   IntervalIntegrable (fun x ↦ log ‖circleMap 0 1 x - 1‖) volume 0 (2 * π) := by | ||||
|   simp_rw [logAffineHelper] | ||||
|   rw [← IntervalIntegrable.comp_mul_left_iff (c := 2) (Ne.symm (NeZero.ne' 2))] | ||||
|   simp | ||||
|  | ||||
|   sorry | ||||
|  | ||||
|  | ||||
| lemma int₁ : | ||||
|   ∫ x in (0)..(2 * π), log ‖circleMap 0 1 x - 1‖ = 0 := by | ||||
|  | ||||
|   simp_rw [logAffineHelper] | ||||
|   simp | ||||
|  | ||||
|   have : ∫ (x : ℝ) in (0)..2 * π, log (4 * sin (x / 2) ^ 2) =  2 * ∫ (x : ℝ) in (0)..π, log (4 * sin x ^ 2) := by | ||||
|   | ||||
| @@ -254,6 +254,19 @@ theorem intervalIntegral.integral_congr_volume | ||||
|   _ = 0 := volume_singleton | ||||
|  | ||||
|  | ||||
| theorem IntervalIntegrable.integral_congr_Ioo | ||||
|   {E : Type u_3} [NormedAddCommGroup E] [NormedSpace ℝ E] | ||||
|   {f g : ℝ → E} | ||||
|   {a b : ℝ} | ||||
|   (hab : a ≤ b) | ||||
|   (hfg : Set.EqOn f g (Set.Ioo a b)) : | ||||
|   IntervalIntegrable f volume a b ↔ IntervalIntegrable g volume a b := by | ||||
|  | ||||
|   rw [intervalIntegrable_iff_integrableOn_Ioo_of_le hab] | ||||
|   rw [MeasureTheory.integrableOn_congr_fun hfg measurableSet_Ioo] | ||||
|   rw [← intervalIntegrable_iff_integrableOn_Ioo_of_le hab] | ||||
|  | ||||
|  | ||||
|  | ||||
| lemma integral_log_sin₀ : ∫ (x : ℝ) in (0)..π, log (sin x) = 2 * ∫ (x : ℝ) in (0)..(π / 2), log (sin x) := by | ||||
|   rw [← intervalIntegral.integral_add_adjacent_intervals (a := 0) (b := π / 2) (c := π)] | ||||
|   | ||||
		Reference in New Issue
	
	Block a user
	 Stefan Kebekus
					Stefan Kebekus