Skip to content

Commit

Permalink
feat(analysis/calculus/cont_diff): extra lemmas about cont_diff_withi…
Browse files Browse the repository at this point in the history
…n_at (#15309)

* Also some lemmas about `fderiv_within` and `nhds_within`.
* From the sphere eversion project
  • Loading branch information
fpvandoorn committed Jul 13, 2022
1 parent 51f76d0 commit 25f60b4
Show file tree
Hide file tree
Showing 3 changed files with 97 additions and 7 deletions.
63 changes: 63 additions & 0 deletions src/analysis/calculus/cont_diff.lean
Expand Up @@ -452,6 +452,12 @@ lemma cont_diff_within_at.congr_of_eventually_eq
⟨{x ∈ u | f₁ x = f x}, filter.inter_mem hu (mem_nhds_within_insert.2 ⟨hx, h₁⟩), p,
(H.mono (sep_subset _ _)).congr (λ _, and.right)⟩

lemma cont_diff_within_at.congr_of_eventually_eq_insert
(h : cont_diff_within_at 𝕜 n f s x) (h₁ : f₁ =ᶠ[𝓝[insert x s] x] f) :
cont_diff_within_at 𝕜 n f₁ s x :=
h.congr_of_eventually_eq (nhds_within_mono x (subset_insert x s) h₁)
(mem_of_mem_nhds_within (mem_insert x s) h₁ : _)

lemma cont_diff_within_at.congr_of_eventually_eq'
(h : cont_diff_within_at 𝕜 n f s x) (h₁ : f₁ =ᶠ[𝓝[s] x] f) (hx : x ∈ s) :
cont_diff_within_at 𝕜 n f₁ s x :=
Expand Down Expand Up @@ -569,6 +575,38 @@ begin
rw [snoc_last, init_snoc] } } } }
end

/-- One direction of `cont_diff_within_at_succ_iff_has_fderiv_within_at`, but where all derivatives
are taken within the same set. -/
lemma cont_diff_within_at.has_fderiv_within_at_nhds {n : ℕ}
(hf : cont_diff_within_at 𝕜 (n + 1 : ℕ) f s x) :
∃ u ∈ 𝓝[insert x s] x, u ⊆ insert x s ∧ ∃ f' : E → E →L[𝕜] F,
(∀ x ∈ u, has_fderiv_within_at f (f' x) s x) ∧ cont_diff_within_at 𝕜 n f' s x :=
begin
obtain ⟨u, hu, f', huf', hf'⟩ := cont_diff_within_at_succ_iff_has_fderiv_within_at.mp hf,
obtain ⟨w, hw, hxw, hwu⟩ := mem_nhds_within.mp hu,
rw [inter_comm] at hwu,
refine ⟨insert x s ∩ w, inter_mem_nhds_within _ (hw.mem_nhds hxw), inter_subset_left _ _,
f', λ y hy, _, _⟩,
{ refine ((huf' y $ hwu hy).mono hwu).mono_of_mem _,
refine mem_of_superset _ (inter_subset_inter_left _ (subset_insert _ _)),
refine inter_mem_nhds_within _ (hw.mem_nhds hy.2) },
{ exact hf'.mono_of_mem (nhds_within_mono _ (subset_insert _ _) hu) }
end

/-- A version of `cont_diff_within_at_succ_iff_has_fderiv_within_at` where all derivatives
are taken within the same set. This lemma assumes `x ∈ s`. -/
lemma cont_diff_within_at_succ_iff_has_fderiv_within_at_of_mem {n : ℕ} (hx : x ∈ s) :
cont_diff_within_at 𝕜 (n + 1 : ℕ) f s x
↔ ∃ u ∈ 𝓝[s] x, u ⊆ s ∧ ∃ f' : E → E →L[𝕜] F,
(∀ x ∈ u, has_fderiv_within_at f (f' x) s x) ∧ cont_diff_within_at 𝕜 n f' s x :=
begin
split,
{ intro hf, simpa only [insert_eq_of_mem hx] using hf.has_fderiv_within_at_nhds },
rw [cont_diff_within_at_succ_iff_has_fderiv_within_at, insert_eq_of_mem hx],
rintro ⟨u, hu, hus, f', huf', hf'⟩,
exact ⟨u, hu, f', λ y hy, (huf' y hy).mono hus, hf'.mono hus⟩
end

/-! ### Smooth functions within a set -/

variable (𝕜)
Expand Down Expand Up @@ -1093,6 +1131,31 @@ lemma cont_diff_on.continuous_on_fderiv_of_open
continuous_on (λ x, fderiv 𝕜 f x) s :=
((cont_diff_on_succ_iff_fderiv_of_open hs).1 (h.of_le hn)).2.continuous_on

lemma cont_diff_within_at.fderiv_within'
(hf : cont_diff_within_at 𝕜 n f s x) (hs : ∀ᶠ y in 𝓝[insert x s] x, unique_diff_within_at 𝕜 s y)
(hmn : m + 1 ≤ n) :
cont_diff_within_at 𝕜 m (fderiv_within 𝕜 f s) s x :=
begin
have : ∀ k : ℕ, (k + 1 : with_top ℕ) ≤ n → cont_diff_within_at 𝕜 k (fderiv_within 𝕜 f s) s x,
{ intros k hkn,
obtain ⟨v, hv, -, f', hvf', hf'⟩ := (hf.of_le hkn).has_fderiv_within_at_nhds,
apply hf'.congr_of_eventually_eq_insert,
filter_upwards [hv, hs],
exact λ y hy h2y, (hvf' y hy).fderiv_within h2y },
induction m using with_top.rec_top_coe,
{ obtain rfl := eq_top_iff.mpr hmn,
rw [cont_diff_within_at_top],
exact λ m, this m le_top },
exact this m hmn
end

lemma cont_diff_within_at.fderiv_within
(hf : cont_diff_within_at 𝕜 n f s x) (hs : unique_diff_on 𝕜 s)
(hmn : (m + 1 : with_top ℕ) ≤ n) (hxs : x ∈ s) :
cont_diff_within_at 𝕜 m (fderiv_within 𝕜 f s) s x :=
hf.fderiv_within' (by { rw [insert_eq_of_mem hxs], exact eventually_of_mem self_mem_nhds_within hs})
hmn

/-- If a function is at least `C^1`, its bundled derivative (mapping `(x, v)` to `Df(x) v`) is
continuous. -/
lemma cont_diff_on.continuous_on_fderiv_within_apply
Expand Down
26 changes: 25 additions & 1 deletion src/analysis/calculus/fderiv.lean
Expand Up @@ -319,9 +319,13 @@ theorem has_fderiv_at_filter.mono (h : has_fderiv_at_filter f f' x L₂) (hst :
has_fderiv_at_filter f f' x L₁ :=
h.mono hst

theorem has_fderiv_within_at.mono_of_mem (h : has_fderiv_within_at f f' t x) (hst : t ∈ 𝓝[s] x) :
has_fderiv_within_at f f' s x :=
h.mono $ nhds_within_le_iff.mpr hst

theorem has_fderiv_within_at.mono (h : has_fderiv_within_at f f' t x) (hst : s ⊆ t) :
has_fderiv_within_at f f' s x :=
h.mono (nhds_within_mono _ hst)
h.mono $ nhds_within_mono _ hst

theorem has_fderiv_at.has_fderiv_at_filter (h : has_fderiv_at f f' x) (hL : L ≤ 𝓝 x) :
has_fderiv_at_filter f f' x L :=
Expand Down Expand Up @@ -795,6 +799,11 @@ else
by rw [fderiv_within_zero_of_not_differentiable_within_at h,
fderiv_within_zero_of_not_differentiable_within_at h']

lemma filter.eventually_eq.fderiv_within_eq_nhds (hs : unique_diff_within_at 𝕜 s x)
(hL : f₁ =ᶠ[𝓝 x] f) :
fderiv_within 𝕜 f₁ s x = fderiv_within 𝕜 f s x :=
(show f₁ =ᶠ[𝓝[s] x] f, from nhds_within_le_nhds hL).fderiv_within_eq hs (mem_of_mem_nhds hL : _)

lemma fderiv_within_congr (hs : unique_diff_within_at 𝕜 s x)
(hL : ∀y∈s, f₁ y = f y) (hx : f₁ x = f x) :
fderiv_within 𝕜 f₁ s x = fderiv_within 𝕜 f s x :=
Expand Down Expand Up @@ -1104,6 +1113,21 @@ lemma fderiv_within.comp {g : F → G} {t : set F}
fderiv_within 𝕜 (g ∘ f) s x = (fderiv_within 𝕜 g t (f x)).comp (fderiv_within 𝕜 f s x) :=
(hg.has_fderiv_within_at.comp x (hf.has_fderiv_within_at) h).fderiv_within hxs

/-- Ternary version of `fderiv_within.comp`, with equality assumptions of basepoints added, in
order to apply more easily as a rewrite from right-to-left. -/
lemma fderiv_within.comp₃ {g' : G → G'} {g : F → G} {t : set F} {u : set G} {y : F} {y' : G}
(hg' : differentiable_within_at 𝕜 g' u y') (hg : differentiable_within_at 𝕜 g t y)
(hf : differentiable_within_at 𝕜 f s x)
(h2g : maps_to g t u) (h2f : maps_to f s t)
(h3g : g y = y') (h3f : f x = y) (hxs : unique_diff_within_at 𝕜 s x) :
fderiv_within 𝕜 (g' ∘ g ∘ f) s x = (fderiv_within 𝕜 g' u y').comp
((fderiv_within 𝕜 g t y).comp (fderiv_within 𝕜 f s x)) :=
begin
substs h3g h3f,
exact (hg'.has_fderiv_within_at.comp x
(hg.has_fderiv_within_at.comp x (hf.has_fderiv_within_at) h2f) $ h2g.comp h2f).fderiv_within hxs
end

lemma fderiv.comp {g : F → G}
(hg : differentiable_at 𝕜 g (f x)) (hf : differentiable_at 𝕜 f x) :
fderiv 𝕜 (g ∘ f) x = (fderiv 𝕜 g (f x)).comp (fderiv 𝕜 f x) :=
Expand Down
15 changes: 9 additions & 6 deletions src/topology/continuous_on.lean
Expand Up @@ -118,6 +118,14 @@ begin
{ refine λ h u, eventually_congr (h.mono $ λ x h, _), rw [h] }
end

lemma nhds_within_le_iff {s t : set α} {x : α} : 𝓝[s] x ≤ 𝓝[t] x ↔ t ∈ 𝓝[s] x :=
begin
simp_rw [filter.le_def, mem_nhds_within_iff_eventually],
split,
{ exact λ h, (h t $ eventually_of_forall (λ x, id)).mono (λ x, id) },
{ exact λ h u hu, (h.and hu).mono (λ x hx h, hx.2 $ hx.1 h) }
end

lemma preimage_nhds_within_coinduced' {π : α → β} {s : set β} {t : set α} {a : α}
(h : a ∈ t) (ht : is_open t)
(hs : s ∈ @nhds β (topological_space.coinduced (λ x : t, π x) subtype.topological_space) (π a)) :
Expand Down Expand Up @@ -182,12 +190,7 @@ nhds_within_restrict' s (is_open.mem_nhds h₁ h₀)

theorem nhds_within_le_of_mem {a : α} {s t : set α} (h : s ∈ 𝓝[t] a) :
𝓝[t] a ≤ 𝓝[s] a :=
begin
rcases mem_nhds_within.1 h with ⟨u, u_open, au, uts⟩,
have : 𝓝[t] a = 𝓝[t ∩ u] a := nhds_within_restrict _ au u_open,
rw [this, inter_comm],
exact nhds_within_mono _ uts
end
nhds_within_le_iff.mpr h

theorem nhds_within_le_nhds {a : α} {s : set α} : 𝓝[s] a ≤ 𝓝 a :=
by { rw ← nhds_within_univ, apply nhds_within_le_of_mem, exact univ_mem }
Expand Down

0 comments on commit 25f60b4

Please sign in to comment.