nevanlinna/Nevanlinna/cauchyRiemann.lean

65 lines
2.0 KiB
Plaintext
Raw Normal View History

2024-04-28 21:17:45 +02:00
import Mathlib.Analysis.Calculus.LineDeriv.Basic
import Mathlib.Analysis.Complex.RealDeriv
2024-05-07 12:13:28 +02:00
import Nevanlinna.partialDeriv
2024-04-28 21:09:38 +02:00
variable {z : } {f : }
2024-05-06 09:01:43 +02:00
2024-04-29 10:19:21 +02:00
theorem CauchyRiemann₁ : (DifferentiableAt f z)
→ (fderiv f z) Complex.I = Complex.I * (fderiv f z) 1 := by
intro h
2024-04-29 19:48:37 +02:00
rw [DifferentiableAt.fderiv_restrictScalars h]
nth_rewrite 1 [← mul_one Complex.I]
exact ContinuousLinearMap.map_smul_of_tower (fderiv f z) Complex.I 1
2024-04-29 10:19:21 +02:00
2024-05-06 09:01:43 +02:00
2024-04-29 10:19:21 +02:00
theorem CauchyRiemann₂ : (DifferentiableAt f z)
→ lineDeriv f z Complex.I = Complex.I * lineDeriv f z 1 := by
intro h
2024-04-29 19:48:37 +02:00
rw [DifferentiableAt.lineDeriv_eq_fderiv (h.restrictScalars )]
rw [DifferentiableAt.lineDeriv_eq_fderiv (h.restrictScalars )]
2024-04-29 10:19:21 +02:00
exact CauchyRiemann₁ h
2024-05-06 09:01:43 +02:00
2024-04-29 10:19:21 +02:00
theorem CauchyRiemann₃ : (DifferentiableAt f z)
2024-04-29 14:19:58 +02:00
→ (lineDeriv (Complex.reCLM ∘ f) z 1 = lineDeriv (Complex.imCLM ∘ f) z Complex.I)
∧ (lineDeriv (Complex.reCLM ∘ f) z Complex.I = -lineDeriv (Complex.imCLM ∘ f) z 1)
:= by
2024-04-29 10:19:21 +02:00
intro h
2024-04-29 20:53:47 +02:00
have ContinuousLinearMap.comp_lineDeriv : ∀ w : , ∀ l : →L[] , lineDeriv (l ∘ f) z w = l ((fderiv f z) w) := by
intro w l
rw [DifferentiableAt.lineDeriv_eq_fderiv]
rw [fderiv.comp]
2024-04-29 14:19:58 +02:00
simp
2024-04-29 20:53:47 +02:00
fun_prop
exact h.restrictScalars
2024-04-29 20:53:47 +02:00
apply (ContinuousLinearMap.differentiableAt l).comp
exact h.restrictScalars
2024-04-29 10:19:21 +02:00
2024-04-29 20:53:47 +02:00
repeat
rw [ContinuousLinearMap.comp_lineDeriv]
rw [CauchyRiemann₁ h, Complex.I_mul]
simp
2024-05-07 12:13:28 +02:00
2024-05-17 09:19:24 +02:00
theorem CauchyRiemann₄ {F : Type*} [NormedAddCommGroup F] [NormedSpace F] {f : → F} : (Differentiable f)
2024-05-08 15:59:56 +02:00
→ partialDeriv Complex.I f = Complex.I • partialDeriv 1 f := by
2024-05-07 12:13:28 +02:00
intro h
2024-05-08 15:59:56 +02:00
unfold partialDeriv
2024-05-07 12:13:28 +02:00
conv =>
left
intro w
rw [DifferentiableAt.fderiv_restrictScalars (h w)]
simp
rw [← mul_one Complex.I]
rw [← smul_eq_mul]
rw [ContinuousLinearMap.map_smul_of_tower (fderiv f w) Complex.I 1]
conv =>
right
right
intro w
rw [DifferentiableAt.fderiv_restrictScalars (h w)]