nevanlinna/Nevanlinna/holomorphic.primitive.lean

269 lines
9.6 KiB
Plaintext
Raw Blame History

This file contains ambiguous Unicode characters

This file contains Unicode characters that might be confused with other characters. If you think that this is intentional, you can safely ignore this warning. Use the Escape button to reveal them.

import Mathlib.Analysis.Complex.TaylorSeries
import Mathlib.MeasureTheory.Integral.DivergenceTheorem
import Mathlib.MeasureTheory.Function.LocallyIntegrable
import Nevanlinna.cauchyRiemann
import Nevanlinna.partialDeriv
theorem MeasureTheory.integral2_divergence₃
{E : Type u} [NormedAddCommGroup E] [NormedSpace E] [CompleteSpace E]
(f g : × → E)
(h₁f : ContDiff 1 f)
(h₁g : ContDiff 1 g)
(a₁ : )
(a₂ : )
(b₁ : )
(b₂ : ) :
∫ (x : ) in a₁..b₁, ∫ (y : ) in a₂..b₂, ((fderiv f) (x, y)) (1, 0) + ((fderiv g) (x, y)) (0, 1) = (((∫ (x : ) in a₁..b₁, g (x, b₂)) - ∫ (x : ) in a₁..b₁, g (x, a₂)) + ∫ (y : ) in a₂..b₂, f (b₁, y)) - ∫ (y : ) in a₂..b₂, f (a₁, y) := by
apply integral2_divergence_prod_of_hasFDerivWithinAt_off_countable f g (fderiv f) (fderiv g) a₁ a₂ b₁ b₂ ∅
exact Set.countable_empty
-- ContinuousOn f (Set.uIcc a₁ b₁ ×ˢ Set.uIcc a₂ b₂)
exact h₁f.continuous.continuousOn
--
exact h₁g.continuous.continuousOn
--
rw [Set.diff_empty]
intro x _
exact DifferentiableAt.hasFDerivAt ((h₁f.differentiable le_rfl) x)
--
rw [Set.diff_empty]
intro y _
exact DifferentiableAt.hasFDerivAt ((h₁g.differentiable le_rfl) y)
--
apply ContinuousOn.integrableOn_compact
apply IsCompact.prod
exact isCompact_uIcc
exact isCompact_uIcc
apply ContinuousOn.add
apply Continuous.continuousOn
exact Continuous.clm_apply (ContDiff.continuous_fderiv h₁f le_rfl) continuous_const
apply Continuous.continuousOn
exact Continuous.clm_apply (ContDiff.continuous_fderiv h₁g le_rfl) continuous_const
theorem integral_divergence₄
{E : Type u} [NormedAddCommGroup E] [NormedSpace E] [CompleteSpace E]
(f g : → E)
(h₁f : ContDiff 1 f)
(h₁g : ContDiff 1 g)
(a₁ : )
(a₂ : )
(b₁ : )
(b₂ : ) :
∫ (x : ) in a₁..b₁, ∫ (y : ) in a₂..b₂, ((fderiv f) ⟨x, y⟩ ) 1 + ((fderiv g) ⟨x, y⟩) Complex.I = (((∫ (x : ) in a₁..b₁, g ⟨x, b₂⟩) - ∫ (x : ) in a₁..b₁, g ⟨x, a₂⟩) + ∫ (y : ) in a₂..b₂, f ⟨b₁, y⟩) - ∫ (y : ) in a₂..b₂, f ⟨a₁, y⟩ := by
let fr : × → E := f ∘ Complex.equivRealProdCLM.symm
let gr : × → E := g ∘ Complex.equivRealProdCLM.symm
have sfr {x y : } : f { re := x, im := y } = fr (x, y) := by exact rfl
have sgr {x y : } : g { re := x, im := y } = gr (x, y) := by exact rfl
repeat (conv in f { re := _, im := _ } => rw [sfr])
repeat (conv in g { re := _, im := _ } => rw [sgr])
have sfr' {x y : } {z : } : (fderiv f { re := x, im := y }) z = fderiv fr (x, y) (Complex.equivRealProdCLM z) := by
rw [fderiv.comp]
rw [Complex.equivRealProdCLM.symm.fderiv]
tauto
apply Differentiable.differentiableAt
exact h₁f.differentiable le_rfl
exact Complex.equivRealProdCLM.symm.differentiableAt
conv in ⇑(fderiv f { re := _, im := _ }) _ => rw [sfr']
have sgr' {x y : } {z : } : (fderiv g { re := x, im := y }) z = fderiv gr (x, y) (Complex.equivRealProdCLM z) := by
rw [fderiv.comp]
rw [Complex.equivRealProdCLM.symm.fderiv]
tauto
apply Differentiable.differentiableAt
exact h₁g.differentiable le_rfl
exact Complex.equivRealProdCLM.symm.differentiableAt
conv in ⇑(fderiv g { re := _, im := _ }) _ => rw [sgr']
apply MeasureTheory.integral2_divergence₃ fr gr _ _ a₁ a₂ b₁ b₂
-- ContDiff 1 fr
exact (ContinuousLinearEquiv.contDiff_comp_iff (ContinuousLinearEquiv.symm Complex.equivRealProdCLM)).mpr h₁f
-- ContDiff 1 gr
exact (ContinuousLinearEquiv.contDiff_comp_iff (ContinuousLinearEquiv.symm Complex.equivRealProdCLM)).mpr h₁g
theorem integral_divergence₅
{E : Type u} [NormedAddCommGroup E] [NormedSpace E] [CompleteSpace E]
(F : → E)
(hF : Differentiable F)
(lowerLeft upperRight : ) :
(∫ (x : ) in lowerLeft.re..upperRight.re, F ⟨x, lowerLeft.im⟩) + Complex.I • ∫ (x : ) in lowerLeft.im..upperRight.im, F ⟨upperRight.re, x⟩ =
(∫ (x : ) in lowerLeft.re..upperRight.re, F ⟨x, upperRight.im⟩) + Complex.I • ∫ (x : ) in lowerLeft.im..upperRight.im, F ⟨lowerLeft.re, x⟩ := by
let h₁f : ContDiff 1 F := (hF.contDiff : ContDiff 1 F).restrict_scalars
let h₁g : ContDiff 1 (-Complex.I • F) := by
have : -Complex.I • F = fun x ↦ -Complex.I • F x := by rfl
rw [this]
apply ContDiff.comp
exact contDiff_const_smul _
exact h₁f
let A := integral_divergence₄ (-Complex.I • F) F h₁g h₁f lowerLeft.re upperRight.im upperRight.re lowerLeft.im
have {z : } : fderiv F z Complex.I = partialDeriv _ F z := by rfl
conv at A in (fderiv F _) _ => rw [this]
have {z : } : fderiv (-Complex.I • F) z 1 = partialDeriv _ (-Complex.I • F) z := by rfl
conv at A in (fderiv (-Complex.I • F) _) _ => rw [this]
conv at A =>
left
arg 1
intro x
arg 1
intro y
rw [CauchyRiemann₄ hF]
rw [partialDeriv_smul'₂]
simp
simp at A
have {t₁ t₂ t₃ t₄ : E} : 0 = (t₁ - t₂) + t₃ + t₄ → t₁ + t₃ = t₂ - t₄ := by
intro hyp
calc
t₁ + t₃ = t₁ + t₃ - 0 := by rw [sub_zero (t₁ + t₃)]
_ = t₁ + t₃ - (t₁ - t₂ + t₃ + t₄) := by rw [hyp]
_ = t₂ - t₄ := by abel
let B := this A
repeat
rw [intervalIntegral.integral_symm lowerLeft.im upperRight.im] at B
simp at B
exact B
noncomputable def primitive
{E : Type u} [NormedAddCommGroup E] [NormedSpace E] [CompleteSpace E] :
→ ( → E) → ( → E) := by
intro z₀
intro f
exact fun z ↦ (∫ (x : ) in z₀.re..z.re, f ⟨x, z₀.im⟩) + Complex.I • ∫ (x : ) in z₀.im..z.im, f ⟨z.re, x⟩
theorem primitive_zeroAtBasepoint
{E : Type u} [NormedAddCommGroup E] [NormedSpace E] [CompleteSpace E]
(f : → E)
(z₀ : ) :
(primitive z₀ f) z₀ = 0 := by
unfold primitive
simp
theorem primitive_lem1
{E : Type u} [NormedAddCommGroup E] [NormedSpace E] [CompleteSpace E] [IsScalarTower E]
(v : E) :
HasDerivAt (primitive 0 (fun _ ↦ v)) v 0 := by
unfold primitive
simp
have : (fun (z : ) => z.re • v + Complex.I • z.im • v) = (fun (y : ) => ((fun w ↦ w) y) • v) := by
funext z
rw [smul_comm]
rw [← smul_assoc]
simp
have : z.re • v = (z.re : ) • v := by exact rfl
rw [this, ← add_smul]
simp
rw [this]
have hc : HasDerivAt (fun (w : ) ↦ w) 1 0 := by
apply hasDerivAt_id'
nth_rewrite 2 [← (one_smul v)]
exact HasDerivAt.smul_const hc v
theorem primitive_fderivAtBasepoint
{E : Type u} [NormedAddCommGroup E] [NormedSpace E] [CompleteSpace E]
(f : → E)
(hf : Continuous f) :
HasDerivAt (primitive 0 f) (f 0) 0 := by
unfold primitive
simp
apply hasDerivAt_iff_isLittleO.2
simp
rw [Asymptotics.isLittleO_iff]
intro c hc
have {z : } {e : E} : z • e = (∫ (x : ) in (0)..(z.re), e) + Complex.I • ∫ (x : ) in (0)..(z.im), e:= by
simp
rw [smul_comm]
rw [← smul_assoc]
simp
have : z.re • e = (z.re : ) • e := by exact rfl
rw [this, ← add_smul]
simp
conv =>
left
intro x
left
arg 1
arg 2
rw [this]
have {A B C D :E} : (A + B) - (C + D) = (A - C) + (B - D) := by
abel
have t₀ {r : } : IntervalIntegrable (fun x => f { re := x, im := 0 }) MeasureTheory.volume 0 r := by sorry
have t₁ {r : } :IntervalIntegrable (fun x => f 0) MeasureTheory.volume 0 r := by sorry
have t₂ {a b : }: IntervalIntegrable (fun x_1 => f { re := a, im := x_1 }) MeasureTheory.volume 0 b := by sorry
have t₃ {a : } : IntervalIntegrable (fun x => f 0) MeasureTheory.volume 0 a := by sorry
conv =>
left
intro x
left
arg 1
rw [this]
rw [← smul_sub]
rw [← intervalIntegral.integral_sub t₀ t₁]
rw [← intervalIntegral.integral_sub t₂ t₃]
rw [Filter.eventually_iff_exists_mem]
let s := f⁻¹' Metric.ball (f 0) c
use s
constructor
· apply IsOpen.mem_nhds
apply IsOpen.preimage hf
exact Metric.isOpen_ball
apply Set.mem_preimage.mpr
exact Metric.mem_ball_self hc
· intro y hy
have : ‖(∫ (x : ) in (0)..(y.re), f { re := x, im := 0 } - f 0)‖ ≤ c * |y.re| := by
let A := intervalIntegral.norm_integral_le_of_norm_le_const_ae
sorry
calc ‖(∫ (x : ) in (0)..(y.re), f { re := x, im := 0 } - f 0) + Complex.I • ∫ (x : ) in (0)..(y.im), f { re := y.re, im := x } - f 0‖
_ ≤ ‖(∫ (x : ) in (0)..(y.re), f { re := x, im := 0 } - f 0)‖ + ‖Complex.I • ∫ (x : ) in (0)..(y.im), f { re := y.re, im := x } - f 0‖ := by apply norm_add_le
_ ≤ ‖(∫ (x : ) in (0)..(y.re), f { re := x, im := 0 } - f 0)‖ + ‖∫ (x : ) in (0)..(y.im), f { re := y.re, im := x } - f 0‖ := by
simp
rw [norm_smul]
simp
_ ≤ |(∫ (x : ) in (0)..(y.re), ‖f { re := x, im := 0 } - f 0‖)| + |∫ (x : ) in (0)..(y.im), ‖f { re := y.re, im := x } - f 0‖| := by
apply add_le_add
apply intervalIntegral.norm_integral_le_abs_integral_norm
apply intervalIntegral.norm_integral_le_abs_integral_norm
_ ≤
sorry
theorem primitive_additivity
{E : Type u} [NormedAddCommGroup E] [NormedSpace E] [CompleteSpace E]
(f : → E)
(hf : Differentiable f)
(z₀ z₁ : ) :
(primitive z₁ f) = (primitive z₀ f) - (fun z ↦ primitive z₀ f z₁) := by
sorry