Skip to content
This commit does not belong to any branch on this repository, and may belong to a fork outside of the repository.

Commit d8fd180

Browse files
affeldt-aistmkerjean
authored and
mkerjean
committedNov 25, 2020
LinearContinuousBounded
Co-authored-by: Théo Vignon <>
1 parent ab12ecd commit d8fd180

File tree

2 files changed

+204
-0
lines changed

2 files changed

+204
-0
lines changed
 

‎CHANGELOG_UNRELEASED.md

+8
Original file line numberDiff line numberDiff line change
@@ -12,6 +12,14 @@
1212
+ definition `mule` and its notation `*` (scope `ereal_scope`)
1313
+ definition `abse` and its notation `` `| | `` (scope `ereal_scope`)
1414

15+
+ lemmas `setIIl`, `setIIr`, `setCS`, `setSD`, `setDS`, `setDSS`, `setCI`,
16+
`setDUr`, `setDUl`, `setIDA`, `setDD`
17+
- in `normedtype.v`:
18+
+ lemma `closed_ereal_le_ereal`
19+
+ lemma `closed_ereal_ge_ereal`
20+
+ section "LinearContinuousBounded"
21+
+ section "Closed_ball"
22+
1523
### Changed
1624

1725
### Renamed

‎theories/normedtype.v

+196
Original file line numberDiff line numberDiff line change
@@ -2663,6 +2663,19 @@ rewrite [in X in _ <= X]/normr /= mx_normrE.
26632663
by apply/BigmaxBigminr.bigmaxr_gerP; right => /=; exists j.
26642664
Qed.
26652665

2666+
Lemma ereal_nbhs'_le (R : numFieldType) (x : {ereal R}) :
2667+
ereal_nbhs' x --> ereal_nbhs x.
2668+
Proof.
2669+
move: x => [x P [_/posnumP[e] HP] |x P|x P] //=.
2670+
by exists e%:num => // ???; apply: HP.
2671+
Qed.
2672+
2673+
Lemma ereal_nbhs'_le_finite (R : numFieldType) (x : R) :
2674+
ereal_nbhs' x%:E --> nbhs x%:E.
2675+
Proof.
2676+
by move=> P [_/posnumP[e] HP] //=; exists e%:num => // ???; apply: HP.
2677+
Qed.
2678+
26662679
Section open_closed_sets_ereal.
26672680
Variable R : realFieldType (* TODO: generalize to numFieldType? *).
26682681
Implicit Types x y : {ereal R}.
@@ -2844,3 +2857,186 @@ rewrite !nbhs_nearE !near_map !near_nbhs in fxP *; have /= := cfx P fxP.
28442857
rewrite !near_simpl near_withinE near_simpl => Pf; near=> y.
28452858
by have [->|] := eqVneq y x; [by apply: nbhs_singleton|near: y].
28462859
Grab Existential Variables. all: end_near. Qed.
2860+
2861+
(* NB: technical lemma that comes in handy in Banach-Steinhaus' PR *)
2862+
Lemma normrZV (R : numFieldType) (V : normedModType R) (x : V) :
2863+
x != 0 -> `| `| x|^-1 *: x | = 1.
2864+
Proof.
2865+
move=> x0; rewrite normmZ normrV ?unitfE ?normr_eq0 //=.
2866+
by rewrite normr_id mulVr // unitfE normr_eq0.
2867+
Qed.
2868+
2869+
Section LinearContinuousBounded.
2870+
2871+
Variables (R : numFieldType) (V W : normedModType R).
2872+
2873+
2874+
(*drafts for another PR *)
2875+
Definition absorbed (A B : set V) :=
2876+
exists (r : R), A `<=` [ set x : V | exists y : V, B y /\ x = r *: y ].
2877+
2878+
Definition strong_bornology :=
2879+
[ set A | (forall (B : set V), (nbhs (0:V) B -> absorbed A B))].
2880+
2881+
Lemma normed_bornological (A : set V) : bounded_set A <->
2882+
(forall (B : set V), (nbhs (0:V) B -> absorbed A B)).
2883+
Proof.
2884+
Admitted.
2885+
(*To be generalized to any topological space, with a theory of bounded*)
2886+
2887+
2888+
Definition bounded_fun_on (f : V -> W) (B : set (set V)):=
2889+
forall A, (B A -> (exists M : R, \forall x \near (0 :V), (`|f x| <= M))).
2890+
2891+
2892+
Lemma linear_bounded (f : {linear V -> W} ):
2893+
( bounded_on f (nbhs (0 : V))) <->
2894+
(*exists r : R, 0 <= r /\ (forall x : V, `|f x| <= r * `|x|).*)
2895+
\forall r \near +oo, (forall x : V, `|f x| <= r * `|x|).
2896+
Proof.
2897+
split; last first.
2898+
move=> /pinfty_ex_gt0 [r [r0 Bf]]; rewrite /bounded_on; near=> M; rewrite -nearE.
2899+
near=> x; apply: le_trans; first by apply: Bf.
2900+
apply: (le_trans (y :=r)).
2901+
apply: ler_pimulr; apply: ltW; first by [].
2902+
near: x; rewrite nearE.
2903+
suff <-: ball_ normr 0 1 = (fun x : V => `|x| < 1) by apply: nbhs_ball_norm.
2904+
by apply: funext => x /=; rewrite sub0r normrN.
2905+
near: M; rewrite nearE /=; exists r; split; first by apply: gtr0_real.
2906+
by move => x; apply: ltW. (* much longer *)
2907+
rewrite /bounded_on => Bf; near=> r => x.
2908+
have [->|x0] := eqVneq x 0; first by rewrite linear0 !normr0 mulr0.
2909+
rewrite -ler_pdivr_mulr ; last by rewrite normr_gt0.
2910+
rewrite -(normr_id x) mulrC.
2911+
rewrite -normrV ?unitfE ?normr_eq0 // -normmZ -linearZ.
2912+
move:x x0; near: r.
2913+
(* hell on earth *)
2914+
(* how to use hypothesis with near *)
2915+
(* split => [[M /nbhs_ballP[a a0 Ba0]]|[r [r0 fr]]]; last first. *)
2916+
(* exists r; apply/nbhs_ballP; exists 1 => // x. *)
2917+
(* rewrite -ball_normE //= sub0r normrN => x1; apply: (le_trans (fr _)). *)
2918+
(* by rewrite ler_pimulr // ltW. *)
2919+
(* exists (M * 2 * a^-1); split. *)
2920+
(* - rewrite !mulr_ge0 //= ?invr_ge0 ; last by apply: ltW. *)
2921+
(* by move: (Ba0 0 (ballxx 0 a0)); rewrite linear0 normr0. *)
2922+
(* - move=> x; rewrite mulrAC ler_pdivl_mulr //=. *)
2923+
(* have [/eqP ->|x0] := boolP (x == 0). *)
2924+
(* + by rewrite linear0 !(normr0,mulr0,mul0r). *)
2925+
(* + rewrite mulrAC -lter_pdivr_mulr // -lter_pdivr_mulr ?normr_gt0 //. *)
2926+
(* rewrite [X in X <= _](_ : _ = `|f (a / 2 / `|x| *: x)|); last first. *)
2927+
(* rewrite linearZ normmZ 2!normrM (gtr0_norm a0) gtr0_norm //. *)
2928+
(* by rewrite normrV ?unitfE ?normr_eq0 // normr_id -2!mulrA mulrC mulrA. *)
2929+
(* apply: Ba0; rewrite -ball_normE /= sub0r normrN. *)
2930+
(* rewrite normmZ normrM (@normrV _ `|x|) ?unitfE ?normr_eq0 // normr_id. *)
2931+
(* rewrite -mulrA mulVr ?mulr1 ?unitfE ?normr_eq0 //. *)
2932+
(* by rewrite normrM 2?gtr0_norm// gtr_pmulr // invf_lt1 // ltr1n. *)
2933+
(* Qed. *)
2934+
Admitted.
2935+
2936+
2937+
Lemma linear_boundedP (f : {linear V -> W}) :
2938+
(bounded_on f (nbhs (0 : V))) <->
2939+
exists r : R, 0 <= r /\ (forall x : V, `|f x| <= r * `|x|).
2940+
Proof.
2941+
split => [/linear_bounded | B].
2942+
by move=> /pinfty_ex_gt0 [M [M0 B]]; exists M; split; first by apply: ltW.
2943+
apply/linear_bounded; near_simpl. Admitted.
2944+
2945+
Lemma linear_continuous0 (f : {linear V -> W}) :
2946+
{for 0, continuous f} -> bounded_on f (nbhs (0 : V)).
2947+
Proof.
2948+
(* move=> /cvg_distP; rewrite linear0 /= => Bf. *)
2949+
(* apply/linear_boundedP; near=> r => x. *)
2950+
(* have [->|x0] := eqVneq x 0; first by rewrite linear0 !normr0 mulr0. *)
2951+
(* rewrite -ler_pdivr_mulr; last by rewrite normr_gt0. *)
2952+
(* rewrite -(normr_id x) mulrC. *)
2953+
(* rewrite -normrV ?unitfE ?normr_eq0 // -normmZ -linearZ. *)
2954+
(* apply: ltW. move: (Bf r) x x0; near: r. apply/near_pinfty_div2. *)
2955+
move=> /cvg_ballP /(_ _ ltr01).
2956+
rewrite linear0 /= nearE => /nbhs_ex[tp ball_f]; apply/linear_boundedP.
2957+
pose t := tp%:num; exists (2 * t^-1); split => [|x].
2958+
by apply: divr_ge0.
2959+
have [->|x0] := eqVneq x 0; first by rewrite linear0 !normr0 mulr0.
2960+
have /ball_f : ball 0 t ((`|x|^-1 * t /2) *: x).
2961+
apply: ball_sym; rewrite -ball_normE /ball_ subr0 normmZ mulrC 2!normrM.
2962+
rewrite 2!mulrA normrV ?unitfE ?normr_eq0 // normr_id.
2963+
rewrite divrr ?mul1r ?unitfE ?normr_eq0 // gtr0_norm // gtr_pmulr //.
2964+
by rewrite gtr0_norm // invr_lt1 // ?unitfE // ltr1n.
2965+
rewrite -ball_normE //= sub0r normrN linearZ /= normmZ -mulrA normrM.
2966+
rewrite normrV ?unitfE ?normr_eq0 // normr_id -mulrA.
2967+
rewrite ltr_pdivr_mull ?mulr1 ?normr_gt0 // -ltr_pdivl_mull ?normr_gt0 //.
2968+
rewrite gtr0_norm // invf_div => /ltW H. apply: le_trans; first by exact: H.
2969+
by apply: ler_pmul.
2970+
Qed.
2971+
2972+
Lemma linear_bounded0 (f : {linear V -> W}) :
2973+
bounded_on f (nbhs (0 : V)) -> {for 0, continuous f} .
2974+
Proof.
2975+
move=> /linear_boundedP [r]; rewrite le0r=> [[/orP r0]]; case: r0.
2976+
- move/eqP => -> fr; apply: near_cst_continuous; near=> y.
2977+
by move: (fr y); rewrite mul0r normr_le0; apply/eqP.
2978+
- move => r0 fr; apply/cvg_ballP => e e0.
2979+
rewrite nearE linear0 /= nbhs_ballP.
2980+
exists (e / 2 / r); first by rewrite !divr_gt0.
2981+
move=> x; rewrite -2!ball_normE /= 2!sub0r 2!normrN => xr.
2982+
have /le_lt_trans -> // : `|f x| <= e / 2.
2983+
by rewrite (le_trans (fr x)) // mulrC -ler_pdivl_mulr // ltW.
2984+
by rewrite gtr_pmulr // invr_lt1 // ?unitfE // ltr1n.
2985+
Grab Existential Variables. by end_near. Qed.
2986+
2987+
Lemma continuousfor0_continuous (f : {linear V -> W}) :
2988+
{for 0, continuous f} -> continuous f.
2989+
Proof.
2990+
move=> /linear_continuous0 /linear_boundedP [r []].
2991+
rewrite le_eqVlt => /orP[/eqP <- f0|r0 fr x].
2992+
- suff : f = (fun _ => 0) :> (_ -> _) by move=> ->; exact: cst_continuous.
2993+
by rewrite funeqE => x; move: (f0 x); rewrite mul0r normr_le0 => /eqP.
2994+
- rewrite cvg_to_locally => e e0; rewrite nearE /= nbhs_ballP.
2995+
exists (e / r); first by rewrite divr_gt0.
2996+
move=> y; rewrite -!ball_normE //= => xy; rewrite -linearB.
2997+
by rewrite (le_lt_trans (fr (x - y))) // mulrC -ltr_pdivl_mulr.
2998+
Qed.
2999+
3000+
Lemma linear_bounded_continuous (f : {linear V -> W}) :
3001+
bounded_on f (nbhs (0 : V)) <-> continuous f.
3002+
Proof.
3003+
split=> [/linear_bounded0|/(_ 0)/linear_continuous0//].
3004+
exact: continuousfor0_continuous.
3005+
Qed.
3006+
3007+
Definition bounded_fun_norm (f : V -> W) :=
3008+
forall r, exists M, forall x, `|x| <= r -> `|f x| <= M.
3009+
(*bounded_fun_on bounded_sets is harder to use*)
3010+
3011+
(* Definition pointwise_bounded (F : set (V -> W)) := *)
3012+
(* forall x, exists M, forall f, F f -> `|f x| <= M. *)
3013+
3014+
(* Definition uniform_bounded (F : set (V -> W)) := *)
3015+
(* forall r, exists M, forall f, F f -> forall x, `|x| <= r -> `|f x| <= M. *)
3016+
3017+
Lemma bounded_funP (f : {linear V -> W}):
3018+
bounded_fun_norm f <-> bounded_on f (nbhs (0 : V)).
3019+
Proof.
3020+
split => [/(_ 1) [M Bf]|/linear_boundedP [r [r0 Bf]] e]; last first.
3021+
by exists (e * r) => x xe; rewrite (le_trans (Bf _)) // mulrC ler_pmul.
3022+
apply/linear_boundedP.
3023+
have M0 : 0 <= M by move: (Bf 0); rewrite linear0 !normr0 => /(_ ler01).
3024+
exists M; split => //.
3025+
move: M0; rewrite le0r => /orP[/eqP|] M0 x.
3026+
- rewrite M0 mul0r normr_le0.
3027+
have [/eqP ->|x0] := boolP (x == 0); first by rewrite linear0.
3028+
have /Bf : `| `|x|^-1 *: x | <= 1.
3029+
rewrite normmZ normrV ?unitf_gt0 ?normrE //=.
3030+
by rewrite mulrC mulrV ?unitf_gt0 ?normrE.
3031+
rewrite linearZ M0 normr_le0 scaler_eq0 invr_eq0 => /orP[|//].
3032+
by rewrite normr_eq0 => /eqP ->; rewrite linear0.
3033+
- have [/eqP ->|x0] := boolP (x == 0).
3034+
+ by rewrite normr0 mulr0 normr_le0 linear0.
3035+
+ rewrite -ler_pdivr_mulr ?normr_gt0 //= mulrC.
3036+
have xx1 : `| `|x|^-1 | = `|x|^-1.
3037+
by rewrite normrV ?unitfE ?normr_eq0 // normr_id.
3038+
rewrite -xx1 -normmZ -linearZ; apply: (Bf (`|x|^-1 *: x)).
3039+
by rewrite normmZ xx1 mulVr // unitfE normr_eq0.
3040+
Qed.
3041+
3042+
End LinearContinuousBounded.

0 commit comments

Comments
 (0)
Please sign in to comment.