@@ -92,8 +92,8 @@ for each natural `m` is by definition `C^∞` at `0`.
92
92
There is another issue with the definition of `times_cont_diff_within_at 𝕜 n f s x`. We can
93
93
require the existence and good behavior of derivatives up to order `n` on a neighborhood of `x`
94
94
within `s`. However, this does not imply continuity or differentiability within `s` of the function
95
- at `x`. Therefore, we require such existence and good behavior on a neighborhood of `x` within
96
- `s ∪ {x}` (which appears as `insert x s` in this file).
95
+ at `x` when `x` does not belong to `s` . Therefore, we require such existence and good behavior on
96
+ a neighborhood of `x` within `s ∪ {x}` (which appears as `insert x s` in this file).
97
97
98
98
### Side of the composition, and universe issues
99
99
@@ -994,7 +994,7 @@ begin
994
994
end
995
995
996
996
/-- A function is `C^(n + 1)` on a domain with unique derivatives if and only if it is
997
- differentiable there, and its derivative is `C^n`. -/
997
+ differentiable there, and its derivative (expressed with `fderiv_within`) is `C^n`. -/
998
998
theorem times_cont_diff_on_succ_iff_fderiv_within {n : ℕ} (hs : unique_diff_on 𝕜 s) :
999
999
times_cont_diff_on 𝕜 ((n + 1 ) : ℕ) f s ↔
1000
1000
differentiable_on 𝕜 f s ∧ times_cont_diff_on 𝕜 n (λ y, fderiv_within 𝕜 f s y) s :=
@@ -1022,8 +1022,22 @@ begin
1022
1022
λ y hy, (hdiff y hy).has_fderiv_within_at, h x hx⟩ }
1023
1023
end
1024
1024
1025
+ /-- A function is `C^(n + 1)` on an open domain if and only if it is
1026
+ differentiable there, and its derivative (expressed with `fderiv`) is `C^n`. -/
1027
+ theorem times_cont_diff_on_succ_iff_fderiv_of_open {n : ℕ} (hs : is_open s) :
1028
+ times_cont_diff_on 𝕜 ((n + 1 ) : ℕ) f s ↔
1029
+ differentiable_on 𝕜 f s ∧ times_cont_diff_on 𝕜 n (λ y, fderiv 𝕜 f y) s :=
1030
+ begin
1031
+ rw times_cont_diff_on_succ_iff_fderiv_within hs.unique_diff_on,
1032
+ congr' 2 ,
1033
+ rw ← iff_iff_eq,
1034
+ apply times_cont_diff_on_congr,
1035
+ assume x hx,
1036
+ exact fderiv_within_of_open hs hx
1037
+ end
1038
+
1025
1039
/-- A function is `C^∞` on a domain with unique derivatives if and only if it is differentiable
1026
- there, and its derivative is `C^∞`. -/
1040
+ there, and its derivative (expressed with `fderiv_within`) is `C^∞`. -/
1027
1041
theorem times_cont_diff_on_top_iff_fderiv_within (hs : unique_diff_on 𝕜 s) :
1028
1042
times_cont_diff_on 𝕜 ∞ f s ↔
1029
1043
differentiable_on 𝕜 f s ∧ times_cont_diff_on 𝕜 ∞ (λ y, fderiv_within 𝕜 f s y) s :=
@@ -1040,6 +1054,20 @@ begin
1040
1054
exact with_top.coe_le_coe.2 (nat.le_succ n) }
1041
1055
end
1042
1056
1057
+ /-- A function is `C^∞` on a domain with unique derivatives if and only if it is differentiable
1058
+ there, and its derivative (expressed with `fderiv`) is `C^∞`. -/
1059
+ theorem times_cont_diff_on_top_iff_fderiv_of_open (hs : is_open s) :
1060
+ times_cont_diff_on 𝕜 ∞ f s ↔
1061
+ differentiable_on 𝕜 f s ∧ times_cont_diff_on 𝕜 ∞ (λ y, fderiv 𝕜 f y) s :=
1062
+ begin
1063
+ rw times_cont_diff_on_top_iff_fderiv_within hs.unique_diff_on,
1064
+ congr' 2 ,
1065
+ rw ← iff_iff_eq,
1066
+ apply times_cont_diff_on_congr,
1067
+ assume x hx,
1068
+ exact fderiv_within_of_open hs hx
1069
+ end
1070
+
1043
1071
lemma times_cont_diff_on.fderiv_within {m n : with_top ℕ}
1044
1072
(hf : times_cont_diff_on 𝕜 n f s) (hs : unique_diff_on 𝕜 s) (hmn : m + 1 ≤ n) :
1045
1073
times_cont_diff_on 𝕜 m (λ y, fderiv_within 𝕜 f s y) s :=
@@ -1053,11 +1081,21 @@ begin
1053
1081
exact ((times_cont_diff_on_succ_iff_fderiv_within hs).1 (hf.of_le hmn)).2 }
1054
1082
end
1055
1083
1084
+ lemma times_cont_diff_on.fderiv_of_open {m n : with_top ℕ}
1085
+ (hf : times_cont_diff_on 𝕜 n f s) (hs : is_open s) (hmn : m + 1 ≤ n) :
1086
+ times_cont_diff_on 𝕜 m (λ y, fderiv 𝕜 f y) s :=
1087
+ (hf.fderiv_within hs.unique_diff_on hmn).congr (λ x hx, (fderiv_within_of_open hs hx).symm)
1088
+
1056
1089
lemma times_cont_diff_on.continuous_on_fderiv_within {n : with_top ℕ}
1057
1090
(h : times_cont_diff_on 𝕜 n f s) (hs : unique_diff_on 𝕜 s) (hn : 1 ≤ n) :
1058
1091
continuous_on (λ x, fderiv_within 𝕜 f s x) s :=
1059
1092
((times_cont_diff_on_succ_iff_fderiv_within hs).1 (h.of_le hn)).2 .continuous_on
1060
1093
1094
+ lemma times_cont_diff_on.continuous_on_fderiv_of_open {n : with_top ℕ}
1095
+ (h : times_cont_diff_on 𝕜 n f s) (hs : is_open s) (hn : 1 ≤ n) :
1096
+ continuous_on (λ x, fderiv 𝕜 f x) s :=
1097
+ ((times_cont_diff_on_succ_iff_fderiv_of_open hs).1 (h.of_le hn)).2 .continuous_on
1098
+
1061
1099
/-- If a function is at least `C^1`, its bundled derivative (mapping `(x, v)` to `Df(x) v`) is
1062
1100
continuous. -/
1063
1101
lemma times_cont_diff_on.continuous_on_fderiv_within_apply
@@ -2310,3 +2348,116 @@ lemma times_cont_diff.has_strict_fderiv_at
2310
2348
hf.times_cont_diff_at.has_strict_fderiv_at hn
2311
2349
2312
2350
end real
2351
+
2352
+ section deriv
2353
+ /-!
2354
+ ### One dimension
2355
+
2356
+ All results up to now have been expressed in terms of the general Fréchet derivative `fderiv`. For
2357
+ maps defined on the field, the one-dimensional derivative `deriv` is often easier to use. In this
2358
+ paragraph, we reformulate some higher smoothness results in terms of `deriv`.
2359
+ -/
2360
+
2361
+ variables {f₂ : 𝕜 → F} {s₂ : set 𝕜}
2362
+ open continuous_linear_map (smul_right)
2363
+
2364
+ /-- A function is `C^(n + 1)` on a domain with unique derivatives if and only if it is
2365
+ differentiable there, and its derivative (formulated with `deriv_within`) is `C^n`. -/
2366
+ theorem times_cont_diff_on_succ_iff_deriv_within {n : ℕ} (hs : unique_diff_on 𝕜 s₂) :
2367
+ times_cont_diff_on 𝕜 ((n + 1 ) : ℕ) f₂ s₂ ↔
2368
+ differentiable_on 𝕜 f₂ s₂ ∧ times_cont_diff_on 𝕜 n (deriv_within f₂ s₂) s₂ :=
2369
+ begin
2370
+ rw times_cont_diff_on_succ_iff_fderiv_within hs,
2371
+ congr' 2 ,
2372
+ rw ← iff_iff_eq,
2373
+ split,
2374
+ { assume h,
2375
+ have : deriv_within f₂ s₂ = (λ u : 𝕜 →L[𝕜] F, u 1 ) ∘ (fderiv_within 𝕜 f₂ s₂),
2376
+ by { ext x, refl },
2377
+ simp only [this ],
2378
+ apply times_cont_diff.comp_times_cont_diff_on _ h,
2379
+ exact (is_bounded_bilinear_map_apply.is_bounded_linear_map_left _).times_cont_diff },
2380
+ { assume h,
2381
+ have : fderiv_within 𝕜 f₂ s₂ = (λ u, smul_right 1 u) ∘ (λ x, deriv_within f₂ s₂ x),
2382
+ by { ext x, simp [deriv_within] },
2383
+ simp only [this ],
2384
+ apply times_cont_diff.comp_times_cont_diff_on _ h,
2385
+ exact (is_bounded_bilinear_map_smul_right.is_bounded_linear_map_right _).times_cont_diff }
2386
+ end
2387
+
2388
+ /-- A function is `C^(n + 1)` on an open domain if and only if it is
2389
+ differentiable there, and its derivative (formulated with `deriv`) is `C^n`. -/
2390
+ theorem times_cont_diff_on_succ_iff_deriv_of_open {n : ℕ} (hs : is_open s₂) :
2391
+ times_cont_diff_on 𝕜 ((n + 1 ) : ℕ) f₂ s₂ ↔
2392
+ differentiable_on 𝕜 f₂ s₂ ∧ times_cont_diff_on 𝕜 n (deriv f₂) s₂ :=
2393
+ begin
2394
+ rw times_cont_diff_on_succ_iff_deriv_within hs.unique_diff_on,
2395
+ congr' 2 ,
2396
+ rw ← iff_iff_eq,
2397
+ apply times_cont_diff_on_congr,
2398
+ assume x hx,
2399
+ exact deriv_within_of_open hs hx
2400
+ end
2401
+
2402
+ /-- A function is `C^∞` on a domain with unique derivatives if and only if it is differentiable
2403
+ there, and its derivative (formulated with `deriv_within`) is `C^∞`. -/
2404
+ theorem times_cont_diff_on_top_iff_deriv_within (hs : unique_diff_on 𝕜 s₂) :
2405
+ times_cont_diff_on 𝕜 ∞ f₂ s₂ ↔
2406
+ differentiable_on 𝕜 f₂ s₂ ∧ times_cont_diff_on 𝕜 ∞ (deriv_within f₂ s₂) s₂ :=
2407
+ begin
2408
+ split,
2409
+ { assume h,
2410
+ refine ⟨h.differentiable_on le_top, _⟩,
2411
+ apply times_cont_diff_on_top.2 (λ n, ((times_cont_diff_on_succ_iff_deriv_within hs).1 _).2 ),
2412
+ exact h.of_le le_top },
2413
+ { assume h,
2414
+ refine times_cont_diff_on_top.2 (λ n, _),
2415
+ have A : (n : with_top ℕ) ≤ ∞ := le_top,
2416
+ apply ((times_cont_diff_on_succ_iff_deriv_within hs).2 ⟨h.1 , h.2 .of_le A⟩).of_le,
2417
+ exact with_top.coe_le_coe.2 (nat.le_succ n) }
2418
+ end
2419
+
2420
+
2421
+ /-- A function is `C^∞` on an open domain if and only if it is differentiable
2422
+ there, and its derivative (formulated with `deriv`) is `C^∞`. -/
2423
+ theorem times_cont_diff_on_top_iff_deriv_of_open (hs : is_open s₂) :
2424
+ times_cont_diff_on 𝕜 ∞ f₂ s₂ ↔
2425
+ differentiable_on 𝕜 f₂ s₂ ∧ times_cont_diff_on 𝕜 ∞ (deriv f₂) s₂ :=
2426
+ begin
2427
+ rw times_cont_diff_on_top_iff_deriv_within hs.unique_diff_on,
2428
+ congr' 2 ,
2429
+ rw ← iff_iff_eq,
2430
+ apply times_cont_diff_on_congr,
2431
+ assume x hx,
2432
+ exact deriv_within_of_open hs hx
2433
+ end
2434
+
2435
+ lemma times_cont_diff_on.deriv_within {m n : with_top ℕ}
2436
+ (hf : times_cont_diff_on 𝕜 n f₂ s₂) (hs : unique_diff_on 𝕜 s₂) (hmn : m + 1 ≤ n) :
2437
+ times_cont_diff_on 𝕜 m (deriv_within f₂ s₂) s₂ :=
2438
+ begin
2439
+ cases m,
2440
+ { change ∞ + 1 ≤ n at hmn,
2441
+ have : n = ∞, by simpa using hmn,
2442
+ rw this at hf,
2443
+ exact ((times_cont_diff_on_top_iff_deriv_within hs).1 hf).2 },
2444
+ { change (m.succ : with_top ℕ) ≤ n at hmn,
2445
+ exact ((times_cont_diff_on_succ_iff_deriv_within hs).1 (hf.of_le hmn)).2 }
2446
+ end
2447
+
2448
+ lemma times_cont_diff_on.deriv_of_open {m n : with_top ℕ}
2449
+ (hf : times_cont_diff_on 𝕜 n f₂ s₂) (hs : is_open s₂) (hmn : m + 1 ≤ n) :
2450
+ times_cont_diff_on 𝕜 m (deriv f₂) s₂ :=
2451
+ (hf.deriv_within hs.unique_diff_on hmn).congr (λ x hx, (deriv_within_of_open hs hx).symm)
2452
+
2453
+ lemma times_cont_diff_on.continuous_on_deriv_within {n : with_top ℕ}
2454
+ (h : times_cont_diff_on 𝕜 n f₂ s₂) (hs : unique_diff_on 𝕜 s₂) (hn : 1 ≤ n) :
2455
+ continuous_on (deriv_within f₂ s₂) s₂ :=
2456
+ ((times_cont_diff_on_succ_iff_deriv_within hs).1 (h.of_le hn)).2 .continuous_on
2457
+
2458
+ lemma times_cont_diff_on.continuous_on_deriv_of_open {n : with_top ℕ}
2459
+ (h : times_cont_diff_on 𝕜 n f₂ s₂) (hs : is_open s₂) (hn : 1 ≤ n) :
2460
+ continuous_on (deriv f₂) s₂ :=
2461
+ ((times_cont_diff_on_succ_iff_deriv_of_open hs).1 (h.of_le hn)).2 .continuous_on
2462
+
2463
+ end deriv
0 commit comments