diff --git a/Mathlib.lean b/Mathlib.lean index c115ab4aae344..86edc0be779ba 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -395,6 +395,7 @@ import Mathlib.Analysis.BoxIntegral.Partition.SubboxInduction import Mathlib.Analysis.BoxIntegral.Partition.Tagged import Mathlib.Analysis.Calculus.FDeriv.Add import Mathlib.Analysis.Calculus.FDeriv.Basic +import Mathlib.Analysis.Calculus.FDeriv.Bilinear import Mathlib.Analysis.Calculus.FDeriv.Comp import Mathlib.Analysis.Calculus.FDeriv.Equiv import Mathlib.Analysis.Calculus.FDeriv.Linear diff --git a/Mathlib/Analysis/Calculus/FDeriv/Bilinear.lean b/Mathlib/Analysis/Calculus/FDeriv/Bilinear.lean new file mode 100644 index 0000000000000..a976343bf3977 --- /dev/null +++ b/Mathlib/Analysis/Calculus/FDeriv/Bilinear.lean @@ -0,0 +1,160 @@ +/- +Copyright (c) 2019 Jeremy Avigad. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Jeremy Avigad, SΓ©bastien GouΓ«zel, Yury Kudryashov + +! This file was ported from Lean 3 source module analysis.calculus.fderiv.bilinear +! leanprover-community/mathlib commit e3fb84046afd187b710170887195d50bada934ee +! Please do not edit these lines, except to modify the commit id +! if you have ported upstream changes. +-/ +import Mathlib.Analysis.Calculus.FDeriv.Prod + +/-! +# The derivative of bounded bilinear maps + +For detailed documentation of the FrΓ©chet derivative, +see the module docstring of `Analysis/Calculus/Fderiv/Basic.lean`. + +This file contains the usual formulas (and existence assertions) for the derivative of +bounded bilinear maps. +-/ + + +open Filter Asymptotics ContinuousLinearMap Set Metric +open Topology Classical NNReal Asymptotics ENNReal + +noncomputable section + +section + +variable {π•œ : Type _} [NontriviallyNormedField π•œ] + +variable {E : Type _} [NormedAddCommGroup E] [NormedSpace π•œ E] + +variable {F : Type _} [NormedAddCommGroup F] [NormedSpace π•œ F] + +variable {G : Type _} [NormedAddCommGroup G] [NormedSpace π•œ G] + +variable {G' : Type _} [NormedAddCommGroup G'] [NormedSpace π•œ G'] + +variable {f fβ‚€ f₁ g : E β†’ F} + +variable {f' fβ‚€' f₁' g' : E β†’L[π•œ] F} + +variable (e : E β†’L[π•œ] F) + +variable {x : E} + +variable {s t : Set E} + +variable {L L₁ Lβ‚‚ : Filter E} + +section BilinearMap + +/-! ### Derivative of a bounded bilinear map -/ + +variable {b : E Γ— F β†’ G} {u : Set (E Γ— F)} + +open NormedField + +-- porting note: todo: rewrite/golf using analytic functions? +theorem IsBoundedBilinearMap.hasStrictFDerivAt (h : IsBoundedBilinearMap π•œ b) (p : E Γ— F) : + HasStrictFDerivAt b (h.deriv p) p := by + simp only [HasStrictFDerivAt] + simp only [← map_add_left_nhds_zero (p, p), isLittleO_map] + set T := (E Γ— F) Γ— E Γ— F + calc + _ = fun x ↦ h.deriv (x.1 - x.2) (x.2.1, x.1.2) := by + ext ⟨⟨x₁, yβ‚βŸ©, ⟨xβ‚‚, yβ‚‚βŸ©βŸ© + rcases p with ⟨x, y⟩ + simp [h.add_left, h.add_right, h.deriv_apply, h.map_sub_left, h.map_sub_right] + abel + -- _ =O[𝓝 (0 : T)] fun x ↦ β€–x.1 - x.2β€– * β€–(x.2.1, x.1.2)β€– := + -- h.toContinuousLinearMap.derivβ‚‚.isBoundedBilinearMap.isBigO_comp + -- _ = o[𝓝 0] fun x ↦ β€–x.1 - x.2β€– * 1 := _ + _ =o[𝓝 (0 : T)] fun x ↦ x.1 - x.2 := by + -- TODO : add 2 `calc` steps instead of the next 3 lines + refine h.toContinuousLinearMap.derivβ‚‚.isBoundedBilinearMap.isBigO_comp.trans_isLittleO ?_ + suffices : (fun x : T ↦ β€–x.1 - x.2β€– * β€–(x.2.1, x.1.2)β€–) =o[𝓝 0] fun x ↦ β€–x.1 - x.2β€– * 1 + Β· simpa only [mul_one, isLittleO_norm_right] using this + refine (isBigO_refl _ _).mul_isLittleO ((isLittleO_one_iff _).2 ?_) + -- TODO: `continuity` fails + exact (continuous_snd.fst.prod_mk continuous_fst.snd).norm.tendsto' _ _ (by simp) + _ = _ := by simp [(Β· ∘ Β·)] +#align is_bounded_bilinear_map.has_strict_fderiv_at IsBoundedBilinearMap.hasStrictFDerivAt + +theorem IsBoundedBilinearMap.hasFDerivAt (h : IsBoundedBilinearMap π•œ b) (p : E Γ— F) : + HasFDerivAt b (h.deriv p) p := + (h.hasStrictFDerivAt p).hasFDerivAt +#align is_bounded_bilinear_map.has_fderiv_at IsBoundedBilinearMap.hasFDerivAt + +theorem IsBoundedBilinearMap.hasFDerivWithinAt (h : IsBoundedBilinearMap π•œ b) (p : E Γ— F) : + HasFDerivWithinAt b (h.deriv p) u p := + (h.hasFDerivAt p).hasFDerivWithinAt +#align is_bounded_bilinear_map.has_fderiv_within_at IsBoundedBilinearMap.hasFDerivWithinAt + +theorem IsBoundedBilinearMap.differentiableAt (h : IsBoundedBilinearMap π•œ b) (p : E Γ— F) : + DifferentiableAt π•œ b p := + (h.hasFDerivAt p).differentiableAt +#align is_bounded_bilinear_map.differentiable_at IsBoundedBilinearMap.differentiableAt + +theorem IsBoundedBilinearMap.differentiableWithinAt (h : IsBoundedBilinearMap π•œ b) (p : E Γ— F) : + DifferentiableWithinAt π•œ b u p := + (h.differentiableAt p).differentiableWithinAt +#align is_bounded_bilinear_map.differentiable_within_at IsBoundedBilinearMap.differentiableWithinAt + +theorem IsBoundedBilinearMap.fderiv (h : IsBoundedBilinearMap π•œ b) (p : E Γ— F) : + fderiv π•œ b p = h.deriv p := + HasFDerivAt.fderiv (h.hasFDerivAt p) +#align is_bounded_bilinear_map.fderiv IsBoundedBilinearMap.fderiv + +theorem IsBoundedBilinearMap.fderivWithin (h : IsBoundedBilinearMap π•œ b) (p : E Γ— F) + (hxs : UniqueDiffWithinAt π•œ u p) : fderivWithin π•œ b u p = h.deriv p := by + rw [DifferentiableAt.fderivWithin (h.differentiableAt p) hxs] + exact h.fderiv p +#align is_bounded_bilinear_map.fderiv_within IsBoundedBilinearMap.fderivWithin + +theorem IsBoundedBilinearMap.differentiable (h : IsBoundedBilinearMap π•œ b) : Differentiable π•œ b := + fun x => h.differentiableAt x +#align is_bounded_bilinear_map.differentiable IsBoundedBilinearMap.differentiable + +theorem IsBoundedBilinearMap.differentiableOn (h : IsBoundedBilinearMap π•œ b) : + DifferentiableOn π•œ b u := + h.differentiable.differentiableOn +#align is_bounded_bilinear_map.differentiable_on IsBoundedBilinearMap.differentiableOn + +variable (B : E β†’L[π•œ] F β†’L[π•œ] G) + +theorem ContinuousLinearMap.hasFDerivWithinAt_of_bilinear {f : G' β†’ E} {g : G' β†’ F} + {f' : G' β†’L[π•œ] E} {g' : G' β†’L[π•œ] F} {x : G'} {s : Set G'} (hf : HasFDerivWithinAt f f' s x) + (hg : HasFDerivWithinAt g g' s x) : + HasFDerivWithinAt (fun y => B (f y) (g y)) + (B.precompR G' (f x) g' + B.precompL G' f' (g x)) s x := + (B.isBoundedBilinearMap.hasFDerivAt (f x, g x)).comp_hasFDerivWithinAt x (hf.prod hg) +#align continuous_linear_map.has_fderiv_within_at_of_bilinear ContinuousLinearMap.hasFDerivWithinAt_of_bilinear + +theorem ContinuousLinearMap.hasFDerivAt_of_bilinear {f : G' β†’ E} {g : G' β†’ F} {f' : G' β†’L[π•œ] E} + {g' : G' β†’L[π•œ] F} {x : G'} (hf : HasFDerivAt f f' x) (hg : HasFDerivAt g g' x) : + HasFDerivAt (fun y => B (f y) (g y)) (B.precompR G' (f x) g' + B.precompL G' f' (g x)) x := + (B.isBoundedBilinearMap.hasFDerivAt (f x, g x)).comp x (hf.prod hg) +#align continuous_linear_map.has_fderiv_at_of_bilinear ContinuousLinearMap.hasFDerivAt_of_bilinear + +theorem ContinuousLinearMap.fderivWithin_of_bilinear {f : G' β†’ E} {g : G' β†’ F} {x : G'} {s : Set G'} + (hf : DifferentiableWithinAt π•œ f s x) (hg : DifferentiableWithinAt π•œ g s x) + (hs : UniqueDiffWithinAt π•œ s x) : + fderivWithin π•œ (fun y => B (f y) (g y)) s x = + B.precompR G' (f x) (fderivWithin π•œ g s x) + B.precompL G' (fderivWithin π•œ f s x) (g x) := + (B.hasFDerivWithinAt_of_bilinear hf.hasFDerivWithinAt hg.hasFDerivWithinAt).fderivWithin hs +#align continuous_linear_map.fderiv_within_of_bilinear ContinuousLinearMap.fderivWithin_of_bilinear + +theorem ContinuousLinearMap.fderiv_of_bilinear {f : G' β†’ E} {g : G' β†’ F} {x : G'} + (hf : DifferentiableAt π•œ f x) (hg : DifferentiableAt π•œ g x) : + fderiv π•œ (fun y => B (f y) (g y)) x = + B.precompR G' (f x) (fderiv π•œ g x) + B.precompL G' (fderiv π•œ f x) (g x) := + (B.hasFDerivAt_of_bilinear hf.hasFDerivAt hg.hasFDerivAt).fderiv +#align continuous_linear_map.fderiv_of_bilinear ContinuousLinearMap.fderiv_of_bilinear + +end BilinearMap + +end