nevanlinna/Nevanlinna/holomorphic.primitive.lean

190 lines
8.0 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
/-
variable {E : Type*} [NormedAddCommGroup E] [NormedSpace E]
variable {F : Type*} [NormedAddCommGroup F] [NormedSpace F] [CompleteSpace F]
variable {G : Type*} [NormedAddCommGroup G] [NormedSpace G] [CompleteSpace G]
noncomputable def Complex.primitive
(f : → F) : → F :=
fun z ↦ ∫ t : in (0)..1, z • f (t * z)
-/
theorem MeasureTheory.integral2_divergence_prod_of_hasFDerivWithinAt_off_countable₁
{E : Type u} [NormedAddCommGroup E] [NormedSpace E] [CompleteSpace E]
(f : × → E)
(g : × → E)
(f' : × × →L[] E)
(g' : × × →L[] E)
(a₁ : )
(a₂ : )
(b₁ : )
(b₂ : )
(s : Set ( × ))
(hs : s.Countable)
(Hcf : ContinuousOn f (Set.uIcc a₁ b₁ ×ˢ Set.uIcc a₂ b₂))
(Hcg : ContinuousOn g (Set.uIcc a₁ b₁ ×ˢ Set.uIcc a₂ b₂))
(Hdf : ∀ x ∈ Set.Ioo (min a₁ b₁) (max a₁ b₁) ×ˢ Set.Ioo (min a₂ b₂) (max a₂ b₂) \ s, HasFDerivAt f (f' x) x)
(Hdg : ∀ x ∈ Set.Ioo (min a₁ b₁) (max a₁ b₁) ×ˢ Set.Ioo (min a₂ b₂) (max a₂ b₂) \ s, HasFDerivAt g (g' x) x)
(Hi : MeasureTheory.IntegrableOn (fun (x : × ) => (f' x) (1, 0) + (g' x) (0, 1)) (Set.uIcc a₁ b₁ ×ˢ Set.uIcc a₂ b₂) MeasureTheory.volume) :
∫ (x : ) in a₁..b₁, ∫ (y : ) in a₂..b₂, (f' (x, y)) (1, 0) + (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
exact
integral2_divergence_prod_of_hasFDerivWithinAt_off_countable f g f' g' a₁ a₂ b₁ b₂ s hs Hcf Hcg
Hdf Hdg Hi
theorem MeasureTheory.integral2_divergence_prod_of_hasFDerivWithinAt_off_countable₂
{E : Type u} [NormedAddCommGroup E] [NormedSpace E] [CompleteSpace E]
(f : × → E)
(g : × → E)
(f' : × × →L[] E)
(g' : × × →L[] E)
(a₁ : )
(a₂ : )
(b₁ : )
(b₂ : )
(Hcf : ContinuousOn f (Set.uIcc a₁ b₁ ×ˢ Set.uIcc a₂ b₂))
(Hcg : ContinuousOn g (Set.uIcc a₁ b₁ ×ˢ Set.uIcc a₂ b₂))
(Hdf : ∀ x ∈ Set.Ioo (min a₁ b₁) (max a₁ b₁) ×ˢ Set.Ioo (min a₂ b₂) (max a₂ b₂), HasFDerivAt f (f' x) x)
(Hdg : ∀ x ∈ Set.Ioo (min a₁ b₁) (max a₁ b₁) ×ˢ Set.Ioo (min a₂ b₂) (max a₂ b₂), HasFDerivAt g (g' x) x)
(Hi : MeasureTheory.IntegrableOn (fun (x : × ) => (f' x) (1, 0) + (g' x) (0, 1)) (Set.uIcc a₁ b₁ ×ˢ Set.uIcc a₂ b₂) MeasureTheory.volume) :
∫ (x : ) in a₁..b₁, ∫ (y : ) in a₂..b₂, (f' (x, y)) (1, 0) + (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 f' g' a₁ a₂ b₁ b₂ ∅
exact Set.countable_empty
assumption
assumption
rwa [Set.diff_empty]
rwa [Set.diff_empty]
assumption
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 : )
(hF : Differentiable F)
(a₁ : )
(a₂ : )
(b₁ : )
(b₂ : ) :
(∫ (x : ) in a₁..b₁, Complex.I • F { re := x, im := b₂ }) +
(∫ (y : ) in a₂..b₂, F { re := b₁, im := y })
=
(∫ (x : ) in a₁..b₁, Complex.I • F { re := x, im := a₂ }) +
(∫ (y : ) in a₂..b₂, F { re := a₁, im := y }) := by
let f := F
let h₁f : ContDiff 1 f := by sorry
let g := Complex.I • F
let h₁g : ContDiff 1 g := by sorry
let A := integral_divergence₄ f g h₁f h₁g a₁ a₂ b₁ b₂
have {z : } : fderiv f z 1 = partialDeriv 1 f z := by rfl
conv at A in (fderiv f _) 1 => rw [this]
have {z : } : fderiv g z Complex.I = partialDeriv Complex.I g z := by rfl
conv at A in (fderiv g _) Complex.I => rw [this]
have : Differentiable g := by sorry
conv at A =>
left
arg 1
intro x
arg 1
intro y
rw [CauchyRiemann₄ this]
rw [partialDeriv_smul'₂]
rw [← smul_assoc]
simp
simp at A
sorry