Skip to content

Commit

Permalink
add derivative theorems for norm
Browse files Browse the repository at this point in the history
  • Loading branch information
lecopivo committed Aug 23, 2024
1 parent a223e5e commit 857a7cd
Show file tree
Hide file tree
Showing 4 changed files with 87 additions and 2 deletions.
1 change: 1 addition & 0 deletions SciLean/Analysis/Calculus/Monad/RevFDerivMonad.lean
Original file line number Diff line number Diff line change
Expand Up @@ -139,6 +139,7 @@ by
rw[revFDerivM_bind f (fun x => pure (g x))
hf (DifferentiableM_pure _ hg)]
simp[revFDerivM_pure g hg]
rfl

@[fun_trans]
theorem let_rule
Expand Down
6 changes: 6 additions & 0 deletions SciLean/Analysis/Calculus/RevFDeriv.lean
Original file line number Diff line number Diff line change
Expand Up @@ -41,6 +41,12 @@ variable
{E : ι → Type _} [∀ i, NormedAddCommGroup (E i)] [∀ i, AdjointSpace K (E i)] [∀ i, CompleteSpace (E i)]


-- Basic simp lemmas -----------------------------------------------------------
--------------------------------------------------------------------------------

@[simp, simp_core]
theorem revFDeriv_fst (f : X → Y) (x : X) : (revFDeriv K f x).1 = f x := by rfl

-- Basic lambda calculus rules -------------------------------------------------
--------------------------------------------------------------------------------

Expand Down
79 changes: 77 additions & 2 deletions SciLean/Analysis/SpecialFunctions/Norm2.lean
Original file line number Diff line number Diff line change
@@ -1,4 +1,5 @@
import SciLean.Analysis.SpecialFunctions.Inner
import SciLean.Analysis.SpecialFunctions.Pow
import SciLean.Meta.GenerateFunProp

namespace SciLean
Expand All @@ -9,6 +10,7 @@ variable
{R} [RealScalar R]
{X} [NormedAddCommGroup X] [NormedSpace R X]
{U} [NormedAddCommGroup U] [AdjointSpace R U] [CompleteSpace U]
{V} [NormedAddCommGroup V] [AdjointSpace R V] [CompleteSpace V]

def_fun_prop with_transitive : Differentiable R (fun u : U => ‖u‖₂²[R]) by
unfold Norm2.norm2; fun_prop [Norm2.norm2]
Expand All @@ -33,11 +35,84 @@ theorem Norm2.norm2.arg_a0.fwdFDeriv_rule :
unfold fwdFDeriv
fun_trans


@[fun_trans]
theorem Norm2.norm2.arg_a0.revFDeriv_rule :
revFDeriv R (fun x : U => ‖x‖₂²[R])
=
fun x => (‖x‖₂²[R], fun dr => (2 * dr)•x) := by
fun x =>
(‖x‖₂²[R], fun dy => (2 * dy) • x) := by
unfold revFDeriv
fun_trans [smul_smul]



-- ‖·‖₂ --------------------------------------------------------------------------------------------
----------------------------------------------------------------------------------------------------

def_fun_prop (x : U) (hx : x ≠ 0) : DifferentiableAt R (norm₂ R) x by
unfold norm₂; simp[Norm2.norm2]; fun_prop (disch:=aesop)

def_fun_prop (f : X → U) (hf : Differentiable R f) (hx' : ∀ x, f x ≠ 0) :
Differentiable R (fun x : X => ‖f x‖₂[R]) by intro x; fun_prop (disch:=aesop)


-- TODO: how can we streamline writing all of these theorems?

@[fun_trans]
theorem norm₂.arg_x.fderiv_rule_at (x : U) (hx : x ≠ 0) :
fderiv R (fun x : U => ‖x‖₂[R]) x
=
fun dx =>L[R] ⟪dx,x⟫[R] / ‖x‖₂[R] := by
unfold norm₂; simp[Norm2.norm2]; fun_trans (disch:=aesop)
ext dx; simp
rw [← AdjointSpace.inner_conj_symm]
simp; ring

@[fun_trans]
theorem norm₂.arg_x.fderiv_rule (f : X → U) (hf : Differentiable R f) (hf' : ∀ x, f x ≠ 0) :
fderiv R (fun x => ‖f x‖₂[R])
=
fun x =>
let y := f x
fun dx =>L[R]
let dy := fderiv R f x dx
⟪dy,y⟫[R] / ‖y‖₂[R] := by
funext x; fun_trans (disch:=aesop)

@[fun_trans]
theorem norm₂.arg_x.fwdFDeriv_rule_at (x : U) (hx : x ≠ 0) :
fwdFDeriv R (fun x : U => ‖x‖₂[R]) x
=
fun dx =>
let y := ‖x‖₂[R]
(y, ⟪dx,x⟫[R] / y) := by
unfold fwdFDeriv; fun_trans (disch:=assumption)

@[fun_trans]
theorem norm₂.arg_x.fwdFDeriv_rule (f : X → U) (hf : Differentiable R f) (hf' : ∀ x, f x ≠ 0) :
fwdFDeriv R (fun x => ‖f x‖₂[R])
=
fun x dx =>
let ydy := fwdFDeriv R f x dx
let yn := ‖ydy.1‖₂[R]
(yn, ⟪ydy.2,ydy.1⟫[R] / yn) := by
unfold fwdFDeriv; fun_trans (disch:=assumption)

@[fun_trans]
theorem norm₂.arg_x.revFDeriv_rule_at (x : U) (hx : x ≠ 0) :
revFDeriv R (fun x : U => ‖x‖₂[R]) x
=
let y := ‖x‖₂[R]
(y, fun dy => (y⁻¹ * dy) • x) := by
unfold revFDeriv; fun_trans (disch:=assumption) [smul_smul]


@[fun_trans]
theorem norm₂.arg_x.revFDeriv_rule (f : U → V) (hf : Differentiable R f) (hf' : ∀ x, f x ≠ 0) :
revFDeriv R (fun x : U => ‖f x‖₂[R])
=
fun x =>
let ydf := revFDeriv R f x
let y := ‖ydf.1‖₂[R]
(y, fun dy => ydf.2 ((y⁻¹ * dy) • ydf.1)) := by
funext x; fun_trans (disch:=apply hf')
3 changes: 3 additions & 0 deletions SciLean/Geometry/FrontierSpeed.lean
Original file line number Diff line number Diff line change
Expand Up @@ -8,6 +8,7 @@ import SciLean.Tactic.Autodiff
-- import SciLean.Tactic.GTrans

set_option linter.unusedVariables false
set_option deprecated.oldSectionVars true

open MeasureTheory Topology Filter FiniteDimensional

Expand Down Expand Up @@ -145,6 +146,8 @@ theorem closedBall₂.arg_xr.frontierSpeed_rule :
unfold closedBall₂
funext w dw x
conv => autodiff
set_option trace.Meta.Tactic.fun_trans true in
conv => autodiff
simp[smul_sub]

@[fun_trans]
Expand Down

0 comments on commit 857a7cd

Please sign in to comment.