Compare commits

...

4 Commits

Author SHA1 Message Date
Stefan Kebekus 38179d24c0 Working… 2024-08-13 16:26:55 +02:00
Stefan Kebekus 0cc0c81508 Update specialFunctions_Integrals.lean 2024-08-13 10:25:03 +02:00
Stefan Kebekus bc8fed96b0 Update specialFunctions_Integrals.lean 2024-08-13 09:46:30 +02:00
Stefan Kebekus 0cb1914b18 Add code of Gareth Ma 2024-08-13 08:42:47 +02:00
4 changed files with 164 additions and 57 deletions

View File

@ -7,8 +7,7 @@ theorem harmonic_meanValue
(ρ R : )
(hR : 0 < R)
(hρ : R < ρ)
(hf : ∀ x ∈ Metric.ball z ρ , HarmonicAt f x)
:
(hf : ∀ x ∈ Metric.ball z ρ , HarmonicAt f x) :
(∫ (x : ) in (0)..2 * Real.pi, f (circleMap z R x)) = 2 * Real.pi * f z
:= by

View File

@ -18,7 +18,7 @@ lemma xx
let A := hf z hz
let B := A.order
exact A.order
exact (A.order : )
else
exact 0

View File

@ -1,54 +0,0 @@
import Mathlib.Analysis.SpecialFunctions.Integrals
theorem intervalIntegral.intervalIntegrable_log'
{a : }
{b : }
{μ : MeasureTheory.Measure }
[MeasureTheory.IsLocallyFiniteMeasure μ]
(ha : 0 < a) :
IntervalIntegrable Real.log μ 0 a
:= by
sorry
theorem integral_log₀
{b : }
(hb : 0 < b) :
∫ (x : ) in (0)..b, Real.log x = b * (Real.log b - 1) := by
apply?
exact integral_log h
open Real Nat Set Finset
open scoped Real Interval
--variable {a b : } (n : )
namespace intervalIntegral
--open MeasureTheory
--variable {f : } {μ ν : Measure } [IsLocallyFiniteMeasure μ] (c d : )
#check integral_mul_deriv_eq_deriv_mul
theorem integral_log₁
(h : (0 : ) ∉ [[a, b]]) :
∫ x in a..b, log x = b * log b - a * log a - b + a := by
have h' : ∀ x ∈ [[a, b]], x ≠ 0 :=
fun x (hx : x ∈ [[a, b]]) => ne_of_mem_of_not_mem hx h
have heq : ∀ x ∈ [[a, b]], x * x⁻¹ = 1 :=
fun x hx => mul_inv_cancel (h' x hx)
let A := fun x hx => hasDerivAt_log (h' x hx)
convert integral_mul_deriv_eq_deriv_mul A (fun x _ => hasDerivAt_id x)
convert integral_mul_deriv_eq_deriv_mul A
(fun x _ => hasDerivAt_id x) (continuousOn_inv₀.mono <|
subset_compl_singleton_iff.mpr h).intervalIntegrable
continuousOn_const.intervalIntegrable using 1 <;>
simp [integral_congr heq, mul_comm, ← sub_add]

View File

@ -0,0 +1,162 @@
import Mathlib.Analysis.SpecialFunctions.Integrals
import Mathlib.Analysis.SpecialFunctions.Log.NegMulLog
import Mathlib.MeasureTheory.Integral.CircleIntegral
import Mathlib.MeasureTheory.Measure.Restrict
open scoped Interval Topology
open Real Filter MeasureTheory intervalIntegral
-- The following theorem was suggested by Gareth Ma on Zulip
example : IntervalIntegrable (log ∘ sin) volume 0 1 := by
have int_log : IntervalIntegrable log volume 0 1 := by sorry
apply IntervalIntegrable.mono_fun' (g := log)
exact int_log
-- AEStronglyMeasurable (log ∘ sin) (volume.restrict (Ι 0 1))
apply ContinuousOn.aestronglyMeasurable
apply ContinuousOn.comp (t := Ι 0 1)
apply ContinuousOn.mono (s := {0}ᶜ)
exact continuousOn_log
intro x hx
by_contra contra
simp at contra
rw [contra] at hx
rw [Set.left_mem_uIoc] at hx
linarith
exact continuousOn_sin
--
rw [Set.uIoc_of_le (zero_le_one' )]
exact fun x hx ↦ ⟨sin_pos_of_pos_of_le_one hx.1 hx.2, sin_le_one x⟩
--
exact measurableSet_uIoc
--
have : ∀ x ∈ (Ι 0 1), ‖(log ∘ sin) x‖ ≤ log x := by sorry
dsimp [EventuallyLE]
rw [MeasureTheory.ae_restrict_iff]
apply MeasureTheory.ae_of_all
exact this
--intro x
rw [MeasureTheory.ae_iff]
simp
rw [MeasureTheory.ae_iff]
simp
sorry
theorem logInt
{t : }
(ht : 0 < t) :
∫ x in (0 : )..t, log x = t * log t - t := by
rw [← integral_add_adjacent_intervals (b := 1)]
trans (-1) + (t * log t - t + 1)
· congr
· -- ∫ x in 0..1, log x = -1, same proof as before
rw [integral_eq_sub_of_hasDerivAt_of_tendsto (f := fun x ↦ x * log x - x) (fa := 0) (fb := -1)]
· simp
· simp
· intro x hx
norm_num at hx
convert (hasDerivAt_mul_log hx.left.ne.symm).sub (hasDerivAt_id x) using 1
norm_num
· rw [← neg_neg log]
apply IntervalIntegrable.neg
apply intervalIntegrable_deriv_of_nonneg (g := fun x ↦ -(x * log x - x))
· exact (continuous_mul_log.continuousOn.sub continuous_id.continuousOn).neg
· intro x hx
norm_num at hx
convert ((hasDerivAt_mul_log hx.left.ne.symm).sub (hasDerivAt_id x)).neg using 1
norm_num
· intro x hx
norm_num at hx
rw [Pi.neg_apply, Left.nonneg_neg_iff]
exact (log_nonpos_iff hx.left).mpr hx.right.le
· have := tendsto_log_mul_rpow_nhds_zero zero_lt_one
simp_rw [rpow_one, mul_comm] at this
-- tendsto_nhdsWithin_of_tendsto_nhds should be under Tendsto namespace
convert this.sub (tendsto_nhdsWithin_of_tendsto_nhds tendsto_id)
norm_num
· rw [(by simp : -1 = 1 * log 1 - 1)]
apply tendsto_nhdsWithin_of_tendsto_nhds
exact (continuousAt_id.mul (continuousAt_log one_ne_zero)).sub continuousAt_id
· -- ∫ x in 1..t, log x = t * log t + 1, just use integral_log_of_pos
rw [integral_log_of_pos zero_lt_one ht]
norm_num
· abel
· -- log is integrable on [[0, 1]]
rw [← neg_neg log]
apply IntervalIntegrable.neg
apply intervalIntegrable_deriv_of_nonneg (g := fun x ↦ -(x * log x - x))
· exact (continuous_mul_log.continuousOn.sub continuous_id.continuousOn).neg
· intro x hx
norm_num at hx
convert ((hasDerivAt_mul_log hx.left.ne.symm).sub (hasDerivAt_id x)).neg using 1
norm_num
· intro x hx
norm_num at hx
rw [Pi.neg_apply, Left.nonneg_neg_iff]
exact (log_nonpos_iff hx.left).mpr hx.right.le
· -- log is integrable on [[0, t]]
simp [Set.mem_uIcc, ht]
lemma int₁₁ : ∫ (x : ) in (0)..π, log (4 * sin x ^ 2) = 0 := by
sorry
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₁₁