nevanlinna/Nevanlinna/holomorphic.primitive.lean

140 lines
6.3 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_prod_of_hasFDerivWithinAt_off_countable₃
{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 h₁x
exact DifferentiableAt.hasFDerivAt ((h₁f.differentiable le_rfl) x)
--
rw [Set.diff_empty]
intro y h₁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
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 h₁x
exact DifferentiableAt.hasFDerivAt ((h₁f.differentiable le_rfl) x)
--
rw [Set.diff_empty]
intro y h₁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