From 3f961407936f02cf6b9cce99ac49a6f59d5b0527 Mon Sep 17 00:00:00 2001 From: damien rouhling Date: Fri, 28 Sep 2018 16:07:45 +0200 Subject: [PATCH 01/11] Remove ball and absRingType --- Rstruct.v | 30 +- _CoqProject | 2 +- classical_sets.v | 45 +- derive.v | 348 ++++--- hierarchy.v | 2432 ++++++++++++++----------------------------- landau.v | 90 +- misc/uniform_bigO.v | 29 +- posnum.v | 7 +- summability.v | 11 +- topology.v | 813 ++++++++++++++- 10 files changed, 1886 insertions(+), 1921 deletions(-) diff --git a/Rstruct.v b/Rstruct.v index 58d165bba6..90c1902443 100644 --- a/Rstruct.v +++ b/Rstruct.v @@ -481,17 +481,21 @@ case: (lerP x y) => H; first by rewrite minr_l // Rmin_left //; apply: RlebP. by rewrite minr_r ?ltrW // Rmin_right //; apply/RlebP; move/ltrW : H. Qed. +Section Max. + +Context (R : realDomainType). + (* bigop pour le max pour des listes non vides ? *) Definition bigmaxr (x0 : R) lr := foldr Num.max (head x0 lr) (behead lr). -Lemma bigmaxr_nil x0 : bigmaxr x0 [::] = x0. +Lemma bigmaxr_nil (x0 : R) : bigmaxr x0 [::] = x0. Proof. by rewrite /bigmaxr. Qed. -Lemma bigmaxr_un x0 x : bigmaxr x0 [:: x] = x. +Lemma bigmaxr_un (x0 x : R) : bigmaxr x0 [:: x] = x. Proof. by rewrite /bigmaxr. Qed. -Lemma bigmaxr_cons x0 x y lr : +Lemma bigmaxr_cons (x0 x y : R) lr : bigmaxr x0 (x :: y :: lr) = Num.max x (bigmaxr x0 (y :: lr)). Proof. rewrite /bigmaxr /=; elim: lr => [/= | a lr /=]; first by rewrite maxrC. @@ -499,7 +503,7 @@ set b := foldr _ _ _; set c := foldr _ _ _ => H. by rewrite [Num.max a b]maxrC maxrA H -maxrA (maxrC c a). Qed. -Lemma bigmaxr_ler x0 lr i : +Lemma bigmaxr_ler (x0 : R) lr i : (i < size lr)%N -> (nth x0 lr i) <= (bigmaxr x0 lr). Proof. case: lr i => [i | x lr]; first by rewrite nth_nil bigmaxr_nil lerr. @@ -512,15 +516,15 @@ by rewrite ler_maxr lerr orbT. Qed. (* Compatibilité avec l'addition *) -Lemma bigmaxr_addr x0 lr x : - bigmaxr (x0 + x) (map (fun y => y + x) lr) = (bigmaxr x0 lr) + x. +Lemma bigmaxr_addr (x0 : R) lr (x : R) : + bigmaxr (x0 + x) (map (fun y : R => y + x) lr) = (bigmaxr x0 lr) + x. Proof. case: lr => [/= | y lr]; first by rewrite bigmaxr_nil. elim: lr y => [y | y lr ihlr z]; first by rewrite /= !bigmaxr_un. by rewrite map_cons !bigmaxr_cons ihlr addr_maxl. Qed. -Lemma bigmaxr_index x0 lr : +Lemma bigmaxr_index (x0 : R) lr : (0 < size lr)%N -> (index (bigmaxr x0 lr) lr < size lr)%N. Proof. case: lr => [//= | x l _]. @@ -530,11 +534,11 @@ case: (z <= y); first by rewrite eq_refl. by case: (y == z); rewrite //. Qed. -Lemma bigmaxr_mem x0 lr : +Lemma bigmaxr_mem (x0 : R) lr : (0 < size lr)%N -> bigmaxr x0 lr \in lr. Proof. by move/(bigmaxr_index x0); rewrite index_mem. Qed. -Lemma bigmaxr_lerP x0 lr x : +Lemma bigmaxr_lerP (x0 : R) lr (x : R) : (0 < size lr)%N -> reflect (forall i, (i < size lr)%N -> (nth x0 lr i) <= x) ((bigmaxr x0 lr) <= x). Proof. @@ -543,7 +547,7 @@ move=> lr_size; apply: (iffP idP) => [le_x i i_size | H]. by move/(nthP x0): (bigmaxr_mem x0 lr_size) => [i i_size <-]; apply: H. Qed. -Lemma bigmaxr_ltrP x0 lr x : +Lemma bigmaxr_ltrP (x0 : R) lr (x : R) : (0 < size lr)%N -> reflect (forall i, (i < size lr)%N -> (nth x0 lr i) < x) ((bigmaxr x0 lr) < x). Proof. @@ -552,7 +556,7 @@ move=> lr_size; apply: (iffP idP) => [lt_x i i_size | H]. by move/(nthP x0): (bigmaxr_mem x0 lr_size) => [i i_size <-]; apply: H. Qed. -Lemma bigmaxrP x0 lr x : +Lemma bigmaxrP (x0 : R) lr (x : R) : (x \in lr /\ forall i, (i < size lr) %N -> (nth x0 lr i) <= x) -> (bigmaxr x0 lr = x). Proof. move=> [] /(nthP x0) [] j j_size j_nth x_ler; apply: ler_asym; apply/andP; split. @@ -571,7 +575,7 @@ apply/negP => /eqP H; apply: neq_i; rewrite -H eq_sym; apply/eqP. by apply: index_uniq. Qed. *) -Lemma bigmaxr_lerif x0 lr : +Lemma bigmaxr_lerif (x0 : R) lr : uniq lr -> forall i, (i < size lr)%N -> (nth x0 lr i) <= (bigmaxr x0 lr) ?= iff (i == index (bigmaxr x0 lr) lr). Proof. @@ -625,4 +629,6 @@ Proof. by move=> inj_f i; rewrite /lerif bmaxrf_ler -(inj_eq inj_f) eq_index_bmaxrf. Qed. +End Max. + End ssreal_struct_contd. diff --git a/_CoqProject b/_CoqProject index 7a4f3889f1..218a9d0ccc 100644 --- a/_CoqProject +++ b/_CoqProject @@ -7,7 +7,7 @@ posnum.v landau.v classical_sets.v Rstruct.v -Rbar.v +# Rbar.v topology.v hierarchy.v forms.v diff --git a/classical_sets.v b/classical_sets.v index 63d4a0a411..95081770ba 100644 --- a/classical_sets.v +++ b/classical_sets.v @@ -1,6 +1,6 @@ From mathcomp Require Import ssreflect ssrfun ssrbool ssrnat eqtype choice. -From mathcomp Require Import ssralg matrix. -Require Import boolp. +From mathcomp Require Import ssralg matrix ssrnum. +Require Import boolp reals. (******************************************************************************) (* This file develops a basic theory of sets and types equipped with a *) @@ -447,6 +447,47 @@ Canonical prod_pointedType (T T' : pointedType) := PointedType (T * T') (point, point). Canonical matrix_pointedType m n (T : pointedType) := PointedType 'M[T]_(m, n) (\matrix_(_, _) point)%R. +Canonical zmod_pointedType (V : zmodType) := PointedType V 0%R. +Canonical ring_pointedType (R : ringType) := + [pointedType of R for zmod_pointedType R]. +Canonical lmod_pointedType (R : ringType) (V : lmodType R) := + [pointedType of V for zmod_pointedType V]. +Canonical lalg_pointedType (R : ringType) (A : lalgType R) := + [pointedType of A for zmod_pointedType A]. +Canonical comRing_pointedType (R : comRingType) := + [pointedType of R for zmod_pointedType R]. +Canonical alg_pointedType (R : comRingType) (A : algType R) := + [pointedType of A for zmod_pointedType A]. +Canonical unitRing_pointedType (R : unitRingType) := + [pointedType of R for zmod_pointedType R]. +Canonical comUnitRing_pointedType (R : comUnitRingType) := + [pointedType of R for zmod_pointedType R]. +Canonical unitAlg_pointedType (R : comUnitRingType) (A : unitAlgType R) := + [pointedType of A for zmod_pointedType A]. +Canonical idomain_pointedType (R : idomainType) := + [pointedType of R for zmod_pointedType R]. +Canonical field_pointedType (F : fieldType) := + [pointedType of F for zmod_pointedType F]. +Canonical decField_pointedType (F : decFieldType) := + [pointedType of F for zmod_pointedType F]. +Canonical closedField_pointedType (F : closedFieldType) := + [pointedType of F for zmod_pointedType F]. +Canonical numDomain_pointedType (R : numDomainType) := + [pointedType of R for zmod_pointedType R]. +Canonical numField_pointedType (R : numFieldType) := + [pointedType of R for zmod_pointedType R]. +Canonical numClosedField_pointedType (R : numClosedFieldType) := + [pointedType of R for zmod_pointedType R]. +Canonical realDomain_pointedType (R : realDomainType) := + [pointedType of R for zmod_pointedType R]. +Canonical realField_pointedType (R : realFieldType) := + [pointedType of R for zmod_pointedType R]. +Canonical archiField_pointedType (R : archiFieldType) := + [pointedType of R for zmod_pointedType R]. +Canonical rcf_pointedType (R : rcfType) := + [pointedType of R for zmod_pointedType R]. +Canonical real_pointedType (R : realType) := + [pointedType of R for zmod_pointedType R]. Notation get := (xget point). diff --git a/derive.v b/derive.v index 7e137c2664..ed3d5671d5 100644 --- a/derive.v +++ b/derive.v @@ -1,9 +1,8 @@ (* mathcomp analysis (c) 2017 Inria and AIST. License: CeCILL-C. *) -Require Import Reals. From mathcomp Require Import ssreflect ssrfun ssrbool ssrnat eqtype choice. From mathcomp Require Import ssralg ssrnum fintype bigop matrix interval. -Require Import boolp reals Rstruct Rbar. -Require Import classical_sets posnum topology hierarchy landau forms. +Require Import boolp reals classical_sets posnum topology hierarchy landau. +Require Import forms. Set Implicit Arguments. Unset Strict Implicit. @@ -11,6 +10,8 @@ Unset Printing Implicit Defensive. Import GRing.Theory Num.Def Num.Theory. Local Open Scope classical_set_scope. +Local Open Scope fun_scope. +Local Open Scope ring_scope. (******************************************************************************) (* This file provides a theory of differentiation. It includes the standard *) @@ -42,7 +43,7 @@ Reserved Notation "f ^` ( n )" (at level 8, format "f ^` ( n )"). Section Differential. -Context {K : absRingType} {V W : normedModType K}. +Context {K : realFieldType} {V W : normedModType K}. Definition diff (F : filter_on V) (_ : phantom (set (set V)) F) (f : V -> W) := (get (fun (df : {linear V -> W}) => continuous df /\ forall x, f x = f (lim F) + df (x - lim F) +o_(x \near F) (x - lim F))). @@ -135,14 +136,14 @@ Notation "'is_diff' F" := (is_diff_def (Phantom _ [filter of F])) Hint Extern 0 (differentiable _ _) => solve[apply: ex_diff] : core. Hint Extern 0 ({for _, continuous _}) => exact: diff_continuous : core. -Lemma differentiableP (V W : normedModType R) (f : V -> W) x : - differentiable f x -> is_diff x f ('d f x). +Lemma differentiableP (R : realFieldType) (V W : normedModType R) (f : V -> W) + x : differentiable f x -> is_diff x f ('d f x). Proof. by move=> ?; apply: DiffDef. Qed. Section jacobian. -Definition jacobian n m (R : absRingType) (f : 'rV[R]_n.+1 -> 'rV[R]_m.+1) p := - lin1_mx ('d f p). +Definition jacobian n m (R : realFieldType) (f : 'rV[R]_n.+1 -> 'rV[R]_m.+1) + p := lin1_mx ('d f p). End jacobian. @@ -150,7 +151,7 @@ Notation "''J' f p" := (jacobian f p). Section DifferentialR. -Context {V W : normedModType R}. +Context (R : realFieldType) {V W : normedModType R}. (* split in multiple bits: - a linear map which is locally bounded is a little o of 1 @@ -161,8 +162,9 @@ Lemma differentiable_continuous (x : V) (f : V -> W) : Proof. move=> /diff_locallyP [dfc]; rewrite -addrA. rewrite (littleo_bigO_eqo (cst (1 : R^o))); last first. - apply/eqOP; near=> k; rewrite /cst [`|[1 : R^o]|]absr1 mulr1. - near=> y; rewrite ltrW //; near: y; apply/locally_normP. + apply/eqOP; near=> k; rewrite /cst [`|[_]|]normr1 mulr1. + near=> y; rewrite ltrW //; near: y. + rewrite !near_simpl -locally_nearE -filter_from_norm_locally. by exists k; [near: k; exists 0|move=> ? /=; rewrite sub0r normmN]. rewrite addfo; first by move=> /eqolim; rewrite flim_shift add0r. by apply/eqolim0P; apply: (flim_trans (dfc 0)); rewrite linear0. @@ -231,11 +233,11 @@ have /eqolimP := df; rewrite -[lim _]/(derive _ _ _). move=> /eqaddoP /(_ e%:num) /(_ [gt0 of e%:num]). apply: filter_app; rewrite /= !near_simpl near_withinE; near=> h => hN0. rewrite /= opprD -![(_ + _ : _ -> _) _]/(_ + _) -![(- _ : _ -> _) _]/(- _). -rewrite /cst /= [`|[1 : R^o]|]absr1 mulr1 => dfv. +rewrite /cst /= [`|[1 : R^o]|]normr1 mulr1 => dfv. rewrite addrA -[X in X + _]scale1r -(@mulVf _ h) //. rewrite mulrC -scalerA -scalerBr normmZ. -rewrite -ler_pdivl_mull; last by rewrite absRE normr_gt0. -by rewrite mulrCA mulVf ?mulr1; last by rewrite absr_eq0. +rewrite -ler_pdivl_mull; last by rewrite normr_gt0. +by rewrite mulrCA mulVf ?mulr1; last by rewrite normr_eq0. Grab Existential Variables. all: end_near. Qed. Lemma derivable_locallyP (f : V -> W) a v : @@ -249,10 +251,9 @@ apply/(@eqolimP _ _ _ (locally'_filter_on _))/eqaddoP => _/posnumP[e]. have /eqaddoP /(_ e%:num) /(_ [gt0 of e%:num]) := df. rewrite /= !(near_simpl, near_withinE); apply: filter_app; near=> h. rewrite /= opprD -![(_ + _ : _ -> _) _]/(_ + _) -![(- _ : _ -> _) _]/(- _). -rewrite /cst /= [`|[1 : R^o]|]absr1 mulr1 addrA => dfv hN0. +rewrite /cst /= [`|[1 : R^o]|]normr1 mulr1 addrA => dfv hN0. rewrite -[X in _ - X]scale1r -(@mulVf _ h) //. -rewrite -scalerA -scalerBr normmZ absRE normfV ler_pdivr_mull ?normr_gt0 //. -by rewrite mulrC. +by rewrite -scalerA -scalerBr normmZ normfV ler_pdivr_mull ?normr_gt0 // mulrC. Grab Existential Variables. all: end_near. Qed. Lemma derivable_locallyx (f : V -> W) a v : @@ -283,29 +284,28 @@ apply: flim_map_lim. pose g1 : R -> W := fun h => (h^-1 * h) *: 'd f a v. pose g2 : R -> W := fun h : R => h^-1 *: k (h *: v ). rewrite (_ : g = g1 + g2) ?funeqE // -(addr0 (_ _ v)); apply: lim_add. - rewrite -(scale1r (_ _ v)); apply: lim_scalel => /= X [e e0]. - rewrite /AbsRing_ball /ball_ /= => eX. - apply/locallyP; rewrite locally_E. - by exists e => //= x _ x0; apply eX; rewrite mulVr // subrr absr0. + rewrite -(scale1r (_ _ v)); apply: lim_scalel. + apply/app_flim_entouragesP => X entX; apply/locallyP. + rewrite locally_E; exists X => // x _ x0; rewrite mulVf //. + exact: entourage_refl. rewrite /g2. have [/eqP ->|v0] := boolP (v == 0). rewrite (_ : (fun _ => _) = cst 0); first exact: cst_continuous. by rewrite funeqE => ?; rewrite scaler0 /k littleo_lim0 // scaler0. apply/flim_normP => e e0. -rewrite nearE /=; apply/locallyP; rewrite locally_E. -have /(littleoP [littleo of k]) /locallyP[i i0 Hi] : 0 < e / (2 * `|[v]|). +rewrite nearE /= locally_simpl /locally' -filter_from_norm_locally. +have /(littleoP [littleo of k]) : 0 < e / (2 * `|[v]|). by rewrite divr_gt0 // pmulr_rgt0 // normm_gt0. +rewrite near_simpl -locally_nearE -filter_from_norm_locally => -[i i0 Hi]. exists (i / `|[v]|); first by rewrite divr_gt0 // normm_gt0. -move=> /= j; rewrite /ball /= /AbsRing_ball /ball_ add0r absrN. -rewrite ltr_pdivl_mulr ?normm_gt0 // => jvi j0. -rewrite add0r normmN normmZ -ltr_pdivl_mull ?normr_gt0 ?invr_neq0 //. -have /Hi/ler_lt_trans -> // : ball 0 i (j *: v). - by rewrite -ball_normE /ball_ add0r normmN (ler_lt_trans _ jvi) // normmZ. +move=> /= j; rewrite normmB subr0 ltr_pdivl_mulr ?normm_gt0 // => jvi j0. +rewrite normmB subr0 normmZ -ltr_pdivl_mull ?normr_gt0 ?invr_neq0 //. +have /Hi/ler_lt_trans -> // : ball norm 0 i (j *: v). + by rewrite /ball add0r normmN (ler_lt_trans _ jvi) // normmZ. rewrite -(mulrC e) -mulrA -ltr_pdivl_mull // mulrA mulVr ?unitfE ?gtr_eqF //. -rewrite absRE normrV ?unitfE // div1r invrK ltr_pdivr_mull; last first. +rewrite normrV ?unitfE // div1r invrK ltr_pdivr_mull; last first. by rewrite pmulr_rgt0 // normm_gt0. -rewrite normmZ absRE mulrC -mulrA. -by rewrite ltr_pmull ?ltr1n // pmulr_rgt0 ?normm_gt0 // normr_gt0. +by rewrite normmZ mulrC -mulrA ltr_pmull ?ltr1n // pmulr_rgt0 normm_gt0. Qed. End DifferentialR. @@ -315,42 +315,43 @@ Notation "''D_' v f c" := (derive f c v). (* printing *) Hint Extern 0 (derivable _ _ _) => solve[apply: ex_derive] : core. Section DifferentialR2. -Implicit Type (V : normedModType R). + +Variable (R : realFieldType) (V : normedModType R). Lemma derivemxE m n (f : 'rV[R]_m.+1 -> 'rV[R]_n.+1) (a v : 'rV[R]_m.+1) : differentiable f a -> 'D_ v f a = v *m jacobian f a. Proof. by move=> /deriveE->; rewrite /jacobian mul_rV_lin1. Qed. -Definition derive1 V (f : R -> V) (a : R) := +Definition derive1 (f : R -> V) (a : R) := lim ((fun h => h^-1 *: (f (h + a) - f a)) @ locally' (0 : R^o)). Local Notation "f ^` ()" := (derive1 f). -Lemma derive1E V (f : R -> V) a : f^`() a = 'D_1 (f : R^o -> _) a. +Lemma derive1E (f : R -> V) a : f^`() a = 'D_1 (f : R^o -> _) a. Proof. rewrite /derive1 /derive; set d := (fun _ : R => _); set d' := (fun _ : R => _). by suff -> : d = d' by []; rewrite funeqE=> h; rewrite /d /d' /= [h%:A](mulr1). Qed. (* Is it necessary? *) -Lemma derive1E' V f a : differentiable (f : R^o -> V) a -> +Lemma derive1E' f a : differentiable (f : R^o -> V) a -> f^`() a = 'd f a 1. Proof. by move=> ?; rewrite derive1E deriveE. Qed. -Definition derive1n V n (f : R -> V) := iter n (@derive1 V) f. +Definition derive1n n (f : R -> V) := iter n derive1 f. Local Notation "f ^` ( n )" := (derive1n n f). -Lemma derive1n0 V (f : R -> V) : f^`(0) = f. +Lemma derive1n0 (f : R -> V) : f^`(0) = f. Proof. by []. Qed. -Lemma derive1n1 V (f : R -> V) : f^`(1) = f^`(). +Lemma derive1n1 (f : R -> V) : f^`(1) = f^`(). Proof. by []. Qed. -Lemma derive1nS V (f : R -> V) n : f^`(n.+1) = f^`(n)^`(). +Lemma derive1nS (f : R -> V) n : f^`(n.+1) = f^`(n)^`(). Proof. by []. Qed. -Lemma derive1Sn V (f : R -> V) n : f^`(n.+1) = f^`()^`(n). +Lemma derive1Sn (f : R -> V) n : f^`(n.+1) = f^`()^`(n). Proof. exact: iterSr. Qed. End DifferentialR2. @@ -360,6 +361,8 @@ Notation "f ^` ( n )" := (derive1n n f). Section DifferentialR3. +Variable (R : realFieldType). + Lemma littleo_linear0 (V W : normedModType R) (f : {linear V -> W}) : (f : V -> W) =o_ (0 : V) id -> f = cst 0 :> (V -> W). Proof. @@ -368,15 +371,15 @@ rewrite funeqE => x; apply/eqP; case: (ler0P `|[x]|) => [|xn0]. by rewrite normm_le0 => /eqP ->; rewrite linear0. rewrite -normm_le0 -(mul0r `|[x]|) -ler_pdivr_mulr //. apply/ler0_addgt0P => _ /posnumP[e]; rewrite ler_pdivr_mulr //. -have /oid /locallyP [_ /posnumP[d] dfe] := posnum_gt0 e. -set k := ((d%:num / 2) / (PosNum xn0)%:num)^-1. +have /oid := posnum_gt0 e. +rewrite !near_simpl -locally_nearE -filter_from_norm_locally. +move=> [_ /posnumP[d] dfe]; set k := ((d%:num / 2) / (PosNum xn0)%:num)^-1. rewrite -{1}(@scalerKV _ _ k _ x) // linearZZ normmZ. -rewrite -ler_pdivl_mull; last by rewrite absRE gtr0_norm. -rewrite mulrCA (@ler_trans _ (e%:num * `|[k^-1 *: x]|)) //; last first. - by rewrite ler_pmul // normmZ absRE normfV. -apply dfe. -rewrite -ball_normE /ball_ sub0r normmN normmZ. -rewrite invrK -ltr_pdivl_mulr // absRE ger0_norm // ltr_pdivr_mulr //. +rewrite -ler_pdivl_mull ?gtr0_norm // mulrCA. +rewrite (@ler_trans _ (e%:num * `|[k^-1 *: x]|)) //; last first. + by rewrite ler_pmul // normmZ normfV gtr0_norm. +apply dfe; rewrite /ball sub0r normmN normmZ. +rewrite invrK -ltr_pdivl_mulr // ger0_norm // ltr_pdivr_mulr //. by rewrite -mulrA mulVf ?lt0r_neq0 // mulr1 [X in _ < X]splitr ltr_addl. Qed. @@ -391,8 +394,7 @@ have hdf h : (f \o shift x = cst (f x) + h +o_ (0 : V) id) -> h = f \o shift x - cst (f x) +o_ (0 : V) id. move=> hdf; apply: eqaddoE. - rewrite hdf -addrA addrCA [cst _ + _]addrC addrK [_ + h]addrC. - rewrite -addrA -[LHS]addr0; congr (_ + _). + rewrite hdf -addrA addrACA [cst _ + _]addrC addrK -[LHS]addr0; congr (_ + _). by apply/eqP; rewrite eq_sym addrC addr_eq0 oppo. rewrite (hdf _ dxf). suff /diff_locally /hdf -> : differentiable f x. @@ -612,19 +614,19 @@ Lemma linear_lipschitz (V' W' : normedModType R) (f : {linear V' -> W'}) : continuous f -> exists2 k, k > 0 & forall x, `|[f x]| <= k * `|[x]|. Proof. move=> /(_ 0); rewrite linear0 => /(_ _ (locally_ball 0 1%:pos)). -move=> /locallyP [_ /posnumP[e] he]; exists (2 / e%:num) => // x. +rewrite !near_simpl -locally_nearE -filter_from_norm_locally. +move=> [_ /posnumP[e] he]; exists (2 / e%:num) => // x. case: (lerP `|[x]| 0) => [|xn0]. by rewrite normm_le0 => /eqP->; rewrite linear0 !normm0 mulr0. set k := 2 / e%:num * (PosNum xn0)%:num. have kn0 : k != 0 by []. have abskgt0 : `|k| > 0 by rewrite normr_gt0. rewrite -[x in X in X <= _](scalerKV kn0) linearZZ normmZ -ler_pdivl_mull //. -suff /he : ball 0 e%:num (k^-1 *: x). - rewrite -ball_normE /= normmB subr0 => /ltrW /ler_trans; apply. - by rewrite absRE ger0_norm // mulVf. -rewrite -ball_normE /= normmB subr0 normmZ. -rewrite absRE normfV ger0_norm // invrM ?unitfE // mulrAC mulVf //. -by rewrite invf_div mul1r [X in _ < X]splitr; apply: ltr_spaddr. +suff /he : ball norm 0 e%:num (k^-1 *: x). + rewrite /ball normmB subr0 => /ltrW /ler_trans; apply. + by rewrite ger0_norm // mulVf. +rewrite /ball /= normmB subr0 normmZ normfV ger0_norm // invfM -mulrA mulVf //. +by rewrite invf_div mulr1 [X in _ < X]splitr; apply: ltr_spaddr. Qed. Lemma linear_eqO (V' W' : normedModType R) (f : {linear V' -> W'}) : @@ -639,47 +641,48 @@ Lemma diff_eqO (V' W' : normedModType R) (x : filter_on V') (f : V' -> W') : Proof. by move=> /diff_continuous /linear_eqO; apply. Qed. (* TODO: generalize *) -Lemma compoO_eqo (K : absRingType) (U V' W' : normedModType K) (f : U -> V') +Lemma compoO_eqo (K : realFieldType) (U V' W' : normedModType K) (f : U -> V') (g : V' -> W') : [o_ (0 : V') id of g] \o [O_ (0 : U) id of f] =o_ (0 : U) id. Proof. apply/eqoP => _ /posnumP[e]. have /bigO_exP [_ /posnumP[k]] := bigOP [bigO of [O_ (0 : U) id of f]]. -have := littleoP [littleo of [o_ (0 : V') id of g]]. -move=> /(_ (e%:num / k%:num)) /(_ _) /locallyP [//|_ /posnumP[d] hd]. +have /(_ (e%:num / k%:num)) := littleoP [littleo of [o_ (0 : V') id of g]]. +rewrite !near_simpl -locally_nearE -filter_from_norm_locally. +move=> /(_ _) [//|_/posnumP[d] hd]. apply: filter_app; near=> x => leOxkx; apply: ler_trans (hd _ _) _; last first. rewrite -ler_pdivl_mull //; apply: ler_trans leOxkx _. by rewrite invf_div mulrA -[_ / _ * _]mulrA mulVf // mulr1. -rewrite -ball_normE /= normmB subr0 (ler_lt_trans leOxkx) //. -rewrite -ltr_pdivl_mull //; near: x; rewrite /= !locally_simpl. -apply/locallyP; exists (k%:num ^-1 * d%:num)=> // x. -by rewrite -ball_normE /= normmB subr0. +rewrite /ball normmB subr0 (ler_lt_trans leOxkx) //. +rewrite -ltr_pdivl_mull //; near: x. +rewrite -locally_nearE -filter_from_norm_locally. +by exists (k%:num ^-1 * d%:num) => // x; rewrite /ball normmB subr0. Grab Existential Variables. all: end_near. Qed. -Lemma compoO_eqox (K : absRingType) (U V' W' : normedModType K) (f : U -> V') +Lemma compoO_eqox (K : realFieldType) (U V' W' : normedModType K) (f : U -> V') (g : V' -> W') : forall x : U, [o_ (0 : V') id of g] ([O_ (0 : U) id of f] x) =o_(x \near 0 : U) x. Proof. by move=> x; rewrite -[X in X = _]/((_ \o _) x) compoO_eqo. Qed. (* TODO: generalize *) -Lemma compOo_eqo (K : absRingType) (U V' W' : normedModType K) (f : U -> V') +Lemma compOo_eqo (K : realFieldType) (U V' W' : normedModType K) (f : U -> V') (g : V' -> W') : [O_ (0 : V') id of g] \o [o_ (0 : U) id of f] =o_ (0 : U) id. Proof. apply/eqoP => _ /posnumP[e]. have /bigO_exP [_ /posnumP[k]] := bigOP [bigO of [O_ (0 : V') id of g]]. -move=> /locallyP [_ /posnumP[d] hd]. -have ekgt0 : e%:num / k%:num > 0 by []. +rewrite !near_simpl -locally_nearE -filter_from_norm_locally. +move=> [_/posnumP[d] hd]; have ekgt0 : e%:num / k%:num > 0 by []. have /(_ _ ekgt0) := littleoP [littleo of [o_ (0 : U) id of f]]. apply: filter_app; near=> x => leoxekx; apply: ler_trans (hd _ _) _; last first. by rewrite -ler_pdivl_mull // mulrA [_^-1 * _]mulrC. -rewrite -ball_normE /= normmB subr0; apply: ler_lt_trans leoxekx _. -rewrite -ltr_pdivl_mull //; near: x; rewrite /= locally_simpl. -apply/locallyP; exists ((e%:num / k%:num) ^-1 * d%:num)=> // x. -by rewrite -ball_normE /= normmB subr0. +rewrite /ball normmB subr0; apply: ler_lt_trans leoxekx _. +rewrite -ltr_pdivl_mull //; near: x. +rewrite !near_simpl -locally_nearE -filter_from_norm_locally. +by exists ((e%:num / k%:num) ^-1 * d%:num)=> // x; rewrite /ball normmB subr0. Grab Existential Variables. all: end_near. Qed. -Lemma compOo_eqox (K : absRingType) (U V' W' : normedModType K) (f : U -> V') +Lemma compOo_eqox (K : realFieldType) (U V' W' : normedModType K) (f : U -> V') (g : V' -> W') : forall x, [O_ (0 : V') id of g] ([o_ (0 : U) id of f] x) =o_(x \near 0 : U) x. Proof. by move=> x; rewrite -[X in X = _]/((_ \o _) x) compOo_eqo. Qed. @@ -722,30 +725,31 @@ Lemma bilinear_schwarz (U V' W' : normedModType R) exists2 k, k > 0 & forall u v, `|[f u v]| <= k * `|[u]| * `|[v]|. Proof. move=> /(_ 0); rewrite linear0r => /(_ _ (locally_ball 0 1%:pos)). -move=> /locallyP [_ /posnumP[e] he]; exists ((2 / e%:num) ^+2) => // u v. +rewrite !near_simpl -locally_nearE -!filter_from_norm_locally. +move=> [[A B] /= [[_/posnumP[eA] sA] [_/posnumP[eB] sB]] sAB]. +set e := minr eA%:num eB%:num; exists ((2 / e) ^+2) => // u v. case: (lerP `|[u]| 0) => [|un0]. by rewrite normm_le0 => /eqP->; rewrite linear0l !normm0 mulr0 mul0r. case: (lerP `|[v]| 0) => [|vn0]. by rewrite normm_le0 => /eqP->; rewrite linear0r !normm0 mulr0. rewrite -[`|[u]|]/((PosNum un0)%:num) -[`|[v]|]/((PosNum vn0)%:num). -set ku := 2 / e%:num * (PosNum un0)%:num. -set kv := 2 / e%:num * (PosNum vn0)%:num. -rewrite -[X in f X](@scalerKV _ _ ku) // linearZl_LR normmZ. -rewrite absRE gtr0_norm // -ler_pdivl_mull //. -rewrite -[X in f _ X](@scalerKV _ _ kv) // linearZr_LR normmZ. -rewrite absRE gtr0_norm // -ler_pdivl_mull //. -suff /he : ball 0 e%:num (ku^-1 *: u, kv^-1 *: v). - rewrite -ball_normE /= normmB subr0 => /ltrW /ler_trans; apply. - rewrite ler_pdivl_mull // mulr1 ler_pdivl_mull //. - by rewrite mulrA [ku * _]mulrAC expr2. -rewrite -ball_normE /= normmB subr0. +set ku := 2 / e * (PosNum un0)%:num. +set kv := 2 / e * (PosNum vn0)%:num. +rewrite -[X in f X](@scalerKV _ _ ku) // linearZl_LR normmZ gtr0_norm //. +rewrite -ler_pdivl_mull // -[X in f _ X](@scalerKV _ _ kv) // linearZr_LR. +rewrite normmZ gtr0_norm // -ler_pdivl_mull //. +suff : ball norm 0 e (ku^-1 *: u, kv^-1 *: v). + rewrite /ball ltr_maxl !ltr_minr => /andP [/andP[/sA uA _] /andP[_ /sB vB]]. + have /sAB /= := conj uA vB; rewrite normmB subr0 => /ltrW /ler_trans; apply. + by rewrite ler_pdivl_mull// mulr1 ler_pdivl_mull// mulrA [ku * _]mulrAC expr2. +rewrite /ball normmB subr0. have -> : (ku^-1 *: u, kv^-1 *: v) = - (e%:num / 2) *: ((PosNum un0)%:num ^-1 *: u, (PosNum vn0)%:num ^-1 *: v). + (e / 2) *: ((PosNum un0)%:num ^-1 *: u, (PosNum vn0)%:num ^-1 *: v). rewrite invrM ?unitfE // [kv ^-1]invrM ?unitfE //. rewrite mulrC -[_ *: u]scalerA [X in X *: v]mulrC -[_ *: v]scalerA. by rewrite invf_div. -rewrite normmZ absRE ger0_norm // -mulrA gtr_pmulr // ltr_pdivr_mull // mulr1. -by rewrite ltr_maxl !normmZ !absRE !ger0_norm // !mulVf // ltr1n. +rewrite normmZ ger0_norm // -mulrA gtr_pmulr // ltr_pdivr_mull // mulr1. +by rewrite ltr_maxl !normmZ !ger0_norm // !mulVf // ltr1n. Qed. Lemma bilinear_eqo (U V' W' : normedModType R) (f : {bilinear U -> V' -> W'}) : @@ -756,7 +760,7 @@ apply/eqoP=> _ /posnumP[e]; near=> x; rewrite (ler_trans (fschwarz _ _))//. rewrite ler_pmul ?pmulr_rge0 //; last by rewrite ler_maxr orbC lerr. rewrite -ler_pdivl_mull //. suff : `|[x]| <= k%:num ^-1 * e%:num by apply: ler_trans; rewrite ler_maxr lerr. -near: x; rewrite !near_simpl; apply/locally_le_locally_norm. +near: x; rewrite !near_simpl -locally_nearE -filter_from_norm_locally. by exists (k%:num ^-1 * e%:num) => // ? /=; rewrite normmB subr0 => /ltrW. Grab Existential Variables. all: end_near. Qed. @@ -767,9 +771,19 @@ Fact dbilin (U V' W' : normedModType R) (f : {bilinear U -> V' -> W'}) p : (fun q => f p.1 q.2 + f q.1 p.2) +o_ (0 : U * V') id. Proof. move=> fc; split=> [q|]. - by apply: (@continuousD _ _ _ (fun q => f p.1 q.2) (fun q => f q.1 p.2)); - move=> A /(fc (_.1, _.2)) /= /locallyP [_ /posnumP[e] fpqe_A]; - apply/locallyP; exists e%:num => // r [??]; apply: (fpqe_A (_.1, _.2)). + apply: (@continuousD _ _ _ (fun q => f p.1 q.2) (fun q => f q.1 p.2)). + move=> A /(fc (_.1, _.2)). + rewrite !near_simpl -!locally_nearE -!filter_from_norm_locally /=. + move=> [PQ [[_/posnumP[eP] sP] [_/posnumP[eQ] sQ]] sPQ]. + exists (setT, PQ.2); first by split; [exists 1|exists eQ%:num]. + move=> xy [_ /= Qy]; apply: (sPQ (_.1,_.2)); split=> //=. + exact/sP/ball_center. + move=> A /(fc (_.1, _.2)). + rewrite !near_simpl -!locally_nearE -!filter_from_norm_locally /=. + move=> [PQ [[_/posnumP[eP] sP] [_/posnumP[eQ] sQ]] sPQ]. + exists (PQ.1, setT); first by split; [exists eP%:num|exists 1]. + move=> xy [/= Px _]; apply: (sPQ (_.1,_.2)); split=> //=. + exact/sQ/ball_center. apply/eqaddoE; rewrite funeqE => q /=. rewrite linearDl !linearDr addrA addrC. rewrite -[f q.1 _ + _ + _]addrA [f q.1 _ + _]addrC addrA [f q.1 _ + _]addrC. @@ -793,20 +807,18 @@ Proof. by move=> fc; apply/diff_locallyP; rewrite diff_bilin //; apply: dbilin p fc. Qed. -Definition Rmult_rev (y x : R) := x * y. -Canonical rev_Rmult := @RevOp _ _ _ Rmult_rev (@GRing.mul [ringType of R]) - (fun _ _ => erefl). +Definition mulr_rev (y x : R) := x * y. +Canonical rev_Rmult := @RevOp _ _ _ mulr_rev *%R (fun _ _ => erefl). -Lemma Rmult_is_linear x : linear (@GRing.mul [ringType of R] x : R^o -> R^o). +Lemma mulr_is_linear (x : R^o) : linear (GRing.mul x). Proof. by move=> ???; rewrite mulrDr scalerAr. Qed. -Canonical Rmult_linear x := Linear (Rmult_is_linear x). +Canonical mulr_linear x := Linear (mulr_is_linear x). -Lemma Rmult_rev_is_linear y : linear (Rmult_rev y : R^o -> R^o). -Proof. by move=> ???; rewrite /Rmult_rev mulrDl scalerAl. Qed. -Canonical Rmult_rev_linear y := Linear (Rmult_rev_is_linear y). +Lemma mulr_rev_is_linear y : linear (mulr_rev y : R^o -> R^o). +Proof. by move=> ???; rewrite /mulr_rev mulrDl scalerAl. Qed. +Canonical mulr_rev_linear y := Linear (mulr_rev_is_linear y). -Canonical Rmult_bilinear := - [bilinear of (@GRing.mul [ringType of [lmodType R of R^o]])]. +Canonical mulr_bilinear := [bilinear of (@GRing.mul [ringType of R^o])]. Global Instance is_diff_Rmult (p : R^o * R^o) : is_diff p (fun q => q.1 * q.2) (fun q => p.1 * q.2 + q.1 * p.2). @@ -815,7 +827,7 @@ apply: DiffDef; first by apply: differentiable_bilin =>?; apply: lim_mult. by rewrite diff_bilin // => ?; apply: lim_mult. Qed. -Lemma eqo_pair (K : absRingType) (U V' W' : normedModType K) (F : filter_on U) +Lemma eqo_pair (K : realFieldType) (U V' W' : normedModType K) (F : filter_on U) (f : U -> V') (g : U -> W') : (fun t => ([o_F id of f] t, [o_F id of g] t)) =o_F id. Proof. @@ -900,33 +912,35 @@ rewrite -[X in X + _]mulr1 -[X in 1 / _ * X](@mulfVK _ (x ^+ 2)); last first. by rewrite sqrf_eq0. rewrite mulrA mulf_div mulr1. have hDx_neq0 : h + x != 0. - near: h; rewrite !locally_simpl; apply/locally_normP. + near: h; rewrite !near_simpl -locally_nearE -filter_from_norm_locally. exists `|x|; first by rewrite normr_gt0. move=> h /=; rewrite normmB subr0 -subr_gt0 => lthx. rewrite -(normm_gt0 (h + x : R^o)) addrC -[h]opprK. apply: ltr_le_trans (ler_distm_dist _ _). - by rewrite absRE ger0_norm normmN //; apply: ltrW. + by rewrite ger0_norm normmN //; apply: ltrW. rewrite addrC -[X in X * _]mulr1 -{2}[1](@mulfVK _ (h + x)) //. rewrite mulrA expr_div_n expr1n mulf_div mulr1 [_ ^+ 2 * _]mulrC -mulrA. -rewrite -mulrDr mulrBr [1 / _ * _]mulrC [`|[ _ ]|]absRE normrM. +rewrite -mulrDr mulrBr [1 / _ * _]mulrC [`|[_]|]normrM. rewrite mulrDl mulrDl opprD addrACA addrA [x * _]mulrC expr2. do 2 ?[rewrite -addrA [- _ + _]addrC subrr addr0]. rewrite div1r normfV [X in _ / X]normrM invfM [X in _ * X]mulrC. rewrite mulrA mulrAC ler_pdivr_mulr ?normr_gt0 ?mulf_neq0 //. rewrite mulrAC ler_pdivr_mulr ?normr_gt0 //. have : `|h * h| <= `|x / 2| * (e%:num * `|x * x| * `|[h : R^o]|). - rewrite !mulrA; near: h; exists (`|x / 2| * e%:num * `|x * x|). + rewrite !mulrA; near: h. + rewrite !near_simpl -locally_nearE -filter_from_norm_locally. + exists (`|x / 2| * e%:num * `|x * x|). by rewrite !pmulr_rgt0 // normr_gt0 mulf_neq0. - by move=> h /ltrW; rewrite absrB subr0 [`|h * _|]normrM => /ler_pmul; apply. + by move=> h /ltrW; rewrite normmB subr0 [`|h * _|]normrM => /ler_pmul; apply. move=> /ler_trans-> //; rewrite [X in X <= _]mulrC ler_pmul ?mulr_ge0 //. -near: h; exists (`|x| / 2); first by rewrite divr_gt0 ?normr_gt0. -move=> h; rewrite /AbsRing_ball /= absrB subr0 => lthhx; rewrite addrC -[h]opprK. -apply: ler_trans (@ler_distm_dist _ [normedModType R of R^o] _ _). -rewrite absRE normmN [X in _ <= X]ger0_norm; last first. - rewrite subr_ge0; apply: ltrW; apply: ltr_le_trans lthhx _. - by rewrite [`|[_]|]splitr ler_addl; apply: divr_ge0. -rewrite ler_subr_addr -ler_subr_addl (splitr `|[x : R^o]|). -by rewrite normrM normfV (@ger0_norm _ 2) // -addrA subrr addr0; apply: ltrW. +near: h; rewrite !near_simpl -locally_nearE -filter_from_norm_locally. +exists (`|x| / 2); first by rewrite divr_gt0 ?normr_gt0. +move=> h; rewrite /ball normmB subr0 => lthhx; rewrite addrC -[h]opprK. +apply: ler_trans (ler_distm_dist _ _); rewrite normmN [X in _ <= X]ger0_norm. + rewrite ler_subr_addr -ler_subr_addl (splitr `|[x : R^o]|). + by rewrite normrM normfV (@ger0_norm _ 2) // -addrA subrr addr0; apply: ltrW. +rewrite subr_ge0; apply: ltrW; apply: ltr_le_trans lthhx _. +by rewrite [`|[_]|]splitr ler_addl; apply: divr_ge0. Grab Existential Variables. all: end_near. Qed. Lemma diff_Rinv (x : R^o) : x != 0 -> @@ -969,7 +983,7 @@ Proof. by move=> df fxn0; apply: differentiable_comp _ (differentiable_Rinv fxn0). Qed. -Lemma exprfunE (T : pointedType) (R : ringType) (f : T -> R) n : +Lemma exprfunE (T : pointedType) (K : ringType) (f : T -> K) n : f ^+ n = (fun x => f x ^+ n). Proof. by elim: n => [|n ihn]; rewrite funeqE=> ?; [rewrite !expr0|rewrite !exprS ihn]. @@ -998,7 +1012,7 @@ End DifferentialR3. Section Derive. -Variable (V W : normedModType R). +Variable (R : realFieldType) (V W : normedModType R). Let der1 (U : normedModType R) (f : R^o -> U) x : derivable f x 1 -> f \o shift x = cst (f x) + ( *:%R^~ (f^`() x)) +o_ (0 : R^o) id. @@ -1200,7 +1214,7 @@ apply: lim_add; last exact: lim_scaler df. apply: flim_comp2 (@lim_mult _ _ _) => /=; last exact: dg. suff : {for 0, continuous (fun h => f(h *: v + x))}. by move=> /continuous_withinNx; rewrite scale0r add0r. -exact/(@differentiable_continuous _ _ (0 : R^o))/derivable1_diffP/derivable1P. +exact/differentiable_continuous/derivable1_diffP/derivable1P. Qed. Lemma deriveM (f g : V -> R^o) (x v : V) : @@ -1251,8 +1265,9 @@ move=> fxn0 df. have /derivable1P/derivable1_diffP/differentiable_continuous := df. move=> /continuous_withinNx; rewrite scale0r add0r => fc. have fn0 : locally' (0 : R^o) [set h | f (h *: v + x) != 0]. - apply: (fc [set x | x != 0]); exists `|[f x]|; first by rewrite normm_gt0. - move=> y; rewrite /AbsRing_ball /= => yltfx. + apply: (fc [set x | x != 0]); rewrite /= -filter_from_norm_locally. + exists `|[f x]|; first by rewrite normm_gt0. + move=> y; rewrite /ball => yltfx. by apply/eqP => y0; move: yltfx; rewrite y0 subr0 ltrr. have : (fun h => - ((1 / f x) * (1 / f (h *: v + x))) *: (h^-1 *: (f (h *: v + x) - f x))) @ locally' (0 : R^o) --> @@ -1260,10 +1275,15 @@ have : (fun h => - ((1 / f x) * (1 / f (h *: v + x))) *: apply: flim_comp2 (@lim_mult _ _ _) => //=. apply: (@lim_opp _ [normedModType R of R^o]); rewrite expr2. exact/lim_scaler/lim_inv. -apply: flim_trans => A [_/posnumP[e] /= Ae]. -move: fn0; apply: filter_app; near=> h => /=. -move=> fhvxn0; have he : AbsRing_ball 0 e%:num h by near: h; exists e%:num. -have hn0 : h != 0 by near: h; exists e%:num. +apply: flim_trans => A. +rewrite {1}/locally' !near_simpl -locally_nearE -filter_from_norm_locally. +move=> [_/posnumP[e] /= Ae]; move: fn0; apply: filter_app; near=> h => /=. +move=> fhvxn0; have he : ball norm 0 e%:num h. + near: h; rewrite /locally' -locally_nearE -filter_from_norm_locally. + by exists e%:num. +have hn0 : h != 0. + near: h; rewrite /locally' -locally_nearE -filter_from_norm_locally. + by exists e%:num. suff <- : - ((1 / f x) * (1 / f (h *: v + x))) *: (h^-1 *: (f (h *: v + x) - f x)) = h^-1 *: (1 / f (h *: v + x) - 1 / f x) by exact: Ae. @@ -1286,7 +1306,7 @@ Qed. End Derive. -Lemma EVT_max (f : R -> R) (a b : R) : +Lemma EVT_max (R : realType) (f : R -> R) (a b : R) : a <= b -> {in `[a, b], continuous f} -> exists2 c, c \in `[a, b] & forall t, t \in `[a, b] -> f t <= f c. Proof. @@ -1298,8 +1318,7 @@ have imf_sup : has_sup imf. apply/compact_bounded/continuous_compact; last exact: segment_compact. by move=> ?; rewrite inE => /asboolP /fcont. exists (M + 1); apply/ubP => y; rewrite !inE => /asboolP /imfltM yltM. - apply/ltrW; apply: ler_lt_trans (yltM _ _); last by rewrite ltr_addl. - by rewrite [ `|[_]| ]absRE ler_norm. + by apply/ltrW; apply: ler_lt_trans (yltM _ _); rewrite ?ler_norm // ltr_addl. case: (pselect (exists2 c, c \in `[a, b] & f c = sup imf)) => [|imf_ltsup]. move=> [c cab fceqsup]; exists c => // t tab. rewrite fceqsup; apply: sup_upper_bound=> //; rewrite !inE; apply/asboolP. @@ -1310,7 +1329,7 @@ have {imf_ltsup} imf_ltsup : forall t, t \in `[a, b] -> f t < sup imf. apply/eqP; rewrite eqr_le supleft sup_upper_bound => //. by rewrite !inE; apply/asboolP/imageP. have invf_cont : {in `[a, b], continuous (fun t => 1 / (sup imf - f t))}. - move=> t tab; apply: lim_inv. + move=> t tab; apply: (@lim_inv _ _ (locally t)). by rewrite neqr_lt subr_gt0 orbC imf_ltsup. by apply: lim_add; [apply: continuous_cst|apply/lim_opp/fcont]. have [M imVfltM] : bounded ((fun t => 1 / (sup imf - f t)) @` @@ -1324,7 +1343,7 @@ move=> [y]; rewrite !inE => /asboolP [t tab <-] {y}. rewrite ltr_subl_addr - ltr_subl_addl. suff : sup imf - f t > k^-1 by move=> /ltrW; rewrite lerNgt => /negbTE ->. rewrite -[X in _ < X]invrK ltr_pinv. - rewrite -div1r; apply: ler_lt_trans (ler_norm _) _; rewrite -absRE. + rewrite -div1r; apply: ler_lt_trans (ler_norm _) _. by apply: imVfltM; [rewrite ltr_maxr ltr_addl ltr01|apply: imageP]. by rewrite inE kgt0 unitfE lt0r_neq0. have invsupft_gt0 : 0 < (sup imf - f t)^-1. @@ -1332,7 +1351,7 @@ have invsupft_gt0 : 0 < (sup imf - f t)^-1. by rewrite inE invsupft_gt0 unitfE lt0r_neq0. Qed. -Lemma EVT_min (f : R -> R) (a b : R) : +Lemma EVT_min (R : realType) (f : R -> R) (a b : R) : a <= b -> {in `[a, b], continuous f} -> exists2 c, c \in `[a, b] & forall t, t \in `[a, b] -> f c <= f t. Proof. @@ -1342,42 +1361,41 @@ have /(EVT_max leab) [c clr fcmax] : {in `[a, b], continuous (- f)}. by exists c => // ? /fcmax; rewrite ler_opp2. Qed. -Lemma cvg_at_rightE (V : normedModType R) (f : R -> V) x : +Lemma cvg_at_rightE (R : realType) (V : normedModType R) (f : R -> V) x : cvg (f @ locally' x) -> lim (f @ locally' x) = lim (f @ at_right x). Proof. -move=> cvfx; apply/Logic.eq_sym. -(* should be inferred *) -have atrF := at_right_proper_filter x. -apply: flim_map_lim => A /cvfx /locallyP [_ /posnumP[e] xe_A]. -by exists e%:num => // y xe_y; rewrite ltr_def => /andP [xney _]; apply: xe_A. +move=> cvfx; apply/esym. (* should be inferred *) +have atrF := at_right_proper_filter x; apply: flim_map_lim => A /cvfx. +rewrite /locally' /at_right -filter_from_norm_locally. +move=> [_/posnumP[e] xe_A]; exists e%:num => // y xe_y. +by rewrite ltr_def => /andP[xney _]; apply: xe_A. Qed. -Lemma cvg_at_leftE (V : normedModType R) (f : R -> V) x : +Lemma cvg_at_leftE (R : realType) (V : normedModType R) (f : R -> V) x : cvg (f @ locally' x) -> lim (f @ locally' x) = lim (f @ at_left x). Proof. -move=> cvfx; apply/Logic.eq_sym. -(* should be inferred *) -have atrF := at_left_proper_filter x. -apply: flim_map_lim => A /cvfx /locallyP [_ /posnumP[e] xe_A]. -exists e%:num => // y xe_y; rewrite ltr_def => /andP [xney _]. -by apply: xe_A => //; rewrite eq_sym. +move=> cvfx; apply/esym. (* should be inferred *) +have atrF := at_left_proper_filter x; apply: flim_map_lim => A /cvfx. +rewrite /locally' /at_left -filter_from_norm_locally. +move=> [_ /posnumP[e] xe_A]; exists e%:num => // y xe_y. +by rewrite ltr_def => /andP [xney _]; apply: xe_A => //; rewrite eq_sym. Qed. -Lemma le0r_flim_map (T : topologicalType) (F : set (set T)) +Lemma le0r_flim_map (R : realFieldType) (T : topologicalType) (F : set (set T)) (FF : ProperFilter F) (f : T -> R^o) : (\forall x \near F, 0 <= f x) -> cvg (f @ F) -> 0 <= lim (f @ F). Proof. move=> fge0 fcv; case: (lerP 0 (lim (f @ F))) => // limlt0; near F => x. have := near fge0 x; rewrite lerNgt => /(_ _) /negbTE<- //; near: x. have normlimgt0 : `|[lim (f @ F)]| > 0 by rewrite normm_gt0 ltr0_neq0. -have /fcv := locally_ball_norm (lim (f @ F)) (PosNum normlimgt0). +have /fcv := locally_ball (lim (f @ F)) (PosNum normlimgt0). rewrite /= !near_simpl; apply: filterS => x. -rewrite /= normmB [ `|[_]| ]absRE => /(ler_lt_trans (ler_norm _)). +rewrite /= normmB => /(ler_lt_trans (ler_norm _)). rewrite ltr_subl_addr => /ltr_le_trans; apply. -by rewrite [ `|[_]| ]absRE ltr0_norm // addrC subrr. +by rewrite [`|[_]|]ltr0_norm // addrC subrr. Grab Existential Variables. all: end_near. Qed. -Lemma ler0_flim_map (T : topologicalType) (F : set (set T)) +Lemma ler0_flim_map (R : realFieldType) (T : topologicalType) (F : set (set T)) (FF : ProperFilter F) (f : T -> R^o) : (\forall x \near F, f x <= 0) -> cvg (f @ F) -> lim (f @ F) <= 0. Proof. @@ -1388,8 +1406,8 @@ rewrite limopp; apply: le0r_flim_map; last by rewrite -limopp; apply: lim_opp. by move: fle0; apply: filterS => x; rewrite oppr_ge0. Qed. -Lemma ler_flim_map (T : topologicalType) (F : set (set T)) (FF : ProperFilter F) - (f g : T -> R^o) : +Lemma ler_flim_map (R : realFieldType) (T : topologicalType) (F : set (set T)) + (FF : ProperFilter F) (f g : T -> R^o) : (\forall x \near F, f x <= g x) -> cvg (f @ F) -> cvg (g @ F) -> lim (f @ F) <= lim (g @ F). Proof. @@ -1401,7 +1419,7 @@ rewrite eqlim; apply: le0r_flim_map; last first. by move: lefg; apply: filterS => x; rewrite subr_ge0. Qed. -Lemma derive1_at_max (f : R^o -> R^o) (a b c : R) : +Lemma derive1_at_max (R : realType) (f : R^o -> R^o) (a b c : R) : a <= b -> (forall t, t \in `]a, b[ -> derivable f t 1) -> c \in `]a, b[ -> (forall t, t \in `]a, b[ -> f t <= f c) -> is_derive (c : R^o) 1 f 0. Proof. @@ -1414,10 +1432,12 @@ apply/eqP; rewrite eqr_le; apply/andP; split. move=> A [e egt0 Ae]; exists e => // x xe xgt0; apply: Ae => //. exact/lt0r_neq0. near=> h; apply: mulr_ge0_le0. - by rewrite invr_ge0; apply: ltrW; near: h; exists 1. + rewrite invr_ge0; apply: ltrW; near: h. + by rewrite /at_right -filter_from_norm_locally; exists 1. rewrite subr_le0 [_%:A]mulr1; apply: cmax; near: h. + rewrite /at_right -filter_from_norm_locally. exists (b - c); first by rewrite subr_gt0 (itvP cab). - move=> h; rewrite /AbsRing_ball /= absrB subr0 absRE. + move=> h; rewrite /ball normmB subr0. move=> /(ler_lt_trans (ler_norm _)); rewrite ltr_subr_addr inE => ->. by move=> /ltr_spsaddl -> //; rewrite (itvP cab). rewrite ['D_1 f c]cvg_at_leftE; last exact: fdrvbl. @@ -1427,15 +1447,17 @@ apply: le0r_flim_map; last first. move=> A [e egt0 Ae]; exists e => // x xe xgt0; apply: Ae => //. exact/ltr0_neq0. near=> h; apply: mulr_le0. - by rewrite invr_le0; apply: ltrW; near: h; exists 1. + rewrite invr_le0; apply: ltrW; near: h. + by rewrite /at_left -filter_from_norm_locally; exists 1. rewrite subr_le0 [_%:A]mulr1; apply: cmax; near: h. +rewrite /at_left -filter_from_norm_locally. exists (c - a); first by rewrite subr_gt0 (itvP cab). -move=> h; rewrite /AbsRing_ball /= absrB subr0 absRE. +move=> h; rewrite /ball normmB subr0. move=> /ltr_normlP []; rewrite ltr_subr_addl ltr_subl_addl inE => -> _. by move=> /ltr_snsaddl -> //; rewrite (itvP cab). Grab Existential Variables. all: end_near. Qed. -Lemma derive1_at_min (f : R^o -> R^o) (a b c : R) : +Lemma derive1_at_min (R : realType) (f : R^o -> R^o) (a b c : R) : a <= b -> (forall t, t \in `]a, b[ -> derivable f t 1) -> c \in `]a, b[ -> (forall t, t \in `]a, b[ -> f c <= f t) -> is_derive (c : R^o) 1 f 0. Proof. @@ -1447,7 +1469,7 @@ apply: derive1_at_max leab _ (cab) _ => t tab; first exact/derivableN/fdrvbl. by rewrite ler_opp2; apply: cmin. Qed. -Lemma Rolle (f : R^o -> R^o) (a b : R) : +Lemma Rolle (R : realType) (f : R^o -> R^o) (a b : R) : a < b -> (forall x, x \in `]a, b[ -> derivable f x 1) -> {in `[a, b], continuous f} -> f a = f b -> exists2 c, c \in `]a, b[ & is_derive (c : R^o) 1 f 0. @@ -1479,7 +1501,7 @@ suff -> : f cmax = f cmin by rewrite fcmin // inE !ltrW ?(itvP sab). by case: cmaxeaVb => ->; case: cmineaVb => ->. Qed. -Lemma MVT (f df : R^o -> R^o) (a b : R) : +Lemma MVT (R : realType) (f df : R^o -> R^o) (a b : R) : a <= b -> (forall x, x \in `]a, b[ -> is_derive (x : R^o) 1 f (df x)) -> {in `[a, b], continuous f} -> exists2 c, c \in `[a, b] & f b - f a = df c * (b - a). @@ -1507,7 +1529,7 @@ move/eqP->; rewrite -mulrA mulVf ?mulr1 //; apply: lt0r_neq0. by rewrite subr_gt0. Qed. -Lemma ler0_derive1_nincr (f : R^o -> R^o) (a b : R) : +Lemma ler0_derive1_nincr (R : realType) (f : R^o -> R^o) (a b : R) : (forall x, x \in `[a, b] -> derivable f x 1) -> (forall x, x \in `[a, b] -> f^`() x <= 0) -> forall x y, a <= x -> x <= y -> y <= b -> f y <= f x. @@ -1519,19 +1541,19 @@ have fdrv z : z \in `]x, y[ -> is_derive (z : R^o) 1 f (f^`()z). apply: DeriveDef; last by rewrite derive1E. apply: fdrvbl; rewrite inE; apply/andP; split; first exact: ler_trans lexz. exact: ler_trans leyb. -have [] := @MVT f (f^`()) x y lexy fdrv. +have [] := @MVT _ f (f^`()) x y lexy fdrv. by move=> ? /itvW /fdrvbl /derivable1_diffP /differentiable_continuous. move=> t /itvW /dfle0 dft dftxy; rewrite -oppr_le0 opprB dftxy. by apply: mulr_le0_ge0 => //; rewrite subr_ge0. Qed. -Lemma le0r_derive1_ndecr (f : R^o -> R^o) (a b : R) : +Lemma le0r_derive1_ndecr (R : realType) (f : R^o -> R^o) (a b : R) : (forall x, x \in `[a, b] -> derivable f x 1) -> (forall x, x \in `[a, b] -> 0 <= f^`() x) -> forall x y, a <= x -> x <= y -> y <= b -> f x <= f y. Proof. move=> fdrvbl dfge0 x y; rewrite -[f _ <= _]ler_opp2. -apply: (@ler0_derive1_nincr (- f)) => t tab; first exact/derivableN/fdrvbl. +apply: (@ler0_derive1_nincr _ (- f)) => t tab; first exact/derivableN/fdrvbl. rewrite derive1E deriveN; last exact: fdrvbl. by rewrite oppr_le0 -derive1E; apply: dfge0. Qed. diff --git a/hierarchy.v b/hierarchy.v index c2ebe1f434..16736c89b5 100644 --- a/hierarchy.v +++ b/hierarchy.v @@ -1,10 +1,8 @@ (* mathcomp analysis (c) 2017 Inria and AIST. License: CeCILL-C. *) -Require Import Reals. From mathcomp Require Import ssreflect ssrfun ssrbool ssrnat eqtype choice. From mathcomp Require Import seq fintype bigop ssralg ssrint ssrnum finmap. From mathcomp Require Import matrix interval zmodp. -Require Import boolp reals Rstruct Rbar. -Require Import classical_sets posnum topology. +Require Import boolp reals Rstruct classical_sets posnum topology. (******************************************************************************) (* This file extends the topological hierarchy with metric-related notions: *) @@ -129,1158 +127,32 @@ Import GRing.Theory Num.Def Num.Theory. Local Open Scope classical_set_scope. -(** * Uniform spaces defined using balls *) - -Definition locally_ {T T'} (ball : T -> R -> set T') (x : T) := - @filter_from R _ [set x | 0 < x] (ball x). - -Lemma locally_E {T T'} (ball : T -> R -> set T') x : - locally_ ball x = @filter_from R _ [set x : R | 0 < x] (ball x). -Proof. by []. Qed. - -Module Uniform. - -Record mixin_of (M : Type) (locally : M -> set (set M)) := Mixin { - ball : M -> R -> M -> Prop ; - ax1 : forall x (e : R), 0 < e -> ball x e x ; - ax2 : forall x y (e : R), ball x e y -> ball y e x ; - ax3 : forall x y z e1 e2, ball x e1 y -> ball y e2 z -> ball x (e1 + e2) z; - ax4 : locally = locally_ ball -}. - -Record class_of (M : Type) := Class { - base : Topological.class_of M; - mixin : mixin_of (Filtered.locally_op base) -}. - -Section ClassDef. - -Structure type := Pack { sort; _ : class_of sort ; _ : Type }. -Local Coercion sort : type >-> Sortclass. -Variables (T : Type) (cT : type). -Definition class := let: Pack _ c _ := cT return class_of cT in c. - -Definition clone c of phant_id class c := @Pack T c T. -Let xT := let: Pack T _ _ := cT in T. -Notation xclass := (class : class_of xT). -Local Coercion base : class_of >-> Topological.class_of. -Local Coercion mixin : class_of >-> mixin_of. - -Definition pack loc (m : @mixin_of T loc) := - fun bT (b : Topological.class_of T) of phant_id (@Topological.class bT) b => - fun m' of phant_id m (m' : @mixin_of T (Filtered.locally_op b)) => - @Pack T (@Class _ b m') T. - -Definition eqType := @Equality.Pack cT xclass xT. -Definition choiceType := @Choice.Pack cT xclass xT. -Definition pointedType := @Pointed.Pack cT xclass xT. -Definition filteredType := @Filtered.Pack cT cT xclass xT. -Definition topologicalType := @Topological.Pack cT xclass xT. - -End ClassDef. - -Module Exports. - -Coercion sort : type >-> Sortclass. -Coercion base : class_of >-> Topological.class_of. -Coercion mixin : class_of >-> mixin_of. -Coercion eqType : type >-> Equality.type. -Canonical eqType. -Coercion choiceType : type >-> Choice.type. -Canonical choiceType. -Coercion pointedType : type >-> Pointed.type. -Canonical pointedType. -Coercion filteredType : type >-> Filtered.type. -Canonical filteredType. -Coercion topologicalType : type >-> Topological.type. -Canonical topologicalType. -Notation uniformType := type. -Notation UniformType T m := (@pack T _ m _ _ idfun _ idfun). -Notation UniformMixin := Mixin. -Notation "[ 'uniformType' 'of' T 'for' cT ]" := (@clone T cT _ idfun) - (at level 0, format "[ 'uniformType' 'of' T 'for' cT ]") : form_scope. -Notation "[ 'uniformType' 'of' T ]" := (@clone T _ _ id) - (at level 0, format "[ 'uniformType' 'of' T ]") : form_scope. - -End Exports. - -End Uniform. - -Export Uniform.Exports. - -Section UniformTopology. - -Lemma my_ball_le (M : Type) (loc : M -> set (set M)) - (m : Uniform.mixin_of loc) : - forall (x : M) (e1 e2 : R), e1 <= e2 -> - forall (y : M), Uniform.ball m x e1 y -> Uniform.ball m x e2 y. -Proof. -move=> x e1 e2 le12 y xe1_y. -move: le12; rewrite ler_eqVlt => /orP [/eqP <- //|]. -rewrite -subr_gt0 => lt12. -rewrite -[e2](subrK e1); apply: Uniform.ax3 xe1_y. -suff : Uniform.ball m x (PosNum lt12)%:num x by []. -exact: Uniform.ax1. -Qed. - -Program Definition topologyOfBallMixin (T : Type) - (loc : T -> set (set T)) (m : Uniform.mixin_of loc) : - Topological.mixin_of loc := topologyOfFilterMixin _ _ _. -Next Obligation. -rewrite (Uniform.ax4 m) locally_E; apply filter_from_proper; last first. - move=> e egt0; exists p; suff : Uniform.ball m p (PosNum egt0)%:num p by []. - exact: Uniform.ax1. -apply: filter_from_filter; first by exists 1%:pos%:num. -move=> e1 e2 e1gt0 e2gt0; exists (Num.min e1 e2). - by have := min_pos_gt0 (PosNum e1gt0) (PosNum e2gt0). -by move=> q pmin_q; split; apply: my_ball_le pmin_q; - rewrite ler_minl lerr // orbC. -Qed. -Next Obligation. -move: H; rewrite (Uniform.ax4 m) locally_E => - [e egt0]; apply. -by have : Uniform.ball m p (PosNum egt0)%:num p by exact: Uniform.ax1. -Qed. -Next Obligation. -move: H; rewrite (Uniform.ax4 m) locally_E => - [e egt0 pe_A]. -exists ((PosNum egt0)%:num / 2) => // q phe_q. -rewrite locally_E; exists ((PosNum egt0)%:num / 2) => // r qhe_r. -by apply: pe_A; rewrite [e]splitr; apply: Uniform.ax3 qhe_r. -Qed. - -End UniformTopology. - -Definition ball {M : uniformType} := Uniform.ball (Uniform.class M). - -Lemma locally_ballE {M : uniformType} : locally_ (@ball M) = locally. -Proof. by case: M=> [?[?[]]]. Qed. - -Lemma filter_from_ballE {M : uniformType} x : - @filter_from R _ [set x : R | 0 < x] (@ball M x) = locally x. -Proof. by rewrite -locally_ballE. Qed. - -Module Export LocallyBall. -Definition locally_simpl := (locally_simpl,@filter_from_ballE,@locally_ballE). -End LocallyBall. - -Lemma locallyP {M : uniformType} (x : M) P : - locally x P <-> locally_ ball x P. -Proof. by rewrite locally_simpl. Qed. - -Section uniformType1. -Context {M : uniformType}. - -Lemma ball_center (x : M) (e : posreal) : ball x e%:num x. -Proof. exact: Uniform.ax1. Qed. -Hint Resolve ball_center. - -Lemma ballxx (x : M) (e : R) : (0 < e)%R -> ball x e x. -Proof. by move=> e_gt0; apply: ball_center (PosNum e_gt0). Qed. - -Lemma ball_sym (x y : M) (e : R) : ball x e y -> ball y e x. -Proof. exact: Uniform.ax2. Qed. - -Lemma ball_triangle (y x z : M) (e1 e2 : R) : - ball x e1 y -> ball y e2 z -> ball x (e1 + e2)%R z. -Proof. exact: Uniform.ax3. Qed. - -Lemma ball_split (z x y : M) (e : R) : - ball x (e / 2)%R z -> ball z (e / 2)%R y -> ball x e y. -Proof. by move=> /ball_triangle h /h; rewrite -splitr. Qed. - -Lemma ball_splitr (z x y : M) (e : R) : - ball z (e / 2)%R x -> ball z (e / 2)%R y -> ball x e y. -Proof. by move=> /ball_sym /ball_split; apply. Qed. - -Lemma ball_splitl (z x y : M) (e : R) : - ball x (e / 2) z -> ball y (e / 2) z -> ball x e y. -Proof. by move=> bxz /ball_sym /(ball_split bxz). Qed. - -Lemma ball_ler (x : M) (e1 e2 : R) : e1 <= e2 -> ball x e1 `<=` ball x e2. -Proof. -move=> le12 y; case: ltrgtP le12 => [//|lte12 _|->//]. -by rewrite -[e2](subrK e1); apply/ball_triangle/ballxx; rewrite subr_gt0. -Qed. - -Lemma ball_le (x : M) (e1 e2 : R) : (e1 <= e2)%coqR -> ball x e1 `<=` ball x e2. -Proof. by move=> /RleP/ball_ler. Qed. - -Definition close (x y : M) : Prop := forall eps : posreal, ball x eps%:num y. - -Lemma close_refl (x : M) : close x x. Proof. by []. Qed. - -Lemma close_sym (x y : M) : close x y -> close y x. -Proof. by move=> ??; apply: ball_sym. Qed. - -Lemma close_trans (x y z : M) : close x y -> close y z -> close x z. -Proof. by move=> cxy cyz eps; apply: ball_split (cxy _) (cyz _). Qed. - -Lemma close_limxx (x y : M) : close x y -> x --> y. -Proof. -move=> cxy P /= /locallyP /= [_/posnumP [eps] epsP]. -apply/locallyP; exists (eps%:num / 2) => // z bxz. -by apply: epsP; apply: ball_splitr (cxy _) bxz. -Qed. - -Definition entourages : set (set (M * M)):= - filter_from [set eps : R | eps > 0] - (fun eps => [set xy | ball xy.1 eps xy.2]). - -Global Instance entourages_filter : ProperFilter entourages. -Proof. -apply filter_from_proper; last by exists (point,point); apply: ballxx. -apply: filter_from_filter; first by exists 1; rewrite ltr01. -move=> _ _ /posnumP[i] /posnumP[j]; exists (minr i j) => // [[/= x y]] bxy. -by eexists => /=; apply: ball_ler bxy; rewrite ler_minl lerr ?orbT. -Qed. -Typeclasses Opaque entourages. - -Definition ball_set (A : set M) e := \bigcup_(p in A) ball p e. -Canonical set_filter_source := - @Filtered.Source Prop _ M (fun A => locally_ ball_set A). - -End uniformType1. - -Section entourages. - -Definition unif_cont (U V : uniformType) (f : U -> V) := - (fun xy => (f xy.1, f xy.2)) @ entourages --> entourages. - -Lemma unif_contP (U V : uniformType) (f : U -> V) : - unif_cont f <-> - forall e, e > 0 -> exists2 d, d > 0 & - forall x, ball x.1 d x.2 -> ball (f x.1) e (f x.2). -Proof. exact: filter_fromP. Qed. - -End entourages. - -Hint Resolve ball_center. -Hint Resolve close_refl. - -(** ** Specific uniform spaces *) - -(** rings with an absolute value *) - -Module AbsRing. - -Record mixin_of (D : ringType) := Mixin { - abs : D -> R; - ax1 : abs 0 = 0 ; - ax2 : abs (- 1) = 1 ; - ax3 : forall x y : D, abs (x + y) <= abs x + abs y ; - ax4 : forall x y : D, abs (x * y) = abs x * abs y ; - ax5 : forall x : D, abs x = 0 -> x = 0 -}. - -Section ClassDef. - -Record class_of (K : Type) := Class { - base : Num.NumDomain.class_of K ; - mixin : mixin_of (Num.NumDomain.Pack base K) -}. -Local Coercion base : class_of >-> Num.NumDomain.class_of. -Local Coercion mixin : class_of >-> mixin_of. - -Structure type := Pack { sort; _ : class_of sort ; _ : Type }. -Local Coercion sort : type >-> Sortclass. - -Variables (T : Type) (cT : type). -Definition class := let: Pack _ c _ := cT return class_of cT in c. -Let xT := let: Pack T _ _ := cT in T. -Notation xclass := (class : class_of xT). -Definition clone c of phant_id class c := @Pack T c T. -Definition pack b0 (m0 : mixin_of (@Num.NumDomain.Pack T b0 T)) := - fun bT b & phant_id (Num.NumDomain.class bT) b => - fun m & phant_id m0 m => Pack (@Class T b m) T. - -Definition eqType := @Equality.Pack cT xclass xT. -Definition choiceType := @Choice.Pack cT xclass xT. -Definition zmodType := @GRing.Zmodule.Pack cT xclass xT. -Definition ringType := @GRing.Ring.Pack cT xclass xT. -Definition comRingType := @GRing.ComRing.Pack cT xclass xT. -Definition unitRingType := @GRing.UnitRing.Pack cT xclass xT. -Definition comUnitRingType := @GRing.ComUnitRing.Pack cT xclass xT. -Definition idomainType := @GRing.IntegralDomain.Pack cT xclass xT. -Definition numDomainType := @Num.NumDomain.Pack cT xclass xT. - -End ClassDef. - -Module Exports. - -Coercion base : class_of >-> Num.NumDomain.class_of. -Coercion mixin : class_of >-> mixin_of. -Coercion sort : type >-> Sortclass. -Coercion eqType : type >-> Equality.type. -Canonical eqType. -Coercion choiceType : type >-> Choice.type. -Canonical choiceType. -Coercion zmodType : type >-> GRing.Zmodule.type. -Canonical zmodType. -Coercion ringType : type >-> GRing.Ring.type. -Canonical ringType. -Coercion comRingType : type >-> GRing.ComRing.type. -Canonical comRingType. -Coercion unitRingType : type >-> GRing.UnitRing.type. -Canonical unitRingType. -Coercion comUnitRingType : type >-> GRing.ComUnitRing.type. -Canonical comUnitRingType. -Coercion idomainType : type >-> GRing.IntegralDomain.type. -Canonical idomainType. -Coercion numDomainType : type >-> Num.NumDomain.type. -Canonical numDomainType. -Notation AbsRingMixin := Mixin. -Notation AbsRingType T m := (@pack T _ m _ _ id _ id). -Notation "[ 'absRingType' 'of' T 'for' cT ]" := (@clone T cT _ idfun) - (at level 0, format "[ 'absRingType' 'of' T 'for' cT ]") : form_scope. -Notation "[ 'absRingType' 'of' T ]" := (@clone T _ _ id) - (at level 0, format "[ 'absRingType' 'of' T ]") : form_scope. -Notation absRingType := type. - -End Exports. - -End AbsRing. - -Export AbsRing.Exports. - -Definition abs {K : absRingType} : K -> R := @AbsRing.abs _ (AbsRing.class K). -Notation "`| x |" := (abs x%R) : R_scope. -Notation "`| x |" := (abs x%R) : real_scope. - -Section AbsRing1. - -Context {K : absRingType}. -Implicit Types x : K. - -Lemma absr0 : `|0 : K| = 0. Proof. exact: AbsRing.ax1. Qed. - -Lemma absrN1: `|- 1 : K| = 1. -Proof. exact: AbsRing.ax2. Qed. - -Lemma ler_abs_add (x y : K) : `|x + y| <= `|x|%real + `|y|%real. -Proof. exact: AbsRing.ax3. Qed. - -Lemma absrM (x y : K) : `|x * y| = `|x|%real * `|y|%real. -Proof. exact: AbsRing.ax4. Qed. - -Lemma absr0_eq0 (x : K) : `|x| = 0 -> x = 0. -Proof. exact: AbsRing.ax5. Qed. - -Lemma absrN x : `|- x| = `|x|. -Proof. -gen have le_absN1 : x / `|- x| <= `|x|. - by rewrite -mulN1r absrM absrN1 mul1r. -by apply/eqP; rewrite eqr_le le_absN1 /= -{1}[x]opprK le_absN1. -Qed. - -Lemma absrB (x y : K) : `|x - y| = `|y - x|. -Proof. by rewrite -absrN opprB. Qed. - -Lemma absr1 : `|1 : K| = 1. Proof. by rewrite -absrN absrN1. Qed. - -Lemma absr_ge0 x : 0 <= `|x|. -Proof. -rewrite -(@pmulr_rge0 _ 2) // mulr2n mulrDl !mul1r. -by rewrite -{2}absrN (ler_trans _ (ler_abs_add _ _)) // subrr absr0. -Qed. - -Lemma absr_eq0 x : (`|x| == 0) = (x == 0). -Proof. by apply/eqP/eqP=> [/absr0_eq0//|->]; rewrite absr0. Qed. - -Lemma absr1_gt0 : `|1 : K| > 0. -Proof. by rewrite ltr_def absr1 oner_eq0 /=. Qed. - -Lemma absrX x n : `|x ^+ n| <= `|x|%real ^+ n. -Proof. -elim: n => [|n IH]; first by rewrite !expr0 absr1. -by rewrite !exprS absrM ler_pmul // absr_ge0. -Qed. - -End AbsRing1. -Hint Resolve absr_ge0. -Hint Resolve absr1_gt0. - -Definition ball_ (V : zmodType) (norm : V -> R) (x : V) - (e : R) := [set y | norm (x - y) < e]. -Arguments ball_ {V} norm x e%R y /. - -Section AbsRing_UniformSpace. - -Context {K : absRingType}. - -Definition AbsRing_ball := ball_ (@abs K). - -Lemma AbsRing_ball_center (x : K) (e : R) : 0 < e -> AbsRing_ball x e x. -Proof. by rewrite /AbsRing_ball /= subrr absr0. Qed. - -Lemma AbsRing_ball_sym (x y : K) (e : R) : - AbsRing_ball x e y -> AbsRing_ball y e x. -Proof. by rewrite /AbsRing_ball /= absrB. Qed. - -(* :TODO: to math-comp *) -Lemma subr_trans (M : zmodType) (z x y : M) : x - y = (x - z) + (z - y). -Proof. by rewrite addrA addrNK. Qed. - -Lemma AbsRing_ball_triangle (x y z : K) (e1 e2 : R) : - AbsRing_ball x e1 y -> AbsRing_ball y e2 z -> AbsRing_ball x (e1 + e2) z. -Proof. -rewrite /AbsRing_ball /= => xy yz. -by rewrite (subr_trans y) (ler_lt_trans (ler_abs_add _ _)) ?ltr_add. -Qed. - -Definition AbsRingUniformMixin := - UniformMixin AbsRing_ball_center AbsRing_ball_sym AbsRing_ball_triangle erefl. - -End AbsRing_UniformSpace. - -(* :TODO: DANGEROUS ! Must change this to include uniform type et al inside absring *) -Coercion absRing_pointedType (K : absRingType) := PointedType K 0. -Canonical absRing_pointedType. -Coercion absRing_filteredType (K : absRingType) := - FilteredType K K (locally_ AbsRing_ball). -Canonical absRing_filteredType. -Coercion absRing_topologicalType (K : absRingType) := - TopologicalType K (topologyOfBallMixin AbsRingUniformMixin). -Canonical absRing_topologicalType. -Coercion absRing_UniformType (K : absRingType) := UniformType K AbsRingUniformMixin. -Canonical absRing_UniformType. - -Lemma ball_absE (K : absRingType) : ball = ball_ (@abs K). -Proof. by []. Qed. - -Lemma locallyN (R : absRingType) (x : R) : - locally (- x) = [set [set - y | y in A] | A in locally x]. -Proof. -rewrite predeqE => A; split=> [[e egt0 oppxe_A]|[B [e egt0 xe_B] <-]]; - last first. - exists e => // y xe_y; exists (- y); last by rewrite opprK. - apply/xe_B. - by rewrite /AbsRing_ball /ball_ opprK -absrN -mulN1r mulrDr !mulN1r. -exists [set - y | y in A]; last first. - rewrite predeqE => y; split=> [[z [t At <- <-]]|Ay]; first by rewrite opprK. - by exists (- y); [exists y|rewrite opprK]. -exists e => // y xe_y; exists (- y); last by rewrite opprK. -by apply/oppxe_A; rewrite /AbsRing_ball /ball_ absrB opprK addrC. -Qed. - -Lemma openN (R : absRingType) (A : set R) : - open A -> open [set - x | x in A]. -Proof. -move=> Aop; rewrite openE => _ [x /Aop x_A <-]. -by rewrite /interior locallyN; exists A. -Qed. - -Lemma closedN (R : absRingType) (A : set R) : - closed A -> closed [set - x | x in A]. -Proof. -move=> Acl x clNAx. -suff /Acl : closure A (- x) by exists (- x)=> //; rewrite opprK. -move=> B oppx_B; have : [set - x | x in A] `&` [set - x | x in B] !=set0. - by apply: clNAx; rewrite -[x]opprK locallyN; exists B. -move=> [y [[z Az oppzey] [t Bt opptey]]]; exists (- y). -by split; [rewrite -oppzey opprK|rewrite -opptey opprK]. -Qed. - -(** real numbers *) - -Program Definition R_AbsRingMixin := - @AbsRing.Mixin _ normr (normr0 _) (normrN1 _) (@ler_norm_add _) _ (@normr0_eq0 _). -Next Obligation. by rewrite normrM. Qed. -Canonical R_absRingType := AbsRingType R R_AbsRingMixin. - -Canonical R_pointedType := [pointedType of R for R_absRingType]. -Canonical R_filteredType := [filteredType R of R for R_absRingType]. -Canonical R_topologicalType := [topologicalType of R for R_absRingType]. -Canonical R_uniformType := [uniformType of R for R_absRingType]. -Canonical Ro_pointedType := [pointedType of R^o for R_absRingType]. -Canonical Ro_filteredType := [filteredType R^o of R^o for R_absRingType]. -Canonical Ro_topologicalType := [topologicalType of R^o for R_absRingType]. -Canonical Ro_uniformType := [uniformType of R^o for R_absRingType]. - -(** locally *) - -Section Locally. -Context {T : uniformType}. - -Lemma locally_ball (x : T) (eps : posreal) : locally x (ball x eps%:num). -Proof. by apply/locallyP; exists eps%:num. Qed. -Hint Resolve locally_ball. - -Lemma forallN {U} (P : set U) : (forall x, ~ P x) = ~ exists x, P x. -Proof. (*boolP*) -rewrite propeqE; split; first by move=> fP [x /fP]. -by move=> nexP x Px; apply: nexP; exists x. -Qed. - -Lemma NNP (P : Prop) : (~ ~ P) <-> P. (*boolP*) -Proof. by split=> [nnp|p /(_ p)//]; have [//|/nnp] := asboolP P. Qed. - -Lemma eqNNP (P : Prop) : (~ ~ P) = P. (*boolP*) -Proof. by rewrite propeqE NNP. Qed. - -Lemma existsN {U} (P : set U) : (exists x, ~ P x) = ~ forall x, P x. (*boolP*) -Proof. -rewrite propeqE; split=> [[x Px] Nall|Nall]; first exact: Px. -by apply/NNP; rewrite -forallN => allP; apply: Nall => x; apply/NNP. -Qed. - -Lemma ex_ball_sig (x : T) (P : set T) : - ~ (forall eps : posreal, ~ (ball x eps%:num `<=` ~` P)) -> - {d : posreal | ball x d%:num `<=` ~` P}. -Proof. -rewrite forallN eqNNP => exNP. -pose D := [set d : R | d > 0 /\ ball x d `<=` ~` P]. -have [|d_gt0] := @getPex _ D; last by exists (PosNum d_gt0). -by move: exNP => [e eP]; exists e%:num. -Qed. - -Lemma locallyC (x : T) (P : set T) : - ~ (forall eps : posreal, ~ (ball x eps%:num `<=` ~` P)) -> - locally x (~` P). -Proof. by move=> /ex_ball_sig [e] ?; apply/locallyP; exists e%:num. Qed. - -Lemma locallyC_ball (x : T) (P : set T) : - locally x (~` P) -> {d : posreal | ball x d%:num `<=` ~` P}. -Proof. -move=> /locallyP xNP; apply: ex_ball_sig. -by have [_ /posnumP[e] eP /(_ _ eP)] := xNP. -Qed. - -Lemma locally_ex (x : T) (P : T -> Prop) : locally x P -> - {d : posreal | forall y, ball x d%:num y -> P y}. -Proof. -move=> /locallyP xP. -pose D := [set d : R | d > 0 /\ forall y, ball x d y -> P y]. -have [|d_gt0 dP] := @getPex _ D; last by exists (PosNum d_gt0). -by move: xP => [e bP]; exists (e : R). -Qed. - -Lemma flim_close {F} {FF : ProperFilter F} (x y : T) : - F --> x -> F --> y -> close x y. -Proof. -move=> Fx Fy e; near F => z; apply: (@ball_splitl _ z); near: z; -by [apply/Fx/locally_ball|apply/Fy/locally_ball]. -Grab Existential Variables. all: end_near. Qed. - -Lemma flimx_close (x y : T) : x --> y -> close x y. -Proof. exact: flim_close. Qed. - -End Locally. -Hint Resolve locally_ball. - -Lemma ler_addgt0Pr (R : realFieldType) (x y : R) : - reflect (forall e, e > 0 -> x <= y + e) (x <= y). -Proof. -apply/(iffP idP)=> [lexy _/posnumP[e] | lexye]; first by rewrite ler_paddr. -case: (lerP x y) => // ltyx. -have /midf_lt [_] := ltyx; rewrite ltrNge -eqbF_neg => /eqP<-. -suff -> : (y + x) / 2 = y + (x - y) / 2. - by apply/lexye/divr_gt0 => //; rewrite subr_gt0. -by rewrite !mulrDl addrC -mulN1r -mulrA mulN1r [RHS]addrC {3}(splitr y) - [RHS]GRing.subrKA. -Qed. - -Lemma ler_addgt0Pl (R : realFieldType) (x y : R) : - reflect (forall e, e > 0 -> x <= e + y) (x <= y). -Proof. -by apply/(equivP (ler_addgt0Pr x y)); split=> lexy e /lexy; rewrite addrC. -Qed. - -Lemma in_segment_addgt0Pr (R : realFieldType) (x y z : R) : - reflect (forall e, e > 0 -> y \in `[(x - e), (z + e)]) (y \in `[x, z]). -Proof. -apply/(iffP idP)=> [xyz _/posnumP[e] | xyz_e]. - rewrite inE; apply/andP; split; last by rewrite ler_paddr // (itvP xyz). - by rewrite ler_subl_addr ler_paddr // (itvP xyz). -rewrite inE; apply/andP. -by split; apply/ler_addgt0Pr => ? /xyz_e /andP []; rewrite ler_subl_addr. -Qed. - -Lemma in_segment_addgt0Pl (R : realFieldType) (x y z : R) : - reflect (forall e, e > 0 -> y \in `[(- e + x), (e + z)]) (y \in `[x, z]). -Proof. -apply/(equivP (in_segment_addgt0Pr x y z)). -by split=> zxy e /zxy; rewrite [z + _]addrC [_ + x]addrC. -Qed. - -Lemma absRE (x : R) : `|x|%real = `|x|%R. -Proof. by []. Qed. - -Lemma Rhausdorff : hausdorff [topologicalType of R]. -Proof. -move=> x y clxy. -apply/eqP; rewrite eqr_le; apply/in_segment_addgt0Pr => _ /posnumP[e]. -rewrite inE -ler_distl -absRE; set he := (e%:num / 2)%:pos. -have [z []] := clxy _ _ (locally_ball x he) (locally_ball y he). -rewrite ball_absE /ball_ absrB => zx_he yz_he. -rewrite (subr_trans z) (ler_trans (ler_abs_add _ _) _)// ltrW//. -by rewrite (splitr e%:num); apply: ltr_add. -Qed. - -(** matrices *) - -Section matrix_Uniform. - -Variables (m n : nat) (T : uniformType). - -Implicit Types x y : 'M[T]_(m, n). - -Definition mx_ball x (e : R) y := forall i j, ball (x i j) e (y i j). - -Lemma mx_ball_center x (e : R) : 0 < e -> mx_ball x e x. -Proof. by move=> ???; apply: ballxx. Qed. - -Lemma mx_ball_sym x y (e : R) : mx_ball x e y -> mx_ball y e x. -Proof. by move=> xe_y ??; apply/ball_sym/xe_y. Qed. - -Lemma mx_ball_triangle x y z (e1 e2 : R) : - mx_ball x e1 y -> mx_ball y e2 z -> mx_ball x (e1 + e2) z. -Proof. -by move=> xe1_y ye2_z ??; apply: ball_triangle; [apply: xe1_y| apply: ye2_z]. -Qed. - -Lemma ltr_bigminr (I : finType) (R : realDomainType) (f : I -> R) (x0 x : R) : - x < x0 -> (forall i, x < f i) -> x < \big[minr/x0]_i f i. -Proof. -move=> ltx0 ltxf; elim/big_ind: _ => // y z ltxy ltxz. -by rewrite ltr_minr ltxy ltxz. -Qed. - -Lemma bigminr_ler (I : finType) (R : realDomainType) (f : I -> R) (x0 : R) i : - \big[minr/x0]_j f j <= f i. -Proof. -have := mem_index_enum i; rewrite unlock; elim: (index_enum I) => //= j l ihl. -by rewrite inE => /orP [/eqP->|/ihl leminlfi]; - rewrite ler_minl ?lerr // leminlfi orbC. -Qed. - -Lemma exists2P A (P Q : A -> Prop) : - (exists2 a, P a & Q a) <-> exists a, P a /\ Q a. -Proof. by split=> [[a ??] | [a []]]; exists a. Qed. - -Lemma mx_locally : locally = locally_ mx_ball. -Proof. -rewrite predeq2E => x A; split; last first. - by move=> [e egt0 xe_A]; exists (fun i j => ball (x i j) (PosNum egt0)%:num). -move=> [P]; rewrite -locally_ballE => x_P sPA. -exists (\big[minr/1]_i \big[minr/1]_j - get (fun e : R => 0 < e /\ ball (x i j) e `<=` P i j)). - apply: ltr_bigminr => // i; apply: ltr_bigminr => // j. - by have /exists2P/getPex [] := x_P i j. -move=> y xmin_y; apply: sPA => i j; have /exists2P/getPex [_] := x_P i j; apply. -apply: ball_ler (xmin_y i j). -by apply: ler_trans (bigminr_ler _ _ i) _; apply: bigminr_ler. -Qed. - -Definition matrix_uniformType_mixin := - Uniform.Mixin mx_ball_center mx_ball_sym mx_ball_triangle mx_locally. - -Canonical matrix_uniformType := - UniformType 'M[T]_(m, n) matrix_uniformType_mixin. - -End matrix_Uniform. - -Lemma coord_continuous {K : absRingType} m n i j : - continuous (fun M : 'M[K]_(m.+1, n.+1) => M i j). -Proof. -move=> /= M s /= /locallyP; rewrite locally_E => -[e e0 es]. -apply/locallyP; rewrite locally_E; exists e => //= N MN; exact/es/MN. -Qed. - -(** product of two uniform spaces *) - -Section prod_Uniform. - -Context {U V : uniformType}. -Implicit Types (x y : U * V). - -Definition prod_point : U * V := (point, point). - -Definition prod_ball x (eps : R) y := - ball (fst x) eps (fst y) /\ ball (snd x) eps (snd y). - -Lemma prod_ball_center x (eps : R) : 0 < eps -> prod_ball x eps x. -Proof. by move=> /posnumP[e]; split. Qed. - -Lemma prod_ball_sym x y (eps : R) : prod_ball x eps y -> prod_ball y eps x. -Proof. by move=> [bxy1 bxy2]; split; apply: ball_sym. Qed. - -Lemma prod_ball_triangle x y z (e1 e2 : R) : - prod_ball x e1 y -> prod_ball y e2 z -> prod_ball x (e1 + e2) z. -Proof. -by move=> [bxy1 bxy2] [byz1 byz2]; split; eapply ball_triangle; eassumption. -Qed. - -Lemma prod_locally : locally = locally_ prod_ball. -Proof. -rewrite predeq2E => -[x y] P; split=> [[[A B] /=[xX yY] XYP] |]; last first. - by move=> [_ /posnumP[eps] epsP]; exists (ball x eps%:num, ball y eps%:num) => /=. -move: xX yY => /locallyP [_ /posnumP[ex] eX] /locallyP [_ /posnumP[ey] eY]. -exists (minr ex%:num ey%:num) => // -[x' y'] [/= xx' yy']. -apply: XYP; split=> /=. - by apply/eX/(ball_ler _ xx'); rewrite ler_minl lerr. -by apply/eY/(ball_ler _ yy'); rewrite ler_minl lerr orbT. -Qed. - -Definition prod_uniformType_mixin := - Uniform.Mixin prod_ball_center prod_ball_sym prod_ball_triangle prod_locally. - -End prod_Uniform. - -Canonical prod_uniformType (U V : uniformType) := - UniformType (U * V) (@prod_uniformType_mixin U V). - -(** Functional metric spaces *) - -Section fct_Uniform. - -Variable (T : choiceType) (U : uniformType). - -Definition fct_ball (x : T -> U) (eps : R) (y : T -> U) := - forall t : T, ball (x t) eps (y t). - -Lemma fct_ball_center (x : T -> U) (e : R) : 0 < e -> fct_ball x e x. -Proof. by move=> /posnumP[{e}e] t. Qed. - -Lemma fct_ball_sym (x y : T -> U) (e : R) : fct_ball x e y -> fct_ball y e x. -Proof. by move=> P t; apply: ball_sym. Qed. - -Lemma fct_ball_triangle (x y z : T -> U) (e1 e2 : R) : - fct_ball x e1 y -> fct_ball y e2 z -> fct_ball x (e1 + e2) z. -Proof. by move=> xy yz t; apply: (@ball_triangle _ (y t)). Qed. - -Definition fct_uniformType_mixin := - UniformMixin fct_ball_center fct_ball_sym fct_ball_triangle erefl. - -Definition fct_topologicalTypeMixin := - topologyOfBallMixin fct_uniformType_mixin. - -Canonical generic_source_filter := @Filtered.Source _ _ _ (locally_ fct_ball). -Canonical fct_topologicalType := - TopologicalType (T -> U) fct_topologicalTypeMixin. -Canonical fct_uniformType := UniformType (T -> U) fct_uniformType_mixin. - -End fct_Uniform. - -(** ** Local predicates *) -(** locally_dist *) -(** Where is it used ??*) - -Definition locally_dist {T : Type} := - locally_ (fun (d : T -> R) delta => [set y | (d y < (delta : R))%R]). - -Global Instance locally_dist_filter T (d : T -> R) : Filter (locally_dist d). -Proof. -apply: filter_from_filter; first by exists 1. -move=> _ _ /posnumP[e1] /posnumP[e2]; exists (minr e1%:num e2%:num) => //. -by move=> P /=; rewrite ltr_minr => /andP [dPe1 dPe2]. -Qed. - -Section Locally_fct. - -Context {T : Type} {U : uniformType}. - -Lemma near_ball (y : U) (eps : posreal) : - \forall y' \near y, ball y eps%:num y'. -Proof. exact: locally_ball. Qed. - -Lemma flim_ballP {F} {FF : Filter F} (y : U) : - F --> y <-> forall eps : R, 0 < eps -> \forall y' \near F, ball y eps y'. -Proof. by rewrite -filter_fromP !locally_simpl /=. Qed. -Definition flim_locally := @flim_ballP. - -Lemma flim_ballPpos {F} {FF : Filter F} (y : U) : - F --> y <-> forall eps : {posnum R}, \forall y' \near F, ball y eps%:num y'. -Proof. -by split => [/flim_ballP|] pos; [case|apply/flim_ballP=> _/posnumP[eps] //]. -Qed. - -Lemma flim_ball {F} {FF : Filter F} (y : U) : - F --> y -> forall eps : R, 0 < eps -> \forall y' \near F, ball y eps y'. -Proof. by move/flim_ballP. Qed. - -Lemma app_flim_locally {F} {FF : Filter F} (f : T -> U) y : - f @ F --> y <-> forall eps : R, 0 < eps -> \forall x \near F, ball y eps (f x). -Proof. exact: flim_ballP. Qed. - -Lemma flimi_ballP {F} {FF : Filter F} (f : T -> U -> Prop) y : - f `@ F --> y <-> - forall eps : R, 0 < eps -> \forall x \near F, exists z, f x z /\ ball y eps z. -Proof. -split=> [Fy _/posnumP[eps] |Fy P] /=; first exact/Fy/locally_ball. -move=> /locallyP[_ /posnumP[eps] subP]. -rewrite near_simpl near_mapi; near=> x. -have [//|z [fxz yz]] := near (Fy _ (posnum_gt0 eps)) x. -by exists z => //; split => //; apply: subP. -Unshelve. all: end_near. Qed. -Definition flimi_locally := @flimi_ballP. - -Lemma flimi_ball {F} {FF : Filter F} (f : T -> U -> Prop) y : - f `@ F --> y -> - forall eps : R, 0 < eps -> F [set x | exists z, f x z /\ ball y eps z]. -Proof. by move/flimi_ballP. Qed. - -Lemma flimi_close {F} {FF: ProperFilter F} (f : T -> set U) (l l' : U) : - {near F, is_fun f} -> f `@ F --> l -> f `@ F --> l' -> close l l'. -Proof. -move=> f_prop fFl fFl'. -suff f_totalfun: infer {near F, is_totalfun f} by exact: flim_close fFl fFl'. -apply: filter_app f_prop; near=> x; split=> //=. -by have [//|y [fxy _]] := near (flimi_ball fFl [gt0 of 1]) x; exists y. -Grab Existential Variables. all: end_near. Qed. -Definition flimi_locally_close := @flimi_close. - -End Locally_fct. - -Section Locally_fct2. - -Context {T : Type} {U V : uniformType}. - -Lemma flim_ball2P {F : set (set U)} {G : set (set V)} - {FF : Filter F} {FG : Filter G} (y : U) (z : V): - (F, G) --> (y, z) <-> - forall eps : R, eps > 0 -> \forall y' \near F & z' \near G, - ball y eps y' /\ ball z eps z'. -Proof. exact: flim_ballP. Qed. - -End Locally_fct2. - -Lemma flim_const {T} {U : uniformType} {F : set (set T)} - {FF : Filter F} (a : U) : a @[_ --> F] --> a. -Proof. -move=> P /locallyP[_ /posnumP[eps] subP]; rewrite !near_simpl /=. -by apply: filterE=> ?; apply/subP. -Qed. -Arguments flim_const {T U F FF} a. - -Section Cvg. - -Context {U : uniformType}. - -Lemma close_lim (F1 F2 : set (set U)) {FF2 : ProperFilter F2} : - F1 --> F2 -> F2 --> F1 -> close (lim F1) (lim F2). -Proof. -move=> F12 F21; have [/(flim_trans F21) F2l|dvgF1] := pselect (cvg F1). - by apply: (@flim_close _ F2) => //; apply: cvgP F2l. -have [/(flim_trans F12)/cvgP//|dvgF2] := pselect (cvg F2). -by rewrite dvgP // dvgP //. -Qed. - -Lemma flim_closeP (F : set (set U)) (l : U) : ProperFilter F -> - F --> l <-> ([cvg F in U] /\ close (lim F) l). -Proof. -move=> FF; split=> [Fl|[cvF]Cl]. - by have /cvgP := Fl; split=> //; apply: flim_close. -by apply: flim_trans (close_limxx Cl). -Qed. - -End Cvg. -Arguments close_lim {U} F1 F2 {FF2} _. - -(** ** Complete uniform spaces *) - -(* :TODO: Use cauchy2 alternative to define cauchy? *) -(* Or not: is the fact that cauchy F -/> ProperFilter F a problem? *) -Definition cauchy_ex {T : uniformType} (F : set (set T)) := - forall eps : R, 0 < eps -> exists x, F (ball x eps). - -Definition cauchy {T : uniformType} (F : set (set T)) := - forall e, e > 0 -> \forall x & y \near F, ball x e y. - -Lemma cauchy_entouragesP (T : uniformType) (F : set (set T)) : - Filter F -> cauchy F <-> (F, F) --> entourages. -Proof. -move=> FF; split=> cauchyF; last first. - by move=> _/posnumP[eps]; apply: cauchyF; exists eps%:num. -move=> U [_/posnumP[eps] xyepsU]. -by near=> x; apply: xyepsU; near: x; apply: cauchyF. -Grab Existential Variables. all: end_near. Qed. - -Lemma cvg_cauchy_ex {T : uniformType} (F : set (set T)) : - [cvg F in T] -> cauchy_ex F. -Proof. by move=> Fl _/posnumP[eps]; exists (lim F); apply/Fl/locally_ball. Qed. - -Lemma cauchy_exP (T : uniformType) (F : set (set T)) : Filter F -> - cauchy_ex F -> cauchy F. -Proof. -move=> FF Fc; apply/cauchy_entouragesP => A [_/posnumP[e] sdeA]. -have /Fc [z /= Fze] := [gt0 of e%:num / 2]; near=> x y; apply: sdeA => /=. -by apply: (@ball_splitr _ z); [near: x|near: y]. -Grab Existential Variables. all: end_near. Qed. - -Lemma cauchyP (T : uniformType) (F : set (set T)) : ProperFilter F -> - cauchy F <-> cauchy_ex F. -Proof. -move=> FF; split=> [Fcauchy _/posnumP[e] |/cauchy_exP//]. -by near F => x; exists x; near: x; apply: (@nearP_dep _ _ F F); apply: Fcauchy. -Grab Existential Variables. all: end_near. Qed. - -Lemma cvg_cauchy {T : uniformType} (F : set (set T)) : Filter F -> - [cvg F in T] -> cauchy F. -Proof. by move=> FF /cvg_cauchy_ex /cauchy_exP. Qed. - -Module Complete. - -Definition axiom (T : uniformType) := - forall (F : set (set T)), ProperFilter F -> cauchy F -> F --> lim F. - -Section ClassDef. - -Record class_of (T : Type) := Class { - base : Uniform.class_of T ; - mixin : axiom (Uniform.Pack base T) -}. -Local Coercion base : class_of >-> Uniform.class_of. -Local Coercion mixin : class_of >-> Complete.axiom. - -Structure type := Pack { sort; _ : class_of sort ; _ : Type }. -Local Coercion sort : type >-> Sortclass. - -Variables (T : Type) (cT : type). - -Definition class := let: Pack _ c _ := cT return class_of cT in c. - -Definition clone c of phant_id class c := @Pack T c T. -Let xT := let: Pack T _ _ := cT in T. -Notation xclass := (class : class_of xT). - -Definition pack b0 (m0 : axiom (@Uniform.Pack T b0 T)) := - fun bT b of phant_id (Uniform.class bT) b => - fun m of phant_id m m0 => @Pack T (@Class T b m) T. - -Definition eqType := @Equality.Pack cT xclass xT. -Definition choiceType := @Choice.Pack cT xclass xT. -Definition pointedType := @Pointed.Pack cT xclass xT. -Definition filteredType := @Filtered.Pack cT cT xclass xT. -Definition topologicalType := @Topological.Pack cT xclass xT. -Definition uniformType := @Uniform.Pack cT xclass xT. - -End ClassDef. - -Module Exports. - -Coercion base : class_of >-> Uniform.class_of. -Coercion mixin : class_of >-> axiom. -Coercion sort : type >-> Sortclass. -Coercion eqType : type >-> Equality.type. -Canonical eqType. -Coercion choiceType : type >-> Choice.type. -Canonical choiceType. -Coercion pointedType : type >-> Pointed.type. -Canonical pointedType. -Coercion filteredType : type >-> Filtered.type. -Canonical filteredType. -Coercion topologicalType : type >-> Topological.type. -Canonical topologicalType. -Coercion uniformType : type >-> Uniform.type. -Canonical uniformType. -Notation completeType := type. -Notation "[ 'completeType' 'of' T 'for' cT ]" := (@clone T cT _ idfun) - (at level 0, format "[ 'completeType' 'of' T 'for' cT ]") : form_scope. -Notation "[ 'completeType' 'of' T ]" := (@clone T _ _ id) - (at level 0, format "[ 'completeType' 'of' T ]") : form_scope. -Notation CompleteType T m := (@pack T _ m _ _ idfun _ idfun). - -End Exports. - -End Complete. - -Export Complete.Exports. - -Section completeType1. - -Context {T : completeType}. - -Lemma complete_cauchy (F : set (set T)) (FF : ProperFilter F) : - cauchy F -> F --> lim F. -Proof. by case: T F FF => [? [?]]. Qed. - -End completeType1. -Arguments complete_cauchy {T} F {FF} _. - -Section matrix_Complete. - -Variables (T : completeType) (m n : nat). - -Lemma mx_complete (F : set (set 'M[T]_(m, n))) : - ProperFilter F -> cauchy F -> cvg F. -Proof. -move=> FF Fc. -have /(_ _ _) /complete_cauchy cvF : - cauchy ((fun M : 'M[T]_(m, n) => M _ _) @ F). - by move=> ?? _ /posnumP[e]; rewrite near_simpl; apply: filterS (Fc _ _). -apply/cvg_ex. -exists (\matrix_(i, j) (lim ((fun M : 'M[T]_(m, n) => M i j) @ F) : T)). -apply/flim_ballP => _ /posnumP[e]; near=> M => i j. -rewrite mxE; near F => M' => /=; apply: (@ball_splitl _ (M' i j)). - by near: M'; apply/cvF/locally_ball. -by move: (i) (j); near: M'; near: M; apply: nearP_dep; apply: filterS (Fc _ _). -Grab Existential Variables. all: end_near. Qed. - -Canonical matrix_completeType := CompleteType 'M[T]_(m, n) mx_complete. - -End matrix_Complete. - -Section fun_Complete. - -Context {T : choiceType} {U : completeType}. - -Lemma fun_complete (F : set (set (T -> U))) - {FF : ProperFilter F} : cauchy F -> cvg F. -Proof. -move=> Fc; have /(_ _) /complete_cauchy Ft_cvg : cauchy (@^~_ @ F). - by move=> t e ?; rewrite near_simpl; apply: filterS (Fc _ _). -apply/cvg_ex; exists (fun t => lim (@^~t @ F)). -apply/flim_ballPpos => e; near=> f => t; near F => g => /=. -apply: (@ball_splitl _ (g t)); first by near: g; exact/Ft_cvg/locally_ball. -by move: (t); near: g; near: f; apply: nearP_dep; apply: filterS (Fc _ _). -Grab Existential Variables. all: end_near. Qed. - -Canonical fun_completeType := CompleteType (T -> U) fun_complete. - -End fun_Complete. - -(** ** Limit switching *) - -Section Flim_switch. - -Context {T1 T2 : choiceType}. - -Lemma flim_switch_1 {U : uniformType} - F1 {FF1 : ProperFilter F1} F2 {FF2 : Filter F2} - (f : T1 -> T2 -> U) (g : T2 -> U) (h : T1 -> U) (l : U) : - f @ F1 --> g -> (forall x1, f x1 @ F2 --> h x1) -> h @ F1 --> l -> - g @ F2 --> l. -Proof. -move=> fg fh hl; apply/flim_ballPpos => e. -rewrite near_simpl; near F1 => x1; near=> x2. -apply: (@ball_split _ (h x1)); first by near: x1; apply/hl/locally_ball. -apply: (@ball_splitl _ (f x1 x2)); first by near: x2; apply/fh/locally_ball. -by move: (x2); near: x1; apply/(flim_ball fg). -Grab Existential Variables. all: end_near. Qed. - -Lemma flim_switch_2 {U : completeType} - F1 {FF1 : ProperFilter F1} F2 {FF2 : ProperFilter F2} - (f : T1 -> T2 -> U) (g : T2 -> U) (h : T1 -> U) : - f @ F1 --> g -> (forall x, f x @ F2 --> h x) -> - [cvg h @ F1 in U]. -Proof. -move=> fg fh; apply: complete_cauchy => _/posnumP[e]. -rewrite !near_simpl; near=> x1 y1=> /=; near F2 => x2. -apply: (@ball_splitl _ (f x1 x2)); first by near: x2; apply/fh/locally_ball. -apply: (@ball_split _ (f y1 x2)); first by near: x2; apply/fh/locally_ball. -apply: (@ball_splitr _ (g x2)); move: (x2); [near: y1|near: x1]; -by apply/(flim_ball fg). -Grab Existential Variables. all: end_near. Qed. - -(* Alternative version *) -(* Lemma flim_switch {U : completeType} *) -(* F1 (FF1 : ProperFilter F1) F2 (FF2 : ProperFilter F2) (f : T1 -> T2 -> U) (g : T2 -> U) (h : T1 -> U) : *) -(* [cvg f @ F1 in T2 -> U] -> (forall x, [cvg f x @ F2 in U]) -> *) -(* [/\ [cvg [lim f @ F1] @ F2 in U], [cvg (fun x => [lim f x @ F2]) @ F1 in U] *) -(* & [lim [lim f @ F1] @ F2] = [lim (fun x => [lim f x @ F2]) @ F1]]. *) -Lemma flim_switch {U : completeType} - F1 (FF1 : ProperFilter F1) F2 (FF2 : ProperFilter F2) - (f : T1 -> T2 -> U) (g : T2 -> U) (h : T1 -> U) : - f @ F1 --> g -> (forall x1, f x1 @ F2 --> h x1) -> - exists l : U, h @ F1 --> l /\ g @ F2 --> l. -Proof. -move=> Hfg Hfh; have hcv := !! flim_switch_2 Hfg Hfh. -by exists [lim h @ F1 in U]; split=> //; apply: flim_switch_1 Hfg Hfh hcv. -Qed. - -End Flim_switch. - -(** * Some Topology on [Rbar] *) - -(* NB: already defined in R_scope in Rbar.v *) - -Notation "'+oo'" := p_infty : real_scope. -Notation "'-oo'" := m_infty : real_scope. -Definition Rbar_locally' (a : Rbar) (P : R -> Prop) := - match a with - | Finite a => locally' a P - | +oo => exists M : R, forall x, M < x -> P x - | -oo => exists M : R, forall x, x < M -> P x - end. -Definition Rbar_locally (a : Rbar) (P : R -> Prop) := - match a with - | Finite a => locally a P - | +oo => exists M : R, forall x, M < x -> P x - | -oo => exists M : R, forall x, x < M -> P x - end. - -Canonical Rbar_eqType := EqType Rbar gen_eqMixin. -Canonical Rbar_choiceType := ChoiceType Rbar gen_choiceMixin. -Canonical Rbar_pointed := PointedType Rbar (+oo). -Canonical Rbar_filter := FilteredType R Rbar (Rbar_locally). - -Global Instance Rlocally'_proper (x : R) : ProperFilter (locally' x). -Proof. -apply: Build_ProperFilter => A [_/posnumP[e] Ae]. -exists (x + e%:num / 2); apply: Ae; last first. - by rewrite eq_sym addrC -subr_eq subrr eq_sym. -rewrite /AbsRing_ball /= opprD addrA subrr absrB subr0 absRE ger0_norm //. -by rewrite {2}(splitr e%:num) ltr_spaddl. -Qed. - -Global Instance Rbar_locally'_filter : forall x, ProperFilter (Rbar_locally' x). -Proof. -case=> [x||]; first exact: Rlocally'_proper. - apply Build_ProperFilter. - by move=> P [M gtMP]; exists (M + 1); apply: gtMP; rewrite ltr_addl. - split=> /= [|P Q [MP gtMP] [MQ gtMQ] |P Q sPQ [M gtMP]]; first by exists 0. - by exists (maxr MP MQ) => ?; rewrite ltr_maxl => /andP [/gtMP ? /gtMQ]. - by exists M => ? /gtMP /sPQ. -apply Build_ProperFilter. - by move=> P [M ltMP]; exists (M - 1); apply: ltMP; rewrite gtr_addl oppr_lt0. -split=> /= [|P Q [MP ltMP] [MQ ltMQ] |P Q sPQ [M ltMP]]; first by exists 0. - by exists (minr MP MQ) => ?; rewrite ltr_minr => /andP [/ltMP ? /ltMQ]. -by exists M => ? /ltMP /sPQ. -Qed. -Typeclasses Opaque Rbar_locally'. - - -Global Instance Rbar_locally_filter : forall x, ProperFilter (Rbar_locally x). -Proof. -case=> [x||]. -by apply/locally_filter. -exact: (Rbar_locally'_filter +oo). -exact: (Rbar_locally'_filter -oo). -Qed. -Typeclasses Opaque Rbar_locally. - -Lemma near_pinfty_div2 (A : set R) : - (\forall k \near +oo, A k) -> (\forall k \near +oo, A (k / 2)). -Proof. -by move=> [M AM]; exists (M * 2) => x; rewrite -ltr_pdivl_mulr //; apply: AM. -Qed. - -Lemma locally_pinfty_gt c : \forall x \near +oo, c < x. -Proof. by exists c. Qed. - -Lemma locally_pinfty_ge c : \forall x \near +oo, c <= x. -Proof. by exists c; apply: ltrW. Qed. - -Hint Extern 0 (is_true (0 < _)) => match goal with - H : ?x \is_near (locally +oo) |- _ => - solve[near: x; exists 0 => _/posnumP[x] //] end : core. - (** ** Modules with a norm *) Reserved Notation "`|[ x ]|" (at level 0, x at level 99, format "`|[ x ]|"). +Definition ball (K : numDomainType) (V : zmodType) (norm : V -> K) (x : V) + (e : K) := [set y | norm (x - y) < e]. +Arguments ball {K V} norm x e%R y /. + +Definition entourages_ (K : numDomainType) (V : zmodType) (norm : V -> K) := + filter_from [set e : K | e > 0] + (fun e => [set xy | norm (xy.1 - xy.2) < e]). + Module NormedModule. -Record mixin_of (K : absRingType) (V : lmodType K) loc (m : @Uniform.mixin_of V loc) := Mixin { - norm : V -> R ; +Record mixin_of (K : numDomainType) (V : lmodType K) loc + (m : @Uniform.mixin_of V loc) := Mixin { + norm : V -> K ; ax1 : forall (x y : V), norm (x + y) <= norm x + norm y ; - ax2 : forall (l : K) (x : V), norm (l *: x) = abs l * norm x; - ax3 : Uniform.ball m = ball_ norm; + ax2 : forall (l : K) (x : V), norm (l *: x) = `|l| * norm x; + ax3 : Uniform.entourages m = entourages_ norm; ax4 : forall x : V, norm x = 0 -> x = 0 }. Section ClassDef. -Variable K : absRingType. +Variable K : numDomainType. Record class_of (T : Type) := Class { base : GRing.Lmodule.class_of K T ; @@ -1369,25 +241,25 @@ End NormedModule. Export NormedModule.Exports. -Definition norm {K : absRingType} {V : normedModType K} : V -> R := +Definition norm {K : numDomainType} {V : normedModType K} : V -> K := NormedModule.norm (NormedModule.class _). Notation "`|[ x ]|" := (norm x) : ring_scope. Section NormedModule1. -Context {K : absRingType} {V : normedModType K}. -Implicit Types (l : K) (x y : V) (eps : posreal). +Context {K : numDomainType} {V : normedModType K}. +Implicit Types (l : K) (x y : V) (e : {posnum K}). Lemma ler_normm_add x y : `|[x + y]| <= `|[x]| + `|[y]|. Proof. exact: NormedModule.ax1. Qed. -Lemma normmZ l x : `|[l *: x]| = `|l|%real * `|[x]|. +Lemma normmZ l x : `|[l *: x]| = `|l| * `|[x]|. Proof. exact: NormedModule.ax2. Qed. -Notation ball_norm := (ball_ (@norm K V)). +Notation entourages_norm := (entourages_ (@norm K V)). -Notation locally_norm := (locally_ ball_norm). +Notation locally_norm := (locally_ entourages_norm). -Lemma ball_normE : ball_norm = ball. +Lemma entourages_normE : entourages_norm = entourages. Proof. by rewrite -NormedModule.ax3. Qed. Lemma normm0_eq0 x : `|[x]| = 0 -> x = 0. @@ -1396,7 +268,7 @@ Proof. exact: NormedModule.ax4. Qed. Lemma normmN x : `|[- x]| = `|[x]|. Proof. gen have le_absN1 : x / `|[- x]| <= `|[x]|. - by rewrite -scaleN1r normmZ absrN1 mul1r. + by rewrite -scaleN1r normmZ normrN1 mul1r. by apply/eqP; rewrite eqr_le le_absN1 /= -{1}[x]opprK le_absN1. Qed. @@ -1406,7 +278,7 @@ Proof. by rewrite -normmN opprB. Qed. Lemma normm0 : `|[0 : V]| = 0. Proof. apply/eqP; rewrite eqr_le; apply/andP; split. - by rewrite -{1}(scale0r 0) normmZ absr0 mul0r. + by rewrite -{1}(scale0r 0) normmZ normr0 mul0r. by rewrite -(ler_add2r `|[0 : V]|) add0r -{1}[0 : V]add0r ler_normm_add. Qed. Hint Resolve normm0. @@ -1424,237 +296,482 @@ Lemma normm_gt0 x : (0 < `|[x]|) = (x != 0). Proof. by rewrite ltr_def normm_eq0 normm_ge0 andbT. Qed. Lemma normm_lt0 x : (`|[x]| < 0) = false. -Proof. by rewrite ltrNge normm_ge0. Qed. +Proof. by apply: ler_gtF; rewrite normm_ge0. Qed. Lemma normm_le0 x : (`|[x]| <= 0) = (x == 0). -Proof. by rewrite lerNgt normm_gt0 negbK. Qed. +Proof. by rewrite ler_eqVlt normm_lt0 orbF normm_eq0. Qed. Lemma ler_distm_dist x y : `| `|[x]| - `|[y]| | <= `|[x - y]|. Proof. wlog gt_xy : x y / `|[x]| >= `|[y]| => [hw|]. - by have [/hw//|/ltrW/hw] := lerP `|[y]| `|[x]|; rewrite absRE distrC normmB. -rewrite absRE ger0_norm ?subr_ge0 // ler_subl_addr. + have /orP [/hw//|/hw] := ger_leVge (normm_ge0 y) (normm_ge0 x). + by rewrite distrC normmB. +rewrite ger0_norm ?subr_ge0 // ler_subl_addr. by rewrite -{1}[x](addrNK y) ler_normm_add. Qed. -Lemma distm_lt_split z x y (e : R) : - `|[x - z]| < (e / 2)%R -> `|[z - y]| < (e / 2)%R -> `|[x - y]| < e. -Proof. by have := @ball_split _ z x y e; rewrite -ball_normE. Qed. - -Lemma distm_lt_splitr z x y (e : R) : - `|[z - x]| < (e / 2)%R -> `|[z - y]| < (e / 2)%R -> `|[x - y]| < e. -Proof. by have := @ball_splitr _ z x y e; rewrite -ball_normE. Qed. - -Lemma distm_lt_splitl z x y (e : R) : - `|[x - z]| < (e / 2)%R -> `|[y - z]| < (e / 2)%R -> `|[x - y]| < e. -Proof. by have := @ball_splitl _ z x y e; rewrite -ball_normE. Qed. - -Lemma normm_leW x (e : R) : e > 0 -> `|[x]| <= (e / 2)%R -> `|[x]| < e. -Proof. -move=> /posnumP[{e}e] /ler_lt_trans ->//. -by rewrite [X in _ < X]splitr ltr_spaddl. -Qed. - -Lemma normm_lt_split x y (e : R) : - `|[x]| < (e / 2)%R -> `|[y]| < (e / 2)%R -> `|[x + y]| < e. -Proof. -by move=> xlt ylt; rewrite -[y]opprK (@distm_lt_split 0) ?subr0 ?opprK ?add0r. -Qed. - -Lemma closeE x y : close x y = (x = y). -Proof. -rewrite propeqE; split => [cl_xy|->//]; have [//|neq_xy] := eqVneq x y. -have dxy_gt0 : `|[x - y]| > 0 by rewrite normm_gt0 subr_eq0. -have dxy_ge0 := ltrW dxy_gt0. -have := cl_xy ((PosNum dxy_gt0)%:num / 2)%:pos. -rewrite -ball_normE /= -subr_lt0 ler_gtF //. -rewrite -[X in X - _]mulr1 -mulrBr mulr_ge0 //. -by rewrite subr_ge0 -(@ler_pmul2r _ 2) // mulVf // mul1r ler1n. -Qed. -Lemma eq_close x y : close x y -> x = y. by rewrite closeE. Qed. +Lemma entourages_ball (e : {posnum K}) : + entourages [set xy : V * V | `|[xy.1 - xy.2]| < e%:num]. +Proof. by rewrite -entourages_normE; apply: in_filter_from. Qed. Lemma locally_le_locally_norm x : flim (locally x) (locally_norm x). Proof. -move=> P [_ /posnumP[e] subP]; apply/locallyP. -by eexists; last (move=> y Py; apply/subP; rewrite ball_normE; apply/Py). +move=> P [A entA sAB]; apply/locallyP; exists A => //. +by rewrite -entourages_normE. Qed. Lemma locally_norm_le_locally x : flim (locally_norm x) (locally x). Proof. -move=> P /locallyP [_ /posnumP[e] Pxe]. -by exists e%:num => // y; rewrite ball_normE; apply/Pxe. +by move=> P /locallyP [A entA sAP]; exists A => //; rewrite entourages_normE. Qed. (* NB: this lemmas was not here before *) Lemma locally_locally_norm : locally_norm = locally. Proof. -by rewrite funeqE => x; rewrite /locally_norm ball_normE filter_from_ballE. +by rewrite funeqE => x; rewrite /locally_norm entourages_normE + filter_from_entourageE. Qed. Lemma locally_normP x P : locally x P <-> locally_norm x P. Proof. by rewrite locally_locally_norm. Qed. Lemma filter_from_norm_locally x : - @filter_from R _ [set x : R | 0 < x] (ball_norm x) = locally x. -Proof. by rewrite -locally_locally_norm. Qed. + @filter_from K _ [set x : K | 0 < x] (ball norm x) = locally x. +Proof. +rewrite predeqE => A; split=> [[_/posnumP[e] sxeA] |]. + by rewrite -locally_entourageE; exists [set xy | `|[xy.1 - xy.2]| < e%:num]; + [apply: entourages_ball|move=> ? /sxeA]. +rewrite -locally_locally_norm => - [B [_/posnumP[e] seB] sBA]. +by exists e%:num => // y xye; apply/sBA/seB. +Qed. Lemma locally_normE (x : V) (P : set V) : locally_norm x P = \near x, P x. Proof. by rewrite locally_locally_norm near_simpl. Qed. Lemma filter_from_normE (x : V) (P : set V) : - @filter_from R _ [set x : R | 0 < x] (ball_norm x) P = \near x, P x. + @filter_from K _ [set x : K | 0 < x] (ball norm x) P = \near x, P x. Proof. by rewrite filter_from_norm_locally. Qed. Lemma near_locally_norm (x : V) (P : set V) : (\forall x \near locally_norm x, P x) = \near x, P x. Proof. exact: locally_normE. Qed. -Lemma locally_norm_ball_norm x (e : posreal) : - locally_norm x (ball_norm x e%:num). -Proof. by exists e%:num. Qed. +Lemma locally_norm_ball x (e : {posnum K}) : + locally_norm x (ball norm x e%:num). +Proof. +by rewrite locally_locally_norm -filter_from_norm_locally; exists e%:num. +Qed. -Lemma locally_norm_ball x (eps : posreal) : locally_norm x (ball x eps%:num). -Proof. rewrite locally_locally_norm; by apply: locally_ball. Qed. +Lemma locally_ball (x : V) (e : {posnum K}) : locally x (ball norm x e%:num). +Proof. rewrite -locally_locally_norm; apply: locally_norm_ball. Qed. -Lemma locally_ball_norm (x : V) (eps : posreal) : locally x (ball_norm x eps%:num). -Proof. rewrite -locally_locally_norm; apply: locally_norm_ball_norm. Qed. +(* :TODO: to math-comp *) +Lemma subr_trans (M : zmodType) (z x y : M) : x - y = (x - z) + (z - y). +Proof. by rewrite addrA addrNK. Qed. -Lemma ball_norm_triangle (x y z : V) (e1 e2 : R) : - ball_norm x e1 y -> ball_norm y e2 z -> ball_norm x (e1 + e2) z. +Lemma ball_triangle (x y z : V) (e1 e2 : K) : + ball norm x e1 y -> ball norm y e2 z -> ball norm x (e1 + e2) z. Proof. -rewrite /ball_norm => H1 H2; rewrite (subr_trans y). +rewrite /ball => H1 H2; rewrite (subr_trans y). by rewrite (ler_lt_trans (ler_normm_add _ _)) ?ltr_add. Qed. -Lemma ball_norm_center (x : V) (e : posreal) : ball_norm x e%:num x. -Proof. by rewrite /ball_norm subrr normm0. Qed. +Lemma ball_center (x : V) (e : {posnum K}) : ball norm x e%:num x. +Proof. by rewrite /ball subrr normm0. Qed. -Lemma ball_norm_dec x y (e : R) : {ball_norm x e y} + {~ ball_norm x e y}. +Lemma ball_dec x y (e : K) : {ball norm x e y} + {~ ball norm x e y}. Proof. exact: pselect. Qed. -Lemma ball_norm_sym x y (e : R) : ball_norm x e y -> ball_norm y e x. -Proof. by rewrite /ball_norm -opprB normmN. Qed. +Lemma ball_sym x y (e : K) : ball norm x e y -> ball norm y e x. +Proof. by rewrite /ball -opprB normmN. Qed. -Lemma ball_norm_le x (e1 e2 : R) : - e1 <= e2 -> ball_norm x e1 `<=` ball_norm x e2. +Lemma ball_norm_le x (e1 e2 : K) : + e1 <= e2 -> ball norm x e1 `<=` ball norm x e2. Proof. by move=> e1e2 y /ltr_le_trans; apply. Qed. -Lemma norm_close x y : close x y = (forall eps : posreal, ball_norm x eps%:num y). -Proof. by rewrite propeqE ball_normE. Qed. +Lemma norm_close x y : + close x y = (forall e : {posnum K}, ball norm x e%:num y). +Proof. +rewrite propeqE; split; first by move=> xy e; have /xy := entourages_ball e. +by move=> xy A; rewrite -entourages_normE => - [_/posnumP[e]]; apply; apply: xy. +Qed. + +End NormedModule1. + +Hint Resolve normm_ge0. + +Module Export LocallyNorm. +Definition locally_simpl := (locally_simpl,@locally_locally_norm,@filter_from_norm_locally). +End LocallyNorm. + +Module Export NearNorm. +Definition near_simpl := (@near_simpl, @locally_normE, + @filter_from_normE, @near_locally_norm). +Ltac near_simpl := rewrite ?near_simpl. +End NearNorm. + +Lemma flim_normP (K : numDomainType) (V : normedModType K) {F : set (set V)} + {FF : Filter F} (y : V) : + F --> y <-> forall e, 0 < e -> \forall y' \near F, `|[y - y']| < e. +Proof. by rewrite -filter_fromP /= !locally_simpl. Qed. + +Lemma flim_norm (K : numDomainType) (V : normedModType K) {F : set (set V)} + {FF : Filter F} (y : V) : + F --> y -> forall e, e > 0 -> \forall y' \near F, `|[y - y']| < e. +Proof. by move=> /flim_normP. Qed. +Arguments flim_norm {_ _ F FF}. + +Lemma continuous_flim_norm {K : numDomainType} (V W : normedModType K) + (f : V -> W) x l : + continuous f -> x --> l -> forall e : {posnum K}, `|[f l - f x]| < e. +Proof. +move=> cf xl e. +move/flim_norm: (cf l) => /(_ _ (posnum_gt0 e)). +rewrite nearE => /locallyP; rewrite -entourages_normE => -[A [i i0 sA]]. +apply; apply: sA. +have /@flim_norm : Filter [filter of x] by apply: filter_on_Filter. +move/(_ _ xl _ i0). +rewrite nearE => /locallyP; rewrite -entourages_normE => -[B [j j0 sB]]. +by apply; apply: sB => /=; rewrite subrr normm0. +Qed. + +Section NormedModuleField. + +Context {K : numFieldType} {V : normedModType K}. +Implicit Types (x y : V). + +Lemma distm_lt_split z x y (e : K) : + `|[x - z]| < e / 2 -> `|[z - y]| < e / 2 -> `|[x - y]| < e. +Proof. +move=> xz zy; rewrite -(subrK z x) -addrA (ler_lt_trans (ler_normm_add _ _)) //. +by rewrite [e]splitr ltr_add. +Qed. + +Lemma distm_lt_splitr z x y (e : K) : + `|[z - x]| < e / 2 -> `|[z - y]| < e / 2 -> `|[x - y]| < e. +Proof. by rewrite normmB; apply: distm_lt_split. Qed. + +Lemma distm_lt_splitl z x y (e : K) : + `|[x - z]| < e / 2 -> `|[y - z]| < e / 2 -> `|[x - y]| < e. +Proof. by rewrite (normmB y); apply: distm_lt_split. Qed. + +Lemma normm_leW x (e : K) : e > 0 -> `|[x]| <= e / 2 -> `|[x]| < e. +Proof. +move=> /posnumP[{e}e] /ler_lt_trans ->//. +by rewrite [X in _ < X]splitr ltr_spaddl. +Qed. + +Lemma normm_lt_split x y (e : K) : + `|[x]| < e / 2 -> `|[y]| < e / 2 -> `|[x + y]| < e. +Proof. +by move=> xlt ylt; rewrite -[y]opprK (@distm_lt_split 0) ?subr0 ?opprK ?add0r. +Qed. + +Lemma closeE x y : close x y = (x = y). +Proof. +rewrite propeqE; split => [cl_xy|->//]; have [//|neq_xy] := eqVneq x y. +have dxy_gt0 : `|[x - y]| > 0 by rewrite normm_gt0 subr_eq0. +have dxy_ge0 := ltrW dxy_gt0. +have /cl_xy /= := (@entourages_ball _ V ((PosNum dxy_gt0)%:num / 2)%:pos). +rewrite -subr_lt0 ler_gtF // -[X in X - _]mulr1 -mulrBr mulr_ge0 //. +by rewrite subr_ge0 -(@ler_pmul2r _ 2) // mulVf // mul1r ler1n. +Qed. +Lemma eq_close x y : close x y -> x = y. by rewrite closeE. Qed. -Lemma ball_norm_eq x y : (forall eps : posreal, ball_norm x eps%:num y) -> x = y. +Lemma ball_norm_eq x y : (forall e : {posnum K}, ball norm x e%:num y) -> x = y. Proof. by rewrite -norm_close closeE. Qed. Lemma flim_unique {F} {FF : ProperFilter F} : is_prop [set x : V | F --> x]. Proof. by move=> Fx Fy; rewrite -closeE; apply: flim_close. Qed. -Lemma locally_flim_unique (x y : V) : x --> y -> x = y. -Proof. by rewrite -closeE; apply: flim_close. Qed. +Lemma locally_flim_unique (x y : V) : x --> y -> x = y. +Proof. by rewrite -closeE; apply: flim_close. Qed. + +Lemma lim_id (x : V) : lim x = x. +Proof. by symmetry; apply: locally_flim_unique; apply/cvg_ex; exists x. Qed. + +Lemma flim_lim {F} {FF : ProperFilter F} (l : V) : + F --> l -> lim F = l. +Proof. by move=> Fl; have Fcv := cvgP Fl; apply: flim_unique. Qed. + +Lemma flim_map_lim {T : Type} {F} {FF : ProperFilter F} (f : T -> V) (l : V) : + f @ F --> l -> lim (f @ F) = l. +Proof. exact: flim_lim. Qed. + +Lemma flimi_unique T {F} {FF : ProperFilter F} (f : T -> set V) : + {near F, is_fun f} -> is_prop [set x : V | f `@ F --> x]. +Proof. by move=> ffun fx fy; rewrite -closeE; apply: flimi_close. Qed. + +Lemma flim_normW {F : set (set V)} {FF : Filter F} (y : V) : + (forall e, 0 < e -> \forall y' \near F, `|[y - y']| <= e) -> F --> y. +Proof. +move=> cv; apply/flim_normP => _/posnumP[e]; near=> x. +by apply: normm_leW => //; near: x; apply: cv. +Grab Existential Variables. all: end_near. Qed. + +Lemma flimi_map_lim T {F} {FF : ProperFilter F} (f : T -> V -> Prop) (l : V) : + F (fun x : T => is_prop (f x)) -> + f `@ F --> l -> lim (f `@ F) = l. +Proof. +move=> f_prop f_l; apply: get_unique => // l' f_l'. +exact: flimi_unique _ f_l' f_l. +Qed. + +End NormedModuleField. + +Arguments flim_normW {_ _ F FF}. + +Section NormedUniformity. + +Variable (K : realFieldType) (V : zmodType) (norm : V -> K). +Hypothesis (norm0 : norm 0 = 0). +Hypothesis (normB : forall x y, norm (x - y) = norm (y - x)). +Hypothesis (ler_dist_add : forall z x y, + norm (x - y) <= norm (x - z) + norm (z - y)). + +Program Definition uniformityOfNormMixin := + @Uniform.Mixin V (locally_ (entourages_ norm)) (entourages_ norm) _ _ _ _ _. +Next Obligation. +apply: filter_from_filter; first by exists 1. +move=> _ _ /posnumP[e1] /posnumP[e2]; exists (minr e1%:num e2%:num) => // xy. +by rewrite ltr_minr => /andP. +Qed. +Next Obligation. +by case: H => _/posnumP[e] sA xy xey; apply: sA; rewrite xey subrr norm0. +Qed. +Next Obligation. +case: H => _/posnumP[e] sA; exists e%:num => // xy; rewrite normB => xey. +exact: sA. +Qed. +Next Obligation. +case: H => _/posnumP[e] sA; exists [set xy | norm (xy.1 - xy.2) < e%:num / 2]. + by exists (e%:num / 2). +move=> xz [y /= xy yz]; apply: sA; apply: ler_lt_trans (ler_dist_add y _ _) _. +by rewrite [e%:num]splitr ltr_add. +Qed. + +End NormedUniformity. + +Definition realField_uniformMixin (R : realFieldType) := + uniformityOfNormMixin (@normr0 R) (@distrC _) (@ler_dist_add _). + +Canonical realField_filteredType (R : realFieldType) := + FilteredType R R (locally_ (entourages_ normr)). +Canonical realField_topologicalType (R : realFieldType) := + TopologicalType R (topologyOfEntourageMixin (realField_uniformMixin R)). +Canonical realField_uniformType (R : realFieldType):= + UniformType R (realField_uniformMixin R). + +Definition realField_normedModMixin (R : realFieldType) := + @NormedModule.Mixin _ (GRing.regular_lmodType R) + (locally_ (entourages_ normr)) (realField_uniformMixin _) normr + (@ler_norm_add _) (@normrM _) erefl (@normr0_eq0 _). + +Canonical realFieldo_normedModType (R : realFieldType) := + NormedModType R R^o (realField_normedModMixin _). +Canonical realField_normedModType (R : realFieldType) := + [normedModType R of R for realFieldo_normedModType R]. + +Canonical archiField_filteredType (R : archiFieldType) := + [filteredType R of R for realField_filteredType R]. +Canonical archiField_topologicalType (R : archiFieldType) := + [topologicalType of R for realField_topologicalType R]. +Canonical archiField_uniformType (R : archiFieldType) := + [uniformType of R for realField_uniformType R]. +Canonical archiField_normedModType (R : archiFieldType) := + [normedModType R of R for realField_normedModType R]. + +Canonical rcf_filteredType (R : rcfType) := + [filteredType R of R for realField_filteredType R]. +Canonical rcf_topologicalType (R : rcfType) := + [topologicalType of R for realField_topologicalType R]. +Canonical rcf_uniformType (R : rcfType) := + [uniformType of R for realField_uniformType R]. +Canonical rcf_normedModType (R : rcfType) := + [normedModType R of R for realField_normedModType R]. + +Canonical real_filteredType (R : realType) := + [filteredType R of R for realField_filteredType R]. +Canonical real_topologicalType (R : realType) := + [topologicalType of R for realField_topologicalType R]. +Canonical real_uniformType (R : realType) := + [uniformType of R for realField_uniformType R]. +Canonical real_normedModType (R : realType) := + [normedModType R of R for realField_normedModType R]. + +Section NormedModuleRealField. + +Context {K : realFieldType} {V : normedModType K}. + +Lemma ler_addgt0Pr (x y : K) : reflect (forall e, e > 0 -> x <= y + e) (x <= y). +Proof. +apply/(iffP idP)=> [lexy _/posnumP[e] | lexye]; first by rewrite ler_paddr. +case: (lerP x y) => // ltyx. +have /midf_lt [_] := ltyx; rewrite ltrNge -eqbF_neg => /eqP<-. +suff -> : (y + x) / 2 = y + (x - y) / 2. + by apply/lexye/divr_gt0 => //; rewrite subr_gt0. +by rewrite !mulrDl addrC -mulN1r -mulrA mulN1r [RHS]addrC {3}(splitr y) + [RHS]GRing.subrKA. +Qed. + +Lemma ler_addgt0Pl (x y : K) : reflect (forall e, e > 0 -> x <= e + y) (x <= y). +Proof. +by apply/(equivP (ler_addgt0Pr x y)); split=> lexy e /lexy; rewrite addrC. +Qed. -Lemma lim_id (x : V) : lim x = x. -Proof. by symmetry; apply: locally_flim_unique; apply/cvg_ex; exists x. Qed. +Lemma in_segment_addgt0Pr (x y z : K) : + reflect (forall e, e > 0 -> y \in `[(x - e), (z + e)]) (y \in `[x, z]). +Proof. +apply/(iffP idP)=> [xyz _/posnumP[e] | xyz_e]. + rewrite inE; apply/andP; split; last by rewrite ler_paddr // (itvP xyz). + by rewrite ler_subl_addr ler_paddr // (itvP xyz). +rewrite inE; apply/andP. +by split; apply/ler_addgt0Pr => ? /xyz_e /andP []; rewrite ler_subl_addr. +Qed. -Lemma flim_lim {F} {FF : ProperFilter F} (l : V) : - F --> l -> lim F = l. -Proof. by move=> Fl; have Fcv := cvgP Fl; apply: flim_unique. Qed. +Lemma in_segment_addgt0Pl (x y z : K) : + reflect (forall e, e > 0 -> y \in `[(- e + x), (e + z)]) (y \in `[x, z]). +Proof. +apply/(equivP (in_segment_addgt0Pr x y z)). +by split=> zxy e /zxy; rewrite [z + _]addrC [_ + x]addrC. +Qed. -Lemma flim_map_lim {T : Type} {F} {FF : ProperFilter F} (f : T -> V) (l : V) : - f @ F --> l -> lim (f @ F) = l. -Proof. exact: flim_lim. Qed. +Lemma realField_hausdorff : hausdorff [topologicalType of K]. +Proof. +move=> x y clxy. +apply/eqP; rewrite eqr_le; apply/in_segment_addgt0Pr => _ /posnumP[e]. +rewrite inE -ler_distl; set he := (e%:num / 2)%:pos. +have [z []] := clxy _ _ (locally_ball x he) (locally_ball y he). +rewrite /ball normmB => zx yz; apply: ler_trans (ler_dist_add z _ _) _. +by rewrite ltrW // [e%:num]splitr ltr_add. +Qed. Lemma normedModType_hausdorff : hausdorff V. Proof. -move=> p q clp_q; apply/subr0_eq/normm0_eq0/Rhausdorff => A B pq_A. -rewrite -normm0 -(subrr p) => pp_B. +move=> p q clp_q; apply/subr0_eq/normm0_eq0/realField_hausdorff => A B pq_A. +rewrite -(@normm0 _ V) -(subrr p) => pp_B. suff loc_preim r C : locally `|[p - r]| C -> locally r ((fun r => `|[p - r]|) @^-1` C). have [r []] := clp_q _ _ (loc_preim _ _ pp_B) (loc_preim _ _ pq_A). by exists `|[p - r]|. -move=> [e egt0 pre_C]; apply: locally_le_locally_norm; exists e => // s re_s. -apply: pre_C; apply: ler_lt_trans (ler_distm_dist _ _) _. +move=> [D [_/posnumP[e] seD] sDC]; apply: locally_le_locally_norm. +exists [set pq | `|[pq.1 - pq.2]| < e%:num]. + by rewrite entourages_normE; apply: entourages_ball. +move=> s re_s; apply/sDC/seD => /=; apply: ler_lt_trans (ler_distm_dist _ _) _. by rewrite opprB addrC -subr_trans normmB. Qed. -End NormedModule1. +End NormedModuleRealField. -Module Export LocallyNorm. -Definition locally_simpl := (locally_simpl,@locally_locally_norm,@filter_from_norm_locally). -End LocallyNorm. +(** * Adding infinities to realFieldTypes *) -Module Export NearNorm. -Definition near_simpl := (@near_simpl, @locally_normE, - @filter_from_normE, @near_locally_norm). -Ltac near_simpl := rewrite ?near_simpl. -End NearNorm. +Section RealFieldBar. -Section NormedModule2. +Context {K : realFieldType}. -Context {T : Type} {K : absRingType} {V : normedModType K}. +Inductive realFieldBar := | Finite of K | p_infty | m_infty. -Lemma flimi_unique {F} {FF : ProperFilter F} (f : T -> set V) : - {near F, is_fun f} -> is_prop [set x : V | f `@ F --> x]. -Proof. by move=> ffun fx fy; rewrite -closeE; apply: flimi_close. Qed. +Notation "'+oo'" := p_infty : real_scope. +Notation "'-oo'" := m_infty : real_scope. -Lemma flim_normP {F : set (set V)} {FF : Filter F} (y : V) : - F --> y <-> forall eps : R, 0 < eps -> \forall y' \near F, `|[y - y']| < eps. -Proof. by rewrite -filter_fromP /= !locally_simpl. Qed. +Definition realFieldBar_locally' (a : realFieldBar) (A : set K) := + match a with + | Finite a => locally' a A + | +oo => exists M : K, forall x, M < x -> A x + | -oo => exists M : K, forall x, x < M -> A x + end. +Definition realFieldBar_locally (a : realFieldBar) (A : set K) := + match a with + | Finite a => locally a A + | +oo => exists M : K, forall x, M < x -> A x + | -oo => exists M : K, forall x, x < M -> A x + end. -Lemma flim_normW {F : set (set V)} {FF : Filter F} (y : V) : - (forall eps : R, 0 < eps -> \forall y' \near F, `|[y - y']| <= eps) -> - F --> y. +Canonical realFieldBar_eqType := EqType realFieldBar gen_eqMixin. +Canonical realFieldBar_choiceType := ChoiceType realFieldBar gen_choiceMixin. +Canonical realFieldBar_pointedType := PointedType realFieldBar (+oo). +Canonical realFieldBar_filteredType := + FilteredType K realFieldBar (realFieldBar_locally). + +Global Instance realField_locally'_proper (x : K) : ProperFilter (locally' x). Proof. -move=> cv; apply/flim_normP => _/posnumP[e]; near=> x. -by apply: normm_leW => //; near: x; apply: cv. -Grab Existential Variables. all: end_near. Qed. +apply: Build_ProperFilter => A; rewrite /locally' -locally_locally_norm. +move=> [B [_/posnumP[e] seB] sBA]. +exists (x + e%:num / 2); apply: sBA; last first. + by rewrite eq_sym addrC -subr_eq subrr eq_sym. +apply: seB; rewrite opprD addrA subrr sub0r normmN [ `|[_]| ]ger0_norm //. +by rewrite {2}(splitr e%:num) ltr_spaddl. +Qed. -Lemma flim_norm {F : set (set V)} {FF : Filter F} (y : V) : - F --> y -> forall eps, eps > 0 -> \forall y' \near F, `|[y - y']| < eps. -Proof. by move=> /flim_normP. Qed. +Global Instance realFieldBar_locally'_filter : + forall x, ProperFilter (realFieldBar_locally' x). +Proof. +case=> [x||]; first exact: realField_locally'_proper. + apply Build_ProperFilter. + by move=> P [M gtMP]; exists (M + 1); apply: gtMP; rewrite ltr_addl. + split=> /= [|P Q [MP gtMP] [MQ gtMQ] |P Q sPQ [M gtMP]]; first by exists 0. + by exists (maxr MP MQ) => ?; rewrite ltr_maxl => /andP [/gtMP ? /gtMQ]. + by exists M => ? /gtMP /sPQ. +apply Build_ProperFilter. + by move=> P [M ltMP]; exists (M - 1); apply: ltMP; rewrite gtr_addl oppr_lt0. +split=> /= [|P Q [MP ltMP] [MQ ltMQ] |P Q sPQ [M ltMP]]; first by exists 0. + by exists (minr MP MQ) => ?; rewrite ltr_minr => /andP [/ltMP ? /ltMQ]. +by exists M => ? /ltMP /sPQ. +Qed. +Typeclasses Opaque realFieldBar_locally'. -Lemma flim_bounded {F : set (set V)} {FF : Filter F} (y : V) : - F --> y -> \forall M \near +oo, \forall y' \near F, `|[y']|%real < M. +Global Instance realFieldBar_locally_filter : + forall x, ProperFilter (realFieldBar_locally x). Proof. -move=> /flim_norm Fy; exists `|[y]| => M. -rewrite -subr_gt0 => subM_gt0; have := Fy _ subM_gt0. -apply: filterS => y' yy'; rewrite -(@ltr_add2r _ (- `|[y]|)). -rewrite (ler_lt_trans _ yy') //. -by rewrite (ler_trans _ (ler_distm_dist _ _)) // absRE distrC ler_norm. +case=> [x||]; first exact/locally_filter. + exact: (realFieldBar_locally'_filter +oo). +exact: (realFieldBar_locally'_filter -oo). Qed. +Typeclasses Opaque realFieldBar_locally. -Lemma flimi_map_lim {F} {FF : ProperFilter F} (f : T -> V -> Prop) (l : V) : - F (fun x : T => is_prop (f x)) -> - f `@ F --> l -> lim (f `@ F) = l. +Lemma near_pinfty_div2 (A : set K) : + (\forall k \near +oo, A k) -> (\forall k \near +oo, A (k / 2)). Proof. -move=> f_prop f_l; apply: get_unique => // l' f_l'. -exact: flimi_unique _ f_l' f_l. +by move=> [M AM]; exists (M * 2) => x; rewrite -ltr_pdivl_mulr //; apply: AM. Qed. -End NormedModule2. -Hint Resolve normm_ge0. -Arguments flim_norm {_ _ F FF}. -Arguments flim_bounded {_ _ F FF}. +Lemma locally_pinfty_gt c : \forall x \near +oo, c < x. +Proof. by exists c. Qed. + +Lemma locally_pinfty_ge c : \forall x \near +oo, c <= x. +Proof. by exists c; apply: ltrW. Qed. -Lemma continuous_flim_norm {K : absRingType} (V W : normedModType K) (f : V -> W) x l : - continuous f -> x --> l -> forall e : posreal, `|[f l - f x]| < e. +End RealFieldBar. + +Notation "'+oo'" := p_infty : real_scope. +Notation "'-oo'" := m_infty : real_scope. +Notation "'+oo' 'in' K" := (@p_infty K) (at level 0) : real_scope. +Notation "'-oo' 'in' K" := (@m_infty K) (at level 0) : real_scope. +Hint Extern 0 (is_true (0 < _)) => match goal with + H : ?x \is_near (locally +oo) |- _ => + solve[near: x; exists 0 => _/posnumP[x] //] end : core. + +Lemma flim_bounded {K : realFieldType} {V : normedModType K} {F : set (set V)} + {FF : Filter F} (y : V) : + F --> y -> \forall M \near +oo, \forall y' \near F, `|[y']| < M. Proof. -move=> cf xl e. -move/flim_norm: (cf l) => /(_ _ (posnum_gt0 e)). -rewrite nearE /= => /locallyP; rewrite locally_E => -[i i0]; apply. -have /@flim_norm : Filter [filter of x] by apply: filter_on_Filter. -move/(_ _ xl _ i0). -rewrite nearE /= => /locallyP; rewrite locally_E => -[j j0]. -move/(_ _ (ballxx _ j0)); by rewrite -ball_normE. +move=> /flim_norm Fy; exists `|[y]| => M. +rewrite -subr_gt0 => subM_gt0; have := Fy _ subM_gt0. +apply: filterS => y' yy'; rewrite -(@ltr_add2r _ (- `|[y]|)). +rewrite (ler_lt_trans _ yy') // (ler_trans _ (ler_distm_dist _ _)) //. +by rewrite distrC ler_norm. Qed. +Arguments flim_bounded {_ _ F FF}. (** ** Matrices *) Section matrix_normedMod. -Variables (K : absRingType) (m n : nat). +Variables (K : realFieldType) (m n : nat). -(* take m.+1, n.+1 because ball_normE is not provable for m = 0 or n = 0 *) +(* take m.+1,n.+1 because entourages_normE is not provable for m = 0 or n = 0 *) Definition mx_norm (x : 'M[K]_(m.+1, n.+1)) := bigmaxr 0 [seq `|x ij.1 ij.2| | ij : 'I_m.+1 * 'I_n.+1]. @@ -1665,41 +782,62 @@ Program Definition matrix_NormedModMixin := Next Obligation. apply/bigmaxr_lerP=> [|i]; rewrite size_map -cardT mxvec_cast // => ltimn. rewrite (nth_map (ord0, ord0)); last by rewrite -cardT mxvec_cast. -rewrite mxE; apply: ler_trans (ler_abs_add _ _) _. +rewrite mxE; apply: ler_trans (ler_norm_add _ _) _. do 2 ?[rewrite -(nth_map _ 0 (fun p => `|_ p.1 p.2|)) -?cardT ?mxvec_cast //]. by apply: ler_add; apply: bigmaxr_ler; rewrite size_map -cardT mxvec_cast. Qed. Next Obligation. apply/bigmaxrP; split=> [|i]. - have : (0 < size [seq `|x ij.1 ij.2|%real | ij : 'I_m.+1 * 'I_n.+1])%N. + have : (0 < size [seq `|x ij.1 ij.2|%R | ij : 'I_m.+1 * 'I_n.+1])%N. by rewrite size_map -cardT mxvec_cast. move=> /bigmaxr_mem - /(_ 0) /mapP [ij ijP normx]; rewrite [mx_norm _]normx. - by apply/mapP; exists ij => //; rewrite mxE absrM. + by apply/mapP; exists ij => //; rewrite mxE normrM. rewrite size_map -cardT mxvec_cast // => ltimn. rewrite (nth_map (ord0, ord0)); last by rewrite -cardT mxvec_cast. -rewrite mxE absrM ler_pmul // -(nth_map _ 0 (fun p => `|x p.1 p.2|)). +rewrite mxE normrM ler_pmul // -(nth_map _ 0 (fun p => `|x p.1 p.2|)). by apply: bigmaxr_ler; rewrite size_map -cardT mxvec_cast. by rewrite -cardT mxvec_cast. Qed. Next Obligation. -rewrite predeq3E => x e y; split. - move=> xe_y; apply/bigmaxr_ltrP; rewrite size_map -cardT mxvec_cast //. - move=> i iltmn; rewrite (nth_map (ord0, ord0)). - by rewrite !mxE; apply: xe_y. - by rewrite -cardT mxvec_cast. -move=> /bigmaxr_ltrP; rewrite size_map -cardT mxvec_cast => xe_y i j. -have := xe_y _ (index (i, j) (enum [finType of 'I_m.+1 * 'I_n.+1])). +rewrite predeqE => A; split; last first. + move=> [_/posnumP[e] sA]; exists (fun _ _ => [set p | `|p.1 - p.2| < e%:num]). + by move=> _ _; exists e%:num. + move=> [x y] /= xy; apply: sA => /=. + apply/bigmaxr_ltrP; rewrite size_map -cardT mxvec_cast // => i ltimn. + rewrite (nth_map (ord0, ord0)); last by rewrite -cardT mxvec_cast. + by rewrite !mxE; apply: xy. +move=> [P entP sPA]; have {entP} entP : forall i j, exists e, 0 < e /\ + [set pq | `|pq.1 - pq.2| < e] `<=` P i j. + by move=> i j; have [e] := entP i j; exists e. +set e := fun i j => get [set e | 0 < e /\ + [set pq | `|pq.1 - pq.2| < e] `<=` P i j]. +exists (- bigmaxr 0 [seq - e ij.1 ij.2 | ij : 'I_m.+1 * 'I_n.+1]). + rewrite oppr_gt0; apply/bigmaxr_ltrP; rewrite size_map -cardT mxvec_cast //. + move=> i ltimn; rewrite (nth_map (ord0, ord0)); last first. + by rewrite -cardT mxvec_cast. + by rewrite oppr_lt0; set j := _.1; set k := _.2; have /getPex[] := entP j k. +move=> [x y] /= xy; apply: sPA => i j /=. +have s0 : (0 < size [seq `|(x - y) ij.1 ij.2|%R | ij : 'I_m.+1 * 'I_n.+1])%N. + by rewrite size_map -cardT mxvec_cast. have memij : (i, j) \in enum [finType of 'I_m.+1 * 'I_n.+1]. by rewrite mem_enum. +have /getPex [_] := entP i j; apply=> /=; rewrite -[get _]/(e i j). +move: xy => /bigmaxr_ltrP - /(_ s0 (index (i, j) + (enum [finType of 'I_m.+1 * 'I_n.+1]))). +rewrite size_map index_mem => /(_ memij). rewrite (nth_map (ord0, ord0)); last by rewrite index_mem. -by rewrite !mxE !nth_index //=; apply=> //; rewrite -mxvec_cast cardT index_mem. +rewrite !nth_index //= !mxE => /ltr_le_trans; apply. +rewrite ler_oppl; have /(nth_index 0) <- : + - e i j \in [seq - e ij.1 ij.2 | ij : 'I_m.+1 * 'I_n.+1]. + by apply/mapP; exists (i, j). +by apply: bigmaxr_ler; rewrite index_mem; apply/mapP; exists (i, j). Qed. Next Obligation. apply/matrixP => i j; rewrite mxE; apply/eqP. -rewrite -absr_eq0 eqr_le; apply/andP; split; last exact: absr_ge0. +rewrite -normr_eq0 eqr_le; apply/andP; split; last exact: normr_ge0. have /(bigmaxr_ler 0) : (index (i, j) (enum [finType of 'I_m.+1 * 'I_n.+1]) < - size [seq (`|x ij.1 ij.2|)%real | ij : 'I_m.+1 * 'I_n.+1])%N. + size [seq (`|x ij.1 ij.2|)%R | ij : 'I_m.+1 * 'I_n.+1])%N. by rewrite size_map index_mem mem_enum. rewrite -{3}H; apply: ler_trans. rewrite (nth_map (ord0, ord0)); last by rewrite index_mem mem_enum. @@ -1711,11 +849,26 @@ Canonical matrix_normedModType := End matrix_normedMod. +Lemma coord_continuous {R : realFieldType} m n i j : + continuous (fun M : 'M[R]_(m.+1, n.+1) => M i j). +Proof. +move=> M s /=; rewrite -filter_from_norm_locally => -[e e0 es]. +apply/locallyP; rewrite locally_entourageE -filter_from_norm_locally. +exists e => // N MN; apply/es; have /bigmaxr_ltrP MeN := MN. +have /(MeN _) : + (index (i, j) (enum [finType of 'I_m.+1 * 'I_n.+1]) < + size [seq (`|(M - N) ij.1 ij.2|)%R | ij : 'I_m.+1 * 'I_n.+1])%N. + by rewrite size_map index_mem mem_enum. +rewrite (nth_map (ord0, ord0)); last by rewrite index_mem mem_enum. +rewrite nth_index ?mem_enum // !mxE; apply. +by rewrite size_map -cardT mxvec_cast. +Qed. + (** ** Pairs *) Section prod_NormedModule. -Context {K : absRingType} {U V : normedModType K}. +Context {K : realDomainType} {U V : normedModType K}. Definition prod_norm (x : U * V) := maxr `|[x.1]| `|[x.2]|. @@ -1727,13 +880,29 @@ by move=> [xu xv] [yu yv]; rewrite ler_maxl /=; apply/andP; split; Qed. Lemma prod_norm_scal (l : K) (x : U * V) : - prod_norm (l *: x) = abs l * prod_norm x. + prod_norm (l *: x) = `|l| * prod_norm x. Proof. by rewrite /prod_norm !normmZ maxr_pmulr. Qed. -Lemma ball_prod_normE : ball = ball_ prod_norm. -Proof. -rewrite funeq2E => - [xu xv] e; rewrite predeqE => - [yu yv]. -by rewrite /ball /= /prod_ball -!ball_normE /ball_ ltr_maxl; split=> /andP. +Lemma entourages_prod_normE : entourages = entourages_ prod_norm. +Proof. +rewrite predeqE => A; split; last first. + move=> [_/posnumP[e] sA]. + exists ([set u | `|[u.1 - u.2]| < e%:num], [set v | `|[v.1 - v.2]| < e%:num]). + by split=> /=; apply: entourages_ball. + move=> /= uv [uv1e uv2e]; exists ((uv.1.1, uv.2.1), (uv.1.2, uv.2.2)). + by apply: sA; rewrite ltr_maxl uv1e uv2e. + by rewrite /= -!surjective_pairing. +move=> [PQ []]; rewrite -!entourages_normE. +move=> [_/posnumP[eP] sP] [_/posnumP[eQ] sQ] sPQA. +exists (minr eP%:num eQ%:num) => // xy. +rewrite ltr_maxl !ltr_minr => /andP [/andP [xy1P xy1Q] /andP [xy2P xy2Q]]. +have PQxy1 : PQ.1 (xy.1.1, xy.2.1) by apply: sP. +have /(conj PQxy1) : PQ.2 (xy.1.2, xy.2.2) by apply: sQ. +move=> /(sPQA ((xy.1.1, xy.2.1), (xy.1.2, xy.2.2))) [uv Auv]. +move=> /eqP /andP [/andP [/= /eqP uvxy11 /eqP uvxy21] /andP + [/= /eqP uvxy12 /eqP uvxy22]]. +rewrite [xy]surjective_pairing [_.2]surjective_pairing [_.1]surjective_pairing. +by rewrite -uvxy11 -uvxy12 -uvxy21 -uvxy22 -!surjective_pairing. Qed. Lemma prod_norm_eq0 (x : U * V) : prod_norm x = 0 -> x = 0. @@ -1747,23 +916,24 @@ Qed. End prod_NormedModule. -Definition prod_NormedModule_mixin (K : absRingType) (U V : normedModType K) := +Definition prod_NormedModule_mixin (K : realDomainType) + (U V : normedModType K) := @NormedModMixin K _ _ _ (@prod_norm K U V) prod_norm_triangle - prod_norm_scal ball_prod_normE prod_norm_eq0. + prod_norm_scal entourages_prod_normE prod_norm_eq0. -Canonical prod_NormedModule (K : absRingType) (U V : normedModType K) := +Canonical prod_NormedModule (K : realDomainType) (U V : normedModType K) := NormedModType K (U * V) (@prod_NormedModule_mixin K U V). Section NormedModule3. -Context {T : Type} {K : absRingType} {U : normedModType K} +Context {T : Type} {K : realDomainType} {U : normedModType K} {V : normedModType K}. Lemma flim_norm2P {F : set (set U)} {G : set (set V)} {FF : Filter F} {FG : Filter G} (y : U) (z : V): (F, G) --> (y, z) <-> - forall eps : R, 0 < eps -> - \forall y' \near F & z' \near G, `|[(y, z) - (y', z')]| < eps. + forall e : K, 0 < e -> + \forall y' \near F & z' \near G, `|[(y, z) - (y', z')]| < e. Proof. exact: flim_normP. Qed. (* Lemma flim_norm_supP {F : set (set U)} {G : set (set V)} *) @@ -1780,26 +950,18 @@ Proof. exact: flim_normP. Qed. Lemma flim_norm2 {F : set (set U)} {G : set (set V)} {FF : Filter F} {FG : Filter G} (y : U) (z : V): (F, G) --> (y, z) -> - forall eps : R, 0 < eps -> - \forall y' \near F & z' \near G, `|[(y, z) - (y', z')]| < eps. + forall e : K, 0 < e -> + \forall y' \near F & z' \near G, `|[(y, z) - (y', z')]| < e. Proof. by rewrite flim_normP. Qed. End NormedModule3. Arguments flim_norm2 {_ _ _ F G FF FG}. -(** Rings with absolute values are normed modules *) - -Definition AbsRing_NormedModMixin (K : absRingType) := - @NormedModule.Mixin K _ _ _ (abs : K^o -> R) ler_abs_add absrM (ball_absE K) - absr0_eq0. -Canonical AbsRing_NormedModType (K : absRingType) := - NormedModType K K^o (AbsRing_NormedModMixin _). - (** Normed vector spaces have some continuous functions *) Section NVS_continuity. -Context {K : absRingType} {V : normedModType K}. +Context {K : realFieldType} {V : normedModType K}. Lemma add_continuous : continuous (fun z : V * V => z.1 + z.2). Proof. @@ -1811,15 +973,15 @@ Grab Existential Variables. all: end_near. Qed. Lemma scale_continuous : continuous (fun z : K * V => z.1 *: z.2). Proof. move=> [k x]; apply/flim_normP=> _/posnumP[e]. -rewrite !near_simpl /=; near +oo => M; near=> l z => /=. +rewrite !near_simpl /=; near +oo in K => M; near=> l z => /=. rewrite (@distm_lt_split _ _ (k *: z)) // -?(scalerBr, scalerBl) normmZ. - rewrite (ler_lt_trans (ler_pmul _ _ (_ : _ <= `|k|%real + 1) (lerr _))) + rewrite (ler_lt_trans (ler_pmul _ _ (_ : _ <= `|k| + 1) (lerr _))) ?ler_addl //. rewrite -ltr_pdivl_mull // ?(ltr_le_trans ltr01) ?ler_addr //; near: z. by apply: flim_norm; rewrite // mulr_gt0 // ?invr_gt0 ltr_paddl. have zM: `|[z]| < M by near: z; near: M; apply: flim_bounded; apply: flim_refl. rewrite (ler_lt_trans (ler_pmul _ _ (lerr _) (_ : _ <= M))) // ?ltrW//. -by rewrite -ltr_pdivl_mulr //; near: l; apply: (flim_norm (_ : K^o)). +by rewrite -ltr_pdivl_mulr //; near: l; apply: flim_norm. Grab Existential Variables. all: end_near. Qed. Arguments scale_continuous _ _ : clear implicits. @@ -1845,36 +1007,7 @@ End NVS_continuity. Section limit_composition. -Lemma lim_cont {T V U : topologicalType} - (F : set (set T)) (FF : Filter F) - (f : T -> V) (h : V -> U) (a : V) : - {for a, continuous h} -> - f @ F --> a -> (h \o f) @ F --> h a. -Proof. -move=> h_cont fa fb; apply: (flim_trans _ h_cont). -exact: (@flim_comp _ _ _ _ h _ _ _ fa). -Qed. - -Lemma lim_cont2 {T V W U : topologicalType} - (F : set (set T)) (FF : Filter F) - (f : T -> V) (g : T -> W) (h : V -> W -> U) (a : V) (b : W) : - h z.1 z.2 @[z --> (a, b)] --> h a b -> - f @ F --> a -> g @ F --> b -> (fun x => h (f x) (g x)) @ F --> h a b. -Proof. -move=> h_cont fa fb; apply: (flim_trans _ h_cont). -exact: (@flim_comp _ _ _ _ (fun x => h x.1 x.2) _ _ _ (flim_pair fa fb)). -Qed. - -Lemma cst_continuous {T T' : topologicalType} (k : T') - (F : set (set T)) {FF : Filter F} : - (fun _ : T => k) @ F --> k. -Proof. -by move=> x; rewrite !near_simpl => /locally_singleton ?; apply: filterE. -Qed. -Arguments cst_continuous {T T'} k F {FF}. -Hint Resolve cst_continuous. - -Context {K : absRingType} {V : normedModType K} {T : topologicalType}. +Context {K : realFieldType} {V : normedModType K} {T : topologicalType}. Lemma lim_cst (a : V) (F : set (set V)) {FF : Filter F} : (fun=> a) @ F --> a. Proof. exact: cst_continuous. Qed. @@ -1919,13 +1052,47 @@ Lemma continuousN (f : T -> V) x : Proof. by move=> ?; apply: lim_opp. Qed. Lemma lim_mult (x y : K) : z.1 * z.2 @[z --> (x, y)] --> x * y. -Proof. exact: (@scale_continuous _ (AbsRing_NormedModType K)). Qed. +Proof. exact: scale_continuous. Qed. Lemma continuousM (f g : T -> K) x : {for x, continuous f} -> {for x, continuous g} -> {for x, continuous (fun x => f x * g x)}. Proof. by move=> fc gc; apply: flim_comp2 fc gc _; apply: lim_mult. Qed. +(** Continuity of norm *) + +Lemma continuous_norm : continuous (@norm _ V). +Proof. +move=> x; apply/flim_normP => _/posnumP[e] /=. +rewrite !near_simpl -locally_nearE -filter_from_norm_locally. +by exists e%:num => // y Hy; apply/(ler_lt_trans (ler_distm_dist _ _)). +Qed. + +(* :TODO: yet, not used anywhere?! *) +Lemma flim_norm0 {U} {F : set (set U)} {FF : Filter F} (f : U -> V) : + (fun x => `|[f x]|) @ F --> (0 : K) + -> f @ F --> (0 : V). +Proof. +move=> /flim_norm fx0; apply/flim_normP => e egt0. +rewrite near_simpl; have := fx0 _ egt0; rewrite near_simpl. +by apply: filterS => x; rewrite !sub0r !normmN [ `|[_]| ]ger0_norm. +Qed. + +(* TODO: simplify using extremumP when PR merged in mathcomp *) +Lemma cvg_seq_bounded (a : nat -> V) : + [cvg a in V] -> {M : K | forall n, norm (a n) <= M}. +Proof. +move=> a_cvg; suff: exists M, forall n, norm (a n) <= M. + by move=> /getPex; set M := get _; exists M. +near +oo in K => M. +have [//|N _ /(_ _ _) /ltrW a_leM] := !! near (flim_bounded _ a_cvg) M. +exists (maxr M (\big[maxr/M]_(n < N) `|[a (val (rev_ord n))]|)) => /= n. +rewrite ler_maxr; have [nN|nN] := leqP N n; first by rewrite a_leM. +apply/orP; right => {a_leM}; elim: N n nN=> //= N IHN n. +rewrite leq_eqVlt => /orP[/eqP[->] |/IHN a_le]; +by rewrite big_ord_recl subn1 /= ler_maxr ?a_le ?lerr ?orbT. +Grab Existential Variables. all: end_near. Qed. + End limit_composition. (** ** Complete Normed Modules *) @@ -1934,7 +1101,7 @@ Module CompleteNormedModule. Section ClassDef. -Variable K : absRingType. +Variable K : numDomainType. Record class_of (T : Type) := Class { base : NormedModule.class_of K T ; @@ -2012,103 +1179,16 @@ End CompleteNormedModule. Export CompleteNormedModule.Exports. -Section CompleteNormedModule1. - -Context {K : absRingType} {V : completeNormedModType K}. - - -End CompleteNormedModule1. - (** * Extended Types *) -(** * The topology on natural numbers *) - -(* :TODO: ultimately nat could be replaced by any lattice *) -Definition eventually := filter_from setT (fun N => [set n | (N <= n)%N]). -Notation "'\oo'" := eventually : classical_set_scope. - -Canonical eventually_filter_source X := - @Filtered.Source X _ nat (fun f => f @ \oo). - -Global Instance eventually_filter : ProperFilter eventually. -Proof. -eapply @filter_from_proper; last by move=> i _; exists i. -apply: filter_fromT_filter; first by exists 0%N. -move=> i j; exists (maxn i j) => n //=. -by rewrite geq_max => /andP[ltin ltjn]. -Qed. +Definition entourages_set (U : uniformType) (A : set ((set U) * (set U))) := + exists2 B, entourages B & forall PQ, A PQ -> forall p q, + PQ.1 p -> PQ.2 q -> B (p,q). +Canonical set_filter_source (U : uniformType) := + @Filtered.Source Prop _ U (fun A => locally_ (@entourages_set U) A). (** * The topology on real numbers *) -(* TODO: Remove R_complete_lim and use lim instead *) -(* Definition R_lim (F : (R -> Prop) -> Prop) : R := *) -(* sup (fun x : R => `[]). *) - -(* move: (Lub_Rbar_correct (fun x : R => F (ball (x + 1) 1))). *) -(* move Hl : (Lub_Rbar _) => l{Hl}; move: l => [x| |] [Hx1 Hx2]. *) -(* - case: (HF (Num.min 2 eps%:num / 2)%:pos) => z Hz. *) -(* have H1 : z - Num.min 2 eps%:num / 2 + 1 <= x + 1. *) -(* rewrite ler_add //; apply/RleP/Hx1. *) -(* apply: filterS Hz. *) -(* rewrite /ball /= => u; rewrite /AbsRing_ball absrB ltr_distl. *) -(* rewrite absrB ltr_distl. *) -(* case/andP => {Hx1 Hx2 FF HF x F} Bu1 Bu2. *) -(* have H : Num.min 2 eps%:num <= 2 by rewrite ler_minl lerr. *) -(* rewrite addrK -addrA Bu1 /= (ltr_le_trans Bu2) //. *) -(* rewrite -addrA ler_add // -addrA addrC ler_subr_addl. *) -(* by rewrite ler_add // ler_pdivr_mulr // ?mul1r. *) -(* have H2 : x + 1 <= z + Num.min 2 eps%:num / 2 + 1. *) -(* rewrite ler_add //; apply/RleP/(Hx2 (Finite _)) => v Hv. *) -(* apply: Rbar_not_lt_le => /RltP Hlt. *) -(* apply: filter_not_empty. *) -(* apply: filterS (filterI Hz Hv). *) -(* rewrite /ball /= => w []; rewrite /AbsRing_ball //. *) -(* rewrite absrB ltr_distl => /andP[_ Hw1]. *) -(* rewrite absrB ltr_distl addrK => /andP[Hw2 _]. *) -(* by move: (ltr_trans (ltr_trans Hw1 Hlt) Hw2); rewrite ltrr. *) -(* apply: filterS Hz. *) -(* rewrite /ball /= => u; rewrite /AbsRing_ball absrB absRE 2!ltr_distl. *) -(* case/andP => {Hx1 Hx2 F FF HF} H H0. *) -(* have H3 : Num.min 2 eps%:num <= eps by rewrite ler_minl lerr orbT. *) -(* apply/andP; split. *) -(* - move: H1; rewrite -ler_subr_addr addrK ler_subl_addr => H1. *) -(* rewrite ltr_subl_addr // (ltr_le_trans H0) //. *) -(* rewrite -ler_subr_addr (ler_trans H1) //. *) -(* rewrite -ler_subr_addl -!addrA (addrC x) !addrA subrK. *) -(* rewrite ler_subr_addr -mulrDl ler_pdivr_mulr //. *) -(* by rewrite -mulr2n -mulr_natl mulrC ler_pmul. *) -(* - move: H2; rewrite -ler_subr_addr addrK. *) -(* move/ler_lt_trans; apply. *) -(* move: H; rewrite // ltr_subl_addr => H. *) -(* rewrite -ltr_subr_addr (ltr_le_trans H) //. *) -(* rewrite addrC -ler_subr_addr -!addrA (addrC u) !addrA subrK. *) -(* rewrite -ler_subl_addr opprK -mulrDl ler_pdivr_mulr // -mulr2n -mulr_natl. *) -(* by rewrite mulrC ler_pmul. *) -(* - case (HF 1%:pos) => y Fy. *) -(* case: (Hx2 (y + 1)) => x Fx. *) -(* apply: Rbar_not_lt_le => Hlt. *) -(* apply: filter_not_empty. *) -(* apply: filterS (filterI Fy Fx) => z [Hz1 Hz2]. *) -(* apply: Rbar_le_not_lt Hlt; apply/RleP. *) -(* rewrite -(ler_add2r (-(y - 1))) opprB !addrA -![in X in _ <= X]addrA. *) -(* rewrite (addrC y) ![in X in _ <= X]addrA subrK. *) -(* suff : `|x + 1 - y|%R <= 1 + 1 by rewrite ler_norml => /andP[]. *) -(* rewrite ltrW // (@subr_trans _ z). *) -(* by rewrite (ler_lt_trans (ler_norm_add _ _)) // ltr_add // distrC. *) -(* - case: (HF 1%:pos) => y Fy. *) -(* case: (Hx1 (y - 1)); by rewrite addrAC addrK. *) -(* Qed. *) -(* Admitted. *) - -Arguments flim_normW {_ _ F FF}. - -Lemma filterP_strong T (F : set (set T)) {FF : Filter F} (P : set T) : - (exists Q : set T, exists FQ : F Q, forall x : T, Q x -> P x) <-> F P. -Proof. -split; last by exists P. -by move=> [Q [FQ QP]]; apply: (filterS QP). -Qed. - (* :TODO: add to mathcomp *) Lemma ltr_distW (R : realDomainType) (x y e : R): (`|x - y|%R < e) -> y - e < x. @@ -2119,40 +1199,41 @@ Lemma ler_distW (R : realDomainType) (x y e : R): (`|x - y|%R <= e) -> y - e <= x. Proof. by rewrite ler_distl => /andP[]. Qed. -Lemma R_complete (F : set (set R)) : ProperFilter F -> cauchy F -> cvg F. +Section real_topology. + +Context {R : realType}. + +Lemma real_complete (F : set (set R)) : ProperFilter F -> cauchy F -> cvg F. Proof. -move=> FF F_cauchy; apply/cvg_ex. +move=> FF Fc; apply/cvg_ex. pose D := \bigcap_(A in F) (down (mem A)). -have /cauchyP /(_ 1) [//|x0 x01] := F_cauchy. +have /Fc := @entourages_ball _ [normedModType R of R] 1%:pos. +rewrite near_simpl -near2_pair => /nearP_dep /filter_ex /= [x0 x01]. have D_has_sup : has_sup (mem D); first split. - exists (x0 - 1); rewrite in_setE => A FA. apply/existsbP; near F => x; first exists x. by rewrite ler_distW 1?distrC 1?ltrW ?andbT ?in_setE //; near: x. - end_near. - exists (x0 + 1); apply/forallbP => x; apply/implyP; rewrite in_setE. move=> /(_ _ x01) /existsbP [y /andP[]]; rewrite in_setE. - rewrite -[ball _ _ _]/(_ (_ < _)) ltr_distl ltr_subl_addr => /andP[/ltrW]. - by move=> /(ler_trans _) yx01 _ /yx01. -exists (sup (mem D)). -apply: (flim_normW (_ : R^o)) => /= _ /posnumP[eps]; near=> x. + by rewrite ltr_distl ltr_subl_addr=> /andP[/ltrW] /(ler_trans _) yx01 _ /yx01. +exists (sup (mem D)); apply: flim_normW => /= _ /posnumP[eps]; near=> x. rewrite ler_distl sup_upper_bound //=. apply: sup_le_ub => //; first by case: D_has_sup. apply/forallbP => y; apply/implyP; rewrite in_setE. - move=> /(_ (ball_ norm x eps%:num) _) /existsbP []. - by near: x; apply: nearP_dep; apply: F_cauchy. - move=> z /andP[]; rewrite in_setE /ball_ ltr_distl ltr_subl_addr. + move=> /(_ (ball norm x eps%:num) _) /existsbP []. + by near: x; apply: nearP_dep; apply: Fc; apply: entourages_ball. + move=> z /andP[]; rewrite in_setE /ball ltr_distl ltr_subl_addr. by move=> /andP [/ltrW /(ler_trans _) le_xeps _ /le_xeps]. rewrite in_setE /D /= => A FA; near F => y. apply/existsbP; exists y; apply/andP; split. by rewrite in_setE; near: y. rewrite ler_subl_addl -ler_subl_addr ltrW //. suff: `|x - y| < eps%:num by rewrite ltr_norml => /andP[_]. -by near: y; near: x; apply: nearP_dep; apply: F_cauchy. +by near: y; near: x; apply: nearP_dep; apply: Fc; apply: entourages_ball. Grab Existential Variables. all: end_near. Qed. -Canonical R_completeType := CompleteType R R_complete. -Canonical R_NormedModule := [normedModType R of R^o]. -Canonical R_CompleteNormedModule := [completeNormedModType R of R^o]. +Canonical real_completeType := CompleteType R real_complete. +Canonical real_completeNormedModType := [completeNormedModType R of R]. Definition at_left x := within (fun u : R => u < x) (locally x). Definition at_right x := within (fun u : R => x < u) (locally x). @@ -2160,70 +1241,38 @@ Definition at_right x := within (fun u : R => x < u) (locally x). Global Instance at_right_proper_filter (x : R) : ProperFilter (at_right x). Proof. -apply: Build_ProperFilter' => -[_ /posnumP[d] /(_ (x + d%:num / 2))]. -apply; last (by rewrite ltr_addl); rewrite /AbsRing_ball /=. -rewrite opprD !addrA subrr add0r absrN absRE normf_div !ger0_norm //. +apply: Build_ProperFilter' => -[A [_/posnumP[d] sA]] /(_ (x + d%:num / 2)). +apply; last by rewrite ltr_addl. +apply: sA; rewrite /= opprD !addrA subrr add0r normrN normf_div !ger0_norm //. by rewrite ltr_pdivr_mulr // ltr_pmulr // (_ : 1 = 1%:R) // ltr_nat. Qed. Global Instance at_left_proper_filter (x : R) : ProperFilter (at_left x). Proof. -apply: Build_ProperFilter' => -[_ /posnumP[d] /(_ (x - d%:num / 2))]. -apply; last (by rewrite ltr_subl_addl ltr_addr); rewrite /AbsRing_ball /=. -rewrite opprD !addrA subrr add0r opprK absRE normf_div !ger0_norm //. +apply: Build_ProperFilter' => -[A [_/posnumP[d] sA]] /(_ (x - d%:num / 2)). +apply; last by rewrite ltr_subl_addl ltr_addr. +apply: sA; rewrite opprD !addrA subrr add0r opprK normf_div !ger0_norm //. by rewrite ltr_pdivr_mulr // ltr_pmulr // (_ : 1 = 1%:R) // ltr_nat. Qed. Typeclasses Opaque at_left at_right. -(** Continuity of norm *) - -Lemma continuous_norm {K : absRingType} {V : normedModType K} : - continuous (@norm _ V). -Proof. -move=> x; apply/(@flim_normP _ [normedModType R of R^o]) => _/posnumP[e] /=. -rewrite !near_simpl; apply/locally_normP; exists e%:num => // y Hy. -exact/(ler_lt_trans (ler_distm_dist _ _)). -Qed. - -(* :TODO: yet, not used anywhere?! *) -Lemma flim_norm0 {U} {K : absRingType} {V : normedModType K} - {F : set (set U)} {FF : Filter F} (f : U -> V) : - (fun x => `|[f x]|) @ F --> (0 : R) - -> f @ F --> (0 : V). -Proof. -move=> /(flim_norm (_ : R^o)) fx0; apply/flim_normP => _/posnumP[e]. -rewrite near_simpl; have := fx0 _ [gt0 of e%:num]; rewrite near_simpl. -by apply: filterS => x; rewrite !sub0r !normmN [ `|[_]| ]ger0_norm. -Qed. - -(* TODO: simplify using extremumP when PR merged in mathcomp *) -Lemma cvg_seq_bounded {K : absRingType} {V : normedModType K} (a : nat -> V) : - [cvg a in V] -> {M : R | forall n, norm (a n) <= M}. -Proof. -move=> a_cvg; suff: exists M, forall n, norm (a n) <= M. - by move=> /getPex; set M := get _; exists M. -near +oo => M. -have [//|N _ /(_ _ _) /ltrW a_leM] := !! near (flim_bounded _ a_cvg) M. -exists (maxr M (\big[maxr/M]_(n < N) `|[a (val (rev_ord n))]|)) => /= n. -rewrite ler_maxr; have [nN|nN] := leqP N n; first by rewrite a_leM. -apply/orP; right => {a_leM}; elim: N n nN=> //= N IHN n. -rewrite leq_eqVlt => /orP[/eqP[->] |/IHN a_le]; -by rewrite big_ord_recl subn1 /= ler_maxr ?a_le ?lerr ?orbT. -Grab Existential Variables. all: end_near. Qed. - (** Some open sets of [R] *) Lemma open_lt (y : R) : open (< y). Proof. -move=> x /=; rewrite -subr_gt0 => yDx_gt0; exists (y - x) => // z. -by rewrite /AbsRing_ball /= absrB ltr_distl addrCA subrr addr0 => /andP[]. +move=> x /=; rewrite -subr_gt0 => yDx_gt0. +rewrite entourages_normE locally_entourageE -filter_from_norm_locally. +exists (y - x) => // z. +by rewrite /ball /= normmB ltr_distl addrCA subrr addr0 => /andP[]. Qed. Hint Resolve open_lt. Lemma open_gt (y : R) : open (> y). Proof. -move=> x /=; rewrite -subr_gt0 => xDy_gt0; exists (x - y) => // z. -by rewrite /AbsRing_ball /= absrB ltr_distl opprB addrCA subrr addr0 => /andP[]. +move=> x /=; rewrite -subr_gt0 => xDy_gt0. +rewrite entourages_normE locally_entourageE -filter_from_norm_locally. +exists (x - y) => // z. +by rewrite /ball /= normmB ltr_distl opprB addrCA subrr addr0 => /andP[]. Qed. Hint Resolve open_gt. @@ -2260,6 +1309,36 @@ Qed. (** properties of segments in [R] *) +Lemma locallyN (x : R) : + locally (- x) = [set [set - y | y in A] | A in locally x]. +Proof. +rewrite predeqE -!filter_from_norm_locally => A; split; last first. + move=> [B [e egt0 xe_B] <-]. + exists e => // y xe_y; exists (- y); last by rewrite opprK. + by apply/xe_B; rewrite /ball opprK -normmN -mulN1r mulrDr !mulN1r. +move=> [e egt0 oppxe_A]; exists [set - y | y in A]; last first. + rewrite predeqE => y; split=> [[z [t At <- <-]] |Ay]; first by rewrite opprK. + by exists (- y); [exists y|rewrite opprK]. +exists e => // y xe_y; exists (- y); last by rewrite opprK. +by apply/oppxe_A; rewrite /ball normmB opprK addrC. +Qed. + +Lemma openN (A : set R) : open A -> open [set - x | x in A]. +Proof. +move=> Aop; rewrite openE => _ [x /Aop x_A <-]. +by rewrite /interior locallyN; exists A. +Qed. + +Lemma closedN (A : set R) : closed A -> closed [set - x | x in A]. +Proof. +move=> Acl x clNAx. +suff /Acl : closure A (- x) by exists (- x)=> //; rewrite opprK. +move=> B oppx_B; have : [set - x | x in A] `&` [set - x | x in B] !=set0. + by apply: clNAx; rewrite -[x]opprK locallyN; exists B. +move=> [y [[z Az oppzey] [t Bt opptey]]]; exists (- y). +by split; [rewrite -oppzey opprK|rewrite -opptey opprK]. +Qed. + Lemma segment_connected (a b : R) : connected [set x | x \in `[a, b]]. Proof. move=> A [y Ay] Aop Acl. @@ -2307,18 +1386,20 @@ have Az : A z. rewrite AeabB; split. suff : {subset `[y, x] <= `[a, b]} by apply. by apply/subitvP; rewrite /= (itvP abx); have /sAab/itvP-> := Ay. - apply: Bcl => D [_ /posnumP[e] ze_D]. - have [t] := sup_adherent Altxsup [gt0 of e%:num]. + apply: Bcl => D; rewrite -filter_from_norm_locally => -[_/posnumP[e] ze_D]. + have [t] := sup_adherent Altxsup (posnum_gt0 e). rewrite in_setE => - [At lttx] ltzet. exists t; split; first by move: At; rewrite AeabB => - []. - apply/ze_D; rewrite /AbsRing_ball /= absRE ltr_distl. + apply/ze_D; rewrite /ball /= ltr_distl. apply/andP; split; last by rewrite -ltr_subl_addr. rewrite ltr_subl_addr; apply: ltr_spaddr => //. by apply/sup_upper_bound => //; rewrite in_setE. have ltzx : 0 < x - z. have : z <= x by rewrite (itvP yxz). by rewrite subr_gt0 ler_eqVlt => /orP [/eqP zex|] //; move: nAx; rewrite -zex. -have := Az; rewrite AeabC => - [_ /Cop [_ /posnumP[e] ze_C]]. +have := Az; rewrite AeabC => -[_ /Cop]. +rewrite entourages_normE locally_entourageE -filter_from_norm_locally. +move => [_ /posnumP[e] ze_C]. suff [t Altxt] : exists2 t, Altx t & z < t. by rewrite ltrNge => /negP; apply; apply/sup_upper_bound. exists (z + (minr (e%:num / 2) ((PosNum ltzx)%:num / 2))); last first. @@ -2327,7 +1408,7 @@ rewrite in_setE; split; last first. rewrite -[(< _) _]ltr_subr_addl ltr_minl; apply/orP; right. by rewrite ltr_pdivr_mulr // mulrDr mulr1 ltr_addl. rewrite AeabC; split; last first. - apply: ze_C; rewrite /AbsRing_ball /ball_ absRE ltr_distl. + apply: ze_C; rewrite /ball ltr_distl. apply/andP; split; last by rewrite -addrA ltr_addl. rewrite -addrA gtr_addl subr_lt0 ltr_minl; apply/orP; left. by rewrite [X in _ < X]splitr ltr_addl. @@ -2366,36 +1447,38 @@ apply: segment_connected. by rewrite eqr_le !(itvP aex). - exists B => //; rewrite openE => x [D' sD [saxUf [i Di fx]]]. have : open (f i) by have /sD := Di; rewrite in_setE => /fop. - rewrite openE => /(_ _ fx) [e egt0 xe_fi]; exists e => // y xe_y. + rewrite openE => /(_ _ fx); rewrite !/(_^°) -!filter_from_norm_locally. + move=> [e egt0 xe_fi]; exists e => // y xe_y. exists D' => //; split; last by exists i => //; apply/xe_fi. move=> z ayz; case: (lerP z x) => [lezx|ltxz]. by apply/saxUf; rewrite inE (itvP ayz) lezx. - exists i=> //; apply/xe_fi; rewrite /AbsRing_ball/ball_ absrB absRE ger0_norm. + exists i=> //; apply/xe_fi; rewrite /ball normmB [ `|[_]|]ger0_norm. have lezy : z <= y by rewrite (itvP ayz). rewrite ltr_subl_addl; apply: ler_lt_trans lezy _; rewrite -ltr_subl_addr. - by have := xe_y; rewrite /AbsRing_ball/ball_ absRE => /ltr_distW. + by have := xe_y; rewrite /ball => /ltr_distW. by rewrite subr_ge0; apply/ltrW. exists A; last by rewrite predeqE => x; split=> [[] | []]. move=> x clAx; have abx : x \in `[a, b]. by apply: segment_closed; have /closureI [] := clAx. split=> //; have /sabUf [i Di fx] := abx. -have /fop := Di; rewrite openE => /(_ _ fx) [_ /posnumP[e] xe_fi]. +have /fop := Di; rewrite openE => /(_ _ fx). +rewrite /(_^°) -filter_from_norm_locally => - [_ /posnumP[e] xe_fi]. have /clAx [y [[aby [D' sD [sayUf _]]] xe_y]] := locally_ball x e. exists (i |` D')%fset; first by move=> j /fset1UP[->|/sD] //; rewrite in_setE. split=> [z axz|]; last first. exists i; first by rewrite !inE eq_refl. - exact/xe_fi/(@ball_center [uniformType of R]). + exact/xe_fi/ball_center. case: (lerP z y) => [lezy|ltyz]. have /sayUf [j Dj fjz] : z \in `[a, y] by rewrite inE (itvP axz) lezy. by exists j => //; rewrite inE orbC Dj. exists i; first by rewrite !inE eq_refl. -apply/xe_fi; rewrite /AbsRing_ball/ball_ absRE ger0_norm; last first. +apply/xe_fi; rewrite /ball [ `|[_]|]ger0_norm; last first. by rewrite subr_ge0 (itvP axz). rewrite ltr_subl_addl -ltr_subl_addr; apply: ltr_trans ltyz. -by apply: ltr_distW; rewrite -absRE absrB. +by apply: ltr_distW; rewrite distrC. Qed. -Lemma ler0_addgt0P (R : realFieldType) (x : R) : +Lemma ler0_addgt0P (K : realFieldType) (x : K) : reflect (forall e, e > 0 -> x <= e) (x <= 0). Proof. apply: (iffP idP) => [lex0 e egt0|lex0]. @@ -2415,7 +1498,7 @@ move=> leab; wlog : f v / f a <= f b. exact: ivt. move=> fcont fabv; have [] := ivt (fun x => - f x) (- v). - by rewrite ler_oppr opprK. - - by move=> x /fcont; apply: (@continuousN _ [normedModType R of R^o]). + - by move=> x /fcont; apply: continuousN. - by rewrite -oppr_max -oppr_min ler_oppr opprK ler_oppr opprK andbC. by move=> c cab /eqP; rewrite eqr_opp => /eqP; exists c. move=> lefab fcont; rewrite minr_l // maxr_r // => /andP []. @@ -2434,30 +1517,37 @@ have supAab : sup A \in `[a, b]. exists (sup A) => //; have lefsupv : f (sup A) <= v. rewrite lerNgt; apply/negP => ltvfsup. have vltfsup : 0 < f (sup A) - v by rewrite subr_gt0. - have /fcont /(_ _ (locally_ball _ (PosNum vltfsup))) [_/posnumP[d] supdfe] - := supAab. - have [t At supd_t] := sup_adherent supA [gt0 of d%:num]. - suff /supdfe : ball (sup A) d%:num t. - rewrite ball_absE /= absRE ltr_norml => /andP [_]. + have /fcont /(_ _ (locally_ball _ (PosNum vltfsup))) := supAab. + rewrite !near_simpl -locally_nearE -filter_from_norm_locally. + move=> [_/posnumP[d] supdfe]. + have [t At supd_t] := sup_adherent supA (posnum_gt0 d). + suff /supdfe : ball norm (sup A) d%:num t. + rewrite /ball ltr_norml => /andP [_]. by rewrite ltr_add2l ltr_oppr opprK ltrNge; have /andP [_ ->] := At. - rewrite ball_absE /= absRE ger0_norm. + rewrite /ball [ `|[_]|]ger0_norm. by rewrite ltr_subl_addr -ltr_subl_addl. by rewrite subr_ge0 sup_upper_bound. apply/eqP; rewrite eqr_le; apply/andP; split=> //. rewrite -subr_le0; apply/ler0_addgt0P => _/posnumP[e]. rewrite ler_subl_addr -ler_subl_addl ltrW //. -have /fcont /(_ _ (locally_ball _ e)) [_/posnumP[d] supdfe] := supAab. +have /fcont /(_ _ (locally_ball _ e)) := supAab. +rewrite !near_simpl -locally_nearE -filter_from_norm_locally. +move=> [_/posnumP[d] supdfe]. have atrF := at_right_proper_filter (sup A); near (at_right (sup A)) => x. -have /supdfe /= : ball (sup A) d%:num x. - by near: x; rewrite /= locally_simpl; exists d%:num => //. -rewrite ball_absE /= absRE => /ltr_distW; apply: ler_lt_trans. +have /supdfe /= : ball norm (sup A) d%:num x. + near: x; rewrite /at_right !near_simpl -locally_nearE. + by rewrite -filter_from_norm_locally; exists d%:num. +move/ltr_distW; apply: ler_lt_trans. rewrite ler_add2r ltrW //; suff : forall t, t \in `](sup A), b] -> v < f t. - apply; rewrite inE; apply/andP; split; first by near: x; exists 1. - near: x; exists (b - sup A). + apply; rewrite inE; apply/andP; split. + near: x; rewrite /at_right !near_simpl -locally_nearE. + by rewrite -filter_from_norm_locally; exists 1. + near: x; rewrite /at_right !near_simpl -locally_nearE. + rewrite -filter_from_norm_locally; exists (b - sup A). rewrite subr_gt0 ltr_def (itvP supAab) andbT; apply/negP => /eqP besup. by move: lefsupv; rewrite lerNgt -besup ltvfb. - move=> t lttb ltsupt; move: lttb; rewrite /AbsRing_ball /= absrB absRE. - by rewrite gtr0_norm ?subr_gt0 // ltr_add2r; apply: ltrW. + move=> t lttb ltsupt; move: lttb; rewrite /ball normmB. + by rewrite [ `|[_]|]gtr0_norm ?subr_gt0 // ltr_add2r; apply: ltrW. move=> t /andP [ltsupt letb]; rewrite ltrNge; apply/negP => leftv. move: ltsupt; rewrite ltrNge => /negP; apply; apply: sup_upper_bound => //. by rewrite inE letb leftv. @@ -2465,15 +1555,20 @@ Grab Existential Variables. all: end_near. Qed. (** Local properties in [R] *) -Lemma locally_interval (P : R -> Prop) (x : R) (a b : Rbar) : - Rbar_lt a x -> Rbar_lt x b -> - (forall (y : R), Rbar_lt a y -> Rbar_lt y b -> P y) -> - locally x P. -Proof. -move => Hax Hxb Hp; case: (Rbar_lt_locally _ _ _ Hax Hxb) => d Hd. -exists d%:num => //= y; rewrite /AbsRing_ball /= absrB. -by move=> /Hd /andP[??]; apply: Hp. -Qed. +(* Lemma locally_interval (P : R -> Prop) (x : R) (a b : Rbar) : *) +(* Rbar_lt a x -> Rbar_lt x b -> *) +(* (forall (y : R), Rbar_lt a y -> Rbar_lt y b -> P y) -> *) +(* locally x P. *) +(* Proof. *) +(* move => Hax Hxb Hp; case: (Rbar_lt_locally _ _ _ Hax Hxb) => d Hd. *) +(* exists d%:num => //= y; rewrite /AbsRing_ball /= absrB. *) +(* by move=> /Hd /andP[??]; apply: Hp. *) +(* Qed. *) + +End real_topology. + +Hint Resolve open_lt. +Hint Resolve open_gt. (** * Topology on [R]² *) @@ -2623,10 +1718,10 @@ Qed. (* by apply (Hd (u, v)) => /=; split; apply sub_abs_ball; rewrite absrB. *) (* Qed. *) -Definition bounded (K : absRingType) (V : normedModType K) (A : set V) := +Definition bounded (K : realFieldType) (V : normedModType K) (A : set V) := \forall M \near +oo, A `<=` [set x | `|[x]| < M]. -Lemma compact_bounded (K : absRingType) (V : normedModType K) (A : set V) : +Lemma compact_bounded (K : realType) (V : normedModType K) (A : set V) : compact A -> bounded A. Proof. rewrite compact_cover => Aco. @@ -2635,8 +1730,9 @@ have covA : A `<=` \bigcup_(n : int) [set p | `|[p]| < n%:~R]. by rewrite rmorphD /= -floorE floorS_gtr. have /Aco [] := covA. move=> n _; rewrite openE => p; rewrite -subr_gt0 => ltpn. - apply/locallyP; exists (n%:~R - `|[p]|) => // q. - rewrite -ball_normE /= ltr_subr_addr normmB; apply: ler_lt_trans. + apply/locallyP; rewrite locally_entourageE -filter_from_norm_locally. + exists (n%:~R - `|[p]|) => // q. + rewrite /ball ltr_subr_addr normmB; apply: ler_lt_trans. by rewrite -{1}(subrK p q) ler_normm_add. move=> D _ DcovA. exists (bigmaxr 0 [seq n%:~R | n <- enum_fset D]). @@ -2703,7 +1799,7 @@ have GC : G [set g | C (\row_j g j)] by exists C. by have [g []] := clGf _ _ GC f_D; exists (\row_j (g j : T)). Qed. -Lemma bounded_closed_compact n (A : set 'rV[R]_n.+1) : +Lemma bounded_closed_compact (R : realType) n (A : set 'rV[R]_n.+1) : bounded A -> closed A -> compact A. Proof. move=> [M normAltM] Acl. @@ -2712,7 +1808,7 @@ have Mnco : compact apply: (@rV_compact _ _ (fun _ => [set x | x \in `[(- (M + 1)), (M + 1)]])). by move=> _; apply: segment_compact. apply: subclosed_compact Acl Mnco _ => v /normAltM normvltM i. -suff /ltrW : `|[v ord0 i : R^o]| < M + 1 by rewrite [ `|[_]| ]absRE ler_norml. +suff /ltrW : `|[v ord0 i]| < M + 1 by rewrite ler_norml. apply: ler_lt_trans (normvltM _ _); last by rewrite ltr_addl. have vinseq : `|v ord0 i| \in [seq `|v ij.1 ij.2| | ij : 'I_1 * 'I_n.+1]. by apply/mapP; exists (ord0, i) => //=; rewrite mem_enum. @@ -2722,121 +1818,121 @@ Qed. (** Open sets in [Rbar] *) -Lemma open_Rbar_lt y : open (fun u : R => Rbar_lt u y). -Proof. -case: y => [y||] /=. -exact: open_lt. -by rewrite trueE; apply: openT. -by rewrite falseE; apply: open0. -Qed. +(* Lemma open_Rbar_lt y : open (fun u : R => Rbar_lt u y). *) +(* Proof. *) +(* case: y => [y||] /=. *) +(* exact: open_lt. *) +(* by rewrite trueE; apply: openT. *) +(* by rewrite falseE; apply: open0. *) +(* Qed. *) -Lemma open_Rbar_gt y : open (fun u : R => Rbar_lt y u). -Proof. -case: y => [y||] /=. -exact: open_gt. -by rewrite falseE; apply: open0. -by rewrite trueE; apply: openT. -Qed. +(* Lemma open_Rbar_gt y : open (fun u : R => Rbar_lt y u). *) +(* Proof. *) +(* case: y => [y||] /=. *) +(* exact: open_gt. *) +(* by rewrite falseE; apply: open0. *) +(* by rewrite trueE; apply: openT. *) +(* Qed. *) -Lemma open_Rbar_lt' x y : Rbar_lt x y -> Rbar_locally x (fun u => Rbar_lt u y). -Proof. -case: x => [x|//|] xy; first exact: open_Rbar_lt. -case: y => [y||//] /= in xy *. -exists y => /= x ? //. -by exists 0. -Qed. +(* Lemma open_Rbar_lt' x y : Rbar_lt x y -> Rbar_locally x (fun u => Rbar_lt u y). *) +(* Proof. *) +(* case: x => [x|//|] xy; first exact: open_Rbar_lt. *) +(* case: y => [y||//] /= in xy *. *) +(* exists y => /= x ? //. *) +(* by exists 0. *) +(* Qed. *) -Lemma open_Rbar_gt' x y : Rbar_lt y x -> Rbar_locally x (fun u => Rbar_lt y u). -Proof. -case: x => [x||] //=; do ?[exact: open_Rbar_gt]; - case: y => [y||] //=; do ?by exists 0. -by exists y => x yx //=. -Qed. +(* Lemma open_Rbar_gt' x y : Rbar_lt y x -> Rbar_locally x (fun u => Rbar_lt y u). *) +(* Proof. *) +(* case: x => [x||] //=; do ?[exact: open_Rbar_gt]; *) +(* case: y => [y||] //=; do ?by exists 0. *) +(* by exists y => x yx //=. *) +(* Qed. *) -Lemma Rbar_locally'_le x : Rbar_locally' x --> Rbar_locally x. -Proof. -move: x => [x P [_/posnumP[e] HP] |x P|x P] //=. -by exists e%:num => // ???; apply: HP. -Qed. +(* Lemma Rbar_locally'_le x : Rbar_locally' x --> Rbar_locally x. *) +(* Proof. *) +(* move: x => [x P [_/posnumP[e] HP] |x P|x P] //=. *) +(* by exists e%:num => // ???; apply: HP. *) +(* Qed. *) -Lemma Rbar_locally'_le_finite (x : R) : Rbar_locally' x --> locally x. -Proof. -by move=> P [_/posnumP[e] HP] //=; exists e%:num => // ???; apply: HP. -Qed. +(* Lemma Rbar_locally'_le_finite (x : R) : Rbar_locally' x --> locally x. *) +(* Proof. *) +(* by move=> P [_/posnumP[e] HP] //=; exists e%:num => // ???; apply: HP. *) +(* Qed. *) (** * Some limits on real functions *) -Definition Rbar_loc_seq (x : Rbar) (n : nat) := match x with - | Finite x => x + (INR n + 1)^-1 - | +oo => INR n - | -oo => - INR n - end. +(* Definition Rbar_loc_seq (x : Rbar) (n : nat) := match x with *) +(* | Finite x => x + (INR n + 1)^-1 *) +(* | +oo => INR n *) +(* | -oo => - INR n *) +(* end. *) -Lemma flim_Rbar_loc_seq x : Rbar_loc_seq x --> Rbar_locally' x. -Proof. -move=> P; rewrite /Rbar_loc_seq. -case: x => /= [x [_/posnumP[delta] Hp] |[delta Hp] |[delta Hp]]; last 2 first. - have /ZnatP [N Nfloor] : ifloor (maxr delta 0) \is a Znat. - by rewrite Znat_def ifloor_ge0 ler_maxr lerr orbC. - exists N.+1 => // n ltNn; apply: Hp. - have /ler_lt_trans : delta <= maxr delta 0 by rewrite ler_maxr lerr. - apply; apply: ltr_le_trans (floorS_gtr _) _; rewrite floorE Nfloor. - by rewrite -(@natrD [ringType of R] N 1) INRE ler_nat addn1. - have /ZnatP [N Nfloor] : ifloor (maxr (- delta) 0) \is a Znat. - by rewrite Znat_def ifloor_ge0 ler_maxr lerr orbC. - exists N.+1 => // n ltNn; apply: Hp; rewrite ltr_oppl. - have /ler_lt_trans : - delta <= maxr (- delta) 0 by rewrite ler_maxr lerr. - apply; apply: ltr_le_trans (floorS_gtr _) _; rewrite floorE Nfloor. - by rewrite -(@natrD [ringType of R] N 1) INRE ler_nat addn1. -have /ZnatP [N Nfloor] : ifloor (delta%:num^-1) \is a Znat. - by rewrite Znat_def ifloor_ge0. -exists N => // n leNn; have gt0Sn : 0 < INR n + 1. - by apply: ltr_spaddr => //; apply/RleP/pos_INR. -apply: Hp; last first. - by rewrite eq_sym addrC -subr_eq subrr eq_sym; apply/invr_neq0/lt0r_neq0. -rewrite /AbsRing_ball /= opprD addrA subrr absrB subr0. -rewrite absRE gtr0_norm; last by rewrite invr_gt0. -rewrite -[X in X < _]mulr1 ltr_pdivr_mull // -ltr_pdivr_mulr // div1r. -apply: ltr_le_trans (floorS_gtr _) _; rewrite floorE Nfloor ler_add //. -by rewrite INRE ler_nat. -Qed. +(* Lemma flim_Rbar_loc_seq x : Rbar_loc_seq x --> Rbar_locally' x. *) +(* Proof. *) +(* move=> P; rewrite /Rbar_loc_seq. *) +(* case: x => /= [x [_/posnumP[delta] Hp] |[delta Hp] |[delta Hp]]; last 2 first. *) +(* have /ZnatP [N Nfloor] : ifloor (maxr delta 0) \is a Znat. *) +(* by rewrite Znat_def ifloor_ge0 ler_maxr lerr orbC. *) +(* exists N.+1 => // n ltNn; apply: Hp. *) +(* have /ler_lt_trans : delta <= maxr delta 0 by rewrite ler_maxr lerr. *) +(* apply; apply: ltr_le_trans (floorS_gtr _) _; rewrite floorE Nfloor. *) +(* by rewrite -(@natrD [ringType of R] N 1) INRE ler_nat addn1. *) +(* have /ZnatP [N Nfloor] : ifloor (maxr (- delta) 0) \is a Znat. *) +(* by rewrite Znat_def ifloor_ge0 ler_maxr lerr orbC. *) +(* exists N.+1 => // n ltNn; apply: Hp; rewrite ltr_oppl. *) +(* have /ler_lt_trans : - delta <= maxr (- delta) 0 by rewrite ler_maxr lerr. *) +(* apply; apply: ltr_le_trans (floorS_gtr _) _; rewrite floorE Nfloor. *) +(* by rewrite -(@natrD [ringType of R] N 1) INRE ler_nat addn1. *) +(* have /ZnatP [N Nfloor] : ifloor (delta%:num^-1) \is a Znat. *) +(* by rewrite Znat_def ifloor_ge0. *) +(* exists N => // n leNn; have gt0Sn : 0 < INR n + 1. *) +(* by apply: ltr_spaddr => //; apply/RleP/pos_INR. *) +(* apply: Hp; last first. *) +(* by rewrite eq_sym addrC -subr_eq subrr eq_sym; apply/invr_neq0/lt0r_neq0. *) +(* rewrite /AbsRing_ball /= opprD addrA subrr absrB subr0. *) +(* rewrite absRE gtr0_norm; last by rewrite invr_gt0. *) +(* rewrite -[X in X < _]mulr1 ltr_pdivr_mull // -ltr_pdivr_mulr // div1r. *) +(* apply: ltr_le_trans (floorS_gtr _) _; rewrite floorE Nfloor ler_add //. *) +(* by rewrite INRE ler_nat. *) +(* Qed. *) (* TODO: express using ball?*) -Lemma continuity_pt_locally f x : continuity_pt f x <-> - forall eps : posreal, locally x (fun u => `|f u - f x| < eps). -Proof. -split=> [fcont e|fcont _/RltP/posnumP[e]]; last first. - have [_/posnumP[d] xd_fxe] := fcont e. - exists d%:num; split; first by apply/RltP; have := [gt0 of d%:num]. - by move=> y [_ /RltP yxd]; apply/RltP/xd_fxe; rewrite /AbsRing_ball /= absrB. -have /RltP egt0 := [gt0 of e%:num]. -have [_ [/RltP/posnumP[d] dx_fxe]] := fcont e%:num egt0. -exists d%:num => // y xyd; case: (eqVneq x y) => [->|xney]. - by rewrite subrr absr0. -apply/RltP/dx_fxe; split; first by split=> //; apply/eqP. -by have /RltP := xyd; rewrite absrB. -Qed. +(* Lemma continuity_pt_locally f x : continuity_pt f x <-> *) +(* forall eps : posreal, locally x (fun u => `|f u - f x| < eps). *) +(* Proof. *) +(* split=> [fcont e|fcont _/RltP/posnumP[e]]; last first. *) +(* have [_/posnumP[d] xd_fxe] := fcont e. *) +(* exists d%:num; split; first by apply/RltP; have := [gt0 of d%:num]. *) +(* by move=> y [_ /RltP yxd]; apply/RltP/xd_fxe; rewrite /AbsRing_ball /= absrB. *) +(* have /RltP egt0 := [gt0 of e%:num]. *) +(* have [_ [/RltP/posnumP[d] dx_fxe]] := fcont e%:num egt0. *) +(* exists d%:num => // y xyd; case: (eqVneq x y) => [->|xney]. *) +(* by rewrite subrr absr0. *) +(* apply/RltP/dx_fxe; split; first by split=> //; apply/eqP. *) +(* by have /RltP := xyd; rewrite absrB. *) +(* Qed. *) -Lemma continuity_pt_flim (f : R -> R) (x : R) : - continuity_pt f x <-> {for x, continuous f}. -Proof. -eapply iff_trans; first exact: continuity_pt_locally. -apply iff_sym. -have FF : Filter (f @ x) by typeclasses eauto. -case: (@flim_ballP _ (f @ x) FF (f x)) => {FF}H1 H2. -(* TODO: in need for lemmas and/or refactoring of already existing lemmas (ball vs. Rabs) *) -split => [{H2} /H1{H1} H1 eps|{H1} H]. -- have {H1} [//|_/posnumP[x0] Hx0] := H1 eps%:num. - exists x0%:num => // Hx0' /Hx0 /=. - by rewrite ball_absE /= absrB; apply. -- apply H2 => _ /posnumP[eps]; move: (H eps) => {H} [_ /posnumP[x0] Hx0]. - exists x0%:num => // y /Hx0 /= {Hx0}Hx0. - by rewrite ball_absE /= absrB. -Qed. +(* Lemma continuity_pt_flim (f : R -> R) (x : R) : *) +(* continuity_pt f x <-> {for x, continuous f}. *) +(* Proof. *) +(* eapply iff_trans; first exact: continuity_pt_locally. *) +(* apply iff_sym. *) +(* have FF : Filter (f @ x) by typeclasses eauto. *) +(* case: (@flim_ballP _ (f @ x) FF (f x)) => {FF}H1 H2. *) +(* (* TODO: in need for lemmas and/or refactoring of already existing lemmas (ball vs. Rabs) *) *) +(* split => [{H2} /H1{H1} H1 eps|{H1} H]. *) +(* - have {H1} [//|_/posnumP[x0] Hx0] := H1 eps%:num. *) +(* exists x0%:num => // Hx0' /Hx0 /=. *) +(* by rewrite ball_absE /= absrB; apply. *) +(* - apply H2 => _ /posnumP[eps]; move: (H eps) => {H} [_ /posnumP[x0] Hx0]. *) +(* exists x0%:num => // y /Hx0 /= {Hx0}Hx0. *) +(* by rewrite ball_absE /= absrB. *) +(* Qed. *) -Lemma continuity_ptE (f : R -> R) (x : R) : - continuity_pt f x <-> {for x, continuous f}. -Proof. exact: continuity_pt_flim. Qed. +(* Lemma continuity_ptE (f : R -> R) (x : R) : *) +(* continuity_pt f x <-> {for x, continuous f}. *) +(* Proof. exact: continuity_pt_flim. Qed. *) Lemma continuous_withinNx {U V : uniformType} (f : U -> V) x : @@ -2852,17 +1948,17 @@ rewrite !near_simpl near_withinE near_simpl => Pf; near=> y. by have [->|] := eqVneq y x; [by apply: locally_singleton|near: y]. Grab Existential Variables. all: end_near. Qed. -Lemma continuity_pt_flim' f x : - continuity_pt f x <-> f @ locally' x --> f x. -Proof. by rewrite continuity_ptE continuous_withinNx. Qed. +(* Lemma continuity_pt_flim' f x : *) +(* continuity_pt f x <-> f @ locally' x --> f x. *) +(* Proof. by rewrite continuity_ptE continuous_withinNx. Qed. *) -Lemma continuity_pt_locally' f x : - continuity_pt f x <-> - forall eps : R, 0 < eps -> locally' x (fun u => `|f x - f u| < eps)%R. -Proof. -by rewrite continuity_pt_flim' (@flim_normP _ [normedModType R of R^o]). -Qed. +(* Lemma continuity_pt_locally' f x : *) +(* continuity_pt f x <-> *) +(* forall eps : R, 0 < eps -> locally' x (fun u => `|f x - f u| < eps)%R. *) +(* Proof. *) +(* by rewrite continuity_pt_flim' (@flim_normP _ [normedModType R of R^o]). *) +(* Qed. *) -Lemma locally_pt_comp (P : R -> Prop) (f : R -> R) (x : R) : - locally (f x) P -> continuity_pt f x -> \near x, P (f x). -Proof. by move=> Lf /continuity_pt_flim; apply. Qed. +(* Lemma locally_pt_comp (P : R -> Prop) (f : R -> R) (x : R) : *) +(* locally (f x) P -> continuity_pt f x -> \near x, P (f x). *) +(* Proof. by move=> Lf /continuity_pt_flim; apply. Qed. *) diff --git a/landau.v b/landau.v index dce736dcb9..197a3bc08b 100644 --- a/landau.v +++ b/landau.v @@ -1,9 +1,7 @@ (* mathcomp analysis (c) 2017 Inria and AIST. License: CeCILL-C. *) -Require Import Reals. From Coq Require Import ssreflect ssrfun ssrbool. From mathcomp Require Import ssrnat eqtype choice fintype bigop ssralg ssrnum. -Require Import boolp reals Rstruct Rbar. -Require Import classical_sets posnum topology hierarchy. +Require Import boolp reals classical_sets posnum topology hierarchy. (******************************************************************************) (* BACHMANN-LANDAU NOTATIONS : BIG AND LITTLE O *) @@ -270,12 +268,11 @@ Reserved Notation "f '=Theta_' F h" F at level 0, h at next level, format "f '=Theta_' F h"). -Delimit Scope R_scope with coqR. Delimit Scope real_scope with real. -Local Close Scope R_scope. +Local Open Scope classical_set_scope. Local Open Scope ring_scope. Local Open Scope real_scope. -Local Open Scope classical_set_scope. +Local Open Scope fun_scope. Section function_space. @@ -346,10 +343,10 @@ Lemma showo : (gen_tag = tt) * (the_tag = tt) * (a_tag = tt). Proof. by []. Qed. (* Tentative to handle small o and big O notations *) Section Domination. -Context {K : absRingType} {T : Type} {V W : normedModType K}. +Context {K : realFieldType} {T : Type} {V W : normedModType K}. Let littleo_def (F : set (set T)) (f : T -> V) (g : T -> W) := - forall eps : R, 0 < eps -> \forall x \near F, `|[f x]| <= eps * `|[g x]|. + forall eps, 0 < eps -> \forall x \near F, `|[f x]| <= eps * `|[g x]|. Structure littleo_type (F : set (set T)) (g : T -> W) := Littleo { littleo_fun :> T -> V; @@ -535,7 +532,7 @@ Lemma scale_littleo_subproof (F : filter_on T) e (df : {o_F e}) a : littleo_def F (a *: (df : _ -> _)) e. Proof. have [->|a0] := eqVneq a 0; first by rewrite scale0r. -move=> _ /posnumP[eps]; have aa := absr_eq0 a; near=> x => /=. +move=> _ /posnumP[eps]; have aa := normr_eq0 a; near=> x => /=. rewrite normmZ -ler_pdivl_mull ?ltr_def ?aa ?a0 //= mulrA; near: x. by apply: littleoP; rewrite mulr_gt0 // invr_gt0 ?ltr_def ?aa ?a0 /=. Grab Existential Variables. all: end_near. Qed. @@ -652,7 +649,7 @@ Canonical the_bigO_bigO (tag : unit) (F : filter_on T) (phF : phantom (set (set T)) F) f h := [bigO of the_bigO tag F phF f h]. Variant bigO_spec (F : set (set T)) (g : T -> W) : (T -> V) -> Prop := - BigOSpec f (k : posreal) + BigOSpec f (k : {posnum K}) of (\forall x \near F, `|[f x]| <= k%:num * `|[g x]|) : bigO_spec F g f. @@ -885,7 +882,7 @@ Hint Resolve littleo_eqO. Arguments bigO {_ _ _ _}. (* NB: see also scaleox *) -Lemma scaleolx (K : absRingType) (V W : normedModType K) {T : Type} +Lemma scaleolx (K : realFieldType) (V W : normedModType K) {T : Type} (F : filter_on T) (a : W) (k : T -> K^o) (e : T -> V) (x : T) : ([o_F e of k] x) *: a = [o_F e of (fun y => [o_F e of k] y *: a)] x. Proof. @@ -901,7 +898,7 @@ Grab Existential Variables. all: end_near. Qed. Section Limit. -Context {K : absRingType} {T : Type} {V W X : normedModType K}. +Context {K : realFieldType} {T : Type} {V W X : normedModType K}. Lemma eqolimP (F : filter_on T) (f : T -> V) (l : V) : f @ F --> l <-> f = cst l +o_F (cst (1 : K^o)). @@ -909,10 +906,10 @@ Proof. split=> fFl. apply/eqaddoP => _/posnumP[eps]; near=> x. rewrite /cst ltrW //= normmB; near: x. - by apply: (flim_norm _ fFl); rewrite mulr_gt0 // ?absr1_gt0. + by apply: (flim_norm _ fFl); rewrite mulr_gt0 // [ `|[_]|]normr1. apply/flim_normP=> _/posnumP[eps]; rewrite /= near_simpl. have lt_eps x : x <= (eps%:num / 2%:R) * `|1 : K^o|%real -> x < eps. - rewrite absr1 mulr1 => /ler_lt_trans; apply. + rewrite normr1 mulr1 => /ler_lt_trans; apply. by rewrite ltr_pdivr_mulr // ltr_pmulr // ltr1n. near=> x; rewrite [X in X x]fFl opprD addNKr normmN lt_eps //; near: x. by rewrite /= !near_simpl; apply: littleoP; rewrite divr_gt0. @@ -1003,7 +1000,7 @@ Arguments bigO_bigO_eqO {K T V W X F}. Section littleo_bigO_transitivity. -Context {K : absRingType} {T : Type} {V W Z : normedModType K}. +Context {K : realFieldType} {T : Type} {V W Z : normedModType K}. Lemma eqaddo_trans (F : filter_on T) (f g h : T -> V) fg gh (e : T -> W): f = g + [o_ F e of fg] -> g = h + [o_F e of gh] -> f = h +o_F e. @@ -1041,13 +1038,13 @@ End littleo_bigO_transitivity. Section rule_of_products_in_R. -Variable pT : pointedType. +Variable (pT : pointedType) (R : rcfType). -Lemma mulo (F : filter_on pT) (h1 h2 f g : pT -> R^o) : +Lemma mulo (F : filter_on pT) (h1 h2 f g : pT -> R) : [o_F h1 of f] * [o_F h2 of g] =o_F (h1 * h2). Proof. rewrite [in RHS]littleoE // => _/posnumP[e]; near=> x. -rewrite [`|[_]|]absrM -(sqr_sqrtr (ltrW [gt0 of e%:num])) expr2. +rewrite [`|[_]|]normrM -(sqr_sqrtr (ltrW [gt0 of e%:num])) expr2. rewrite [`|[_]|]normrM mulrACA ler_pmul //; near: x; by have [/= h] := littleo; apply. Grab Existential Variables. all: end_near. Qed. @@ -1057,7 +1054,7 @@ Lemma mulO (F : filter_on pT) (h1 h2 f g : pT -> R^o) : Proof. rewrite [RHS]bigOE//; have [ O1 k1 Oh1] := bigO; have [ O2 k2 Oh2] := bigO. near=> k; move: Oh1 Oh2; apply: filter_app2; near=> x => leOh1 leOh2. -rewrite [`|[_]|]absrM (ler_trans (ler_pmul _ _ leOh1 leOh2)) //. +rewrite [`|[_]|]normrM // (ler_trans (ler_pmul _ _ leOh1 leOh2)) //. rewrite mulrACA [`|[_]| in X in _ <= X]normrM ler_wpmul2r // ?mulr_ge0 //. by near: k; apply: locally_pinfty_ge. Unshelve. end_near. Grab Existential Variables. end_near. Qed. @@ -1088,19 +1085,20 @@ End Shift. Arguments shift {R} x / y. Notation center c := (shift (- c)). -Lemma near_shift {K : absRingType} {R : normedModType K} +Lemma near_shift {K : realFieldType} {R : normedModType K} (y x : R) (P : set R) : (\near x, P x) = (\forall z \near y, (P \o shift (x - y)) z). Proof. -rewrite propeqE; split=> /= /locally_normP [_/posnumP[e] ye]; -apply/locally_normP; exists e%:num => // t /= et. +rewrite propeqE; split; + rewrite -!locally_nearE -!filter_from_norm_locally => -[_/posnumP[e] ye]; + exists e%:num => // t /= et. apply: ye; rewrite /= !opprD addrA addrACA subrr add0r. by rewrite opprK addrC. have /= := ye (t - (x - y)); rewrite addrNK; apply. by rewrite /= !opprB addrA addrCA subrr addr0. Qed. -Lemma flim_shift {T : Type} {K : absRingType} {R : normedModType K} +Lemma flim_shift {T : Type} {K : realFieldType} {R : normedModType K} (x y : R) (f : R -> T) : (f \o shift x) @ y = f @ (y + x). Proof. @@ -1109,8 +1107,8 @@ by rewrite (_ : _ \o _ = A \o f) // funeqE=> z; rewrite /= opprD addNKr addrNK. Qed. Section Linear3. -Context (U : normedModType R) (V : normedModType R) (s : R -> V -> V) - (s_law : GRing.Scale.law s). +Context (R : realFieldType) (U : normedModType R) (V : normedModType R) + (s : R -> V -> V) (s_law : GRing.Scale.law s). Hypothesis (normm_s : forall k x, `|[s k x]| = `|k| * `|[x]|). (* Split in multiple bits *) @@ -1127,13 +1125,13 @@ apply/flim_normP => _/posnumP[e]; rewrite !near_simpl. rewrite (near_shift 0) /= subr0; near=> y => /=. rewrite -linearB opprD addrC addrNK linearN normmN; near: y. suff flip : \forall k \near +oo, forall x, `|[f x]| <= k * `|[x]|. - near +oo => k; near=> y. + near +oo in R => k; near=> y. rewrite (ler_lt_trans (near flip k _ _)) // -ltr_pdivl_mull //. - near: y; apply/locally_normP. + near: y; rewrite -locally_nearE -filter_from_norm_locally. by eexists; last by move=> ?; rewrite /= sub0r normmN; apply. -have /locally_normP [_/posnumP[d]] := Of1. -rewrite /cst [X in _ * X]absr1 mulr1 => fk; near=> k => y. -case: (ler0P `|[y]|) => [|y0]. +move: Of1; rewrite near_simpl -locally_nearE -filter_from_norm_locally. +move=> [_/posnumP[d]]; rewrite /cst [X in _ * X]normr1 mulr1 => fk. +near=> k => y; case: (ler0P `|[y]|) => [|y0]. by rewrite normm_le0 => /eqP->; rewrite linear0 !normm0 mulr0. have ky0 : 0 <= k0%:num / (k * `|[y]|). by rewrite pmulr_rge0 // invr_ge0 mulr_ge0 // ltrW. @@ -1141,21 +1139,21 @@ rewrite -[X in _ <= X]mulr1 -ler_pdivr_mull ?pmulr_rgt0 //. rewrite -(ler_pmul2l [gt0 of k0%:num]) mulr1 mulrA -[_ / _]ger0_norm //. rewrite -normm_s. have <- : GRing.Scale.op s_law =2 s by rewrite GRing.Scale.opE. -rewrite -linearZ fk //= normmB subr0 normmZ absRE ger0_norm //. +rewrite -linearZ fk //= normmB subr0 normmZ ger0_norm //. rewrite invfM mulrA mulfVK ?lt0r_neq0 // ltr_pdivr_mulr //. by rewrite mulrC -ltr_pdivr_mulr //; near: k; apply: locally_pinfty_gt. Grab Existential Variables. all: end_near. Qed. End Linear3. -Arguments linear_for_continuous {U V s s_law normm_s} f _. +Arguments linear_for_continuous {R U V s s_law normm_s} f _. -Lemma linear_continuous (U : normedModType R) (V : normedModType R) - (f : {linear U -> V}) : +Lemma linear_continuous (R : realFieldType) (U : normedModType R) + (V : normedModType R) (f : {linear U -> V}) : (f : _ -> _) =O_ (0 : U) (cst (1 : R^o)) -> continuous f. Proof. by apply: linear_for_continuous => ??; rewrite normmZ. Qed. -Lemma linear_for_mul_continuous (U : normedModType R) +Lemma linear_for_mul_continuous (R : realFieldType) (U : normedModType R) (f : {linear U -> R | (@GRing.mul [ringType of R^o])}) : (f : _ -> _) =O_ (0 : U) (cst (1 : R^o)) -> continuous f. Proof. by apply: linear_for_continuous => ??; rewrite normmZ. Qed. @@ -1167,7 +1165,7 @@ Notation "f '~~_' F g" := (f == g +o_ F g) Section asymptotic_equivalence. -Context {K : absRingType} {T : Type} {V W : normedModType K}. +Context {K : realFieldType} {T : Type} {V W : normedModType K}. Implicit Types F : filter_on T. Lemma equivOLR F (f g : T -> V) : f ~_F g -> f =O_F g. @@ -1219,7 +1217,7 @@ End asymptotic_equivalence. Section big_omega. -Context {K : absRingType} {T : Type} {V : normedModType K}. +Context {K : realFieldType} {T : Type} {V : normedModType K}. Implicit Types W : normedModType K. Let bigOmega_def W (F : set (set T)) (f : T -> V) (g : T -> W) := @@ -1284,7 +1282,7 @@ Canonical the_bigOmega_bigOmega (tag : unit) (F : filter_on T) (phF : phantom (set (set T)) F) f h := [bigOmega of the_bigOmega tag F phF f h]. Variant bigOmega_spec {W} (F : set (set T)) (g : T -> W) : (T -> V) -> Prop := - BigOmegaSpec f (k : posreal) of + BigOmegaSpec f (k : {posnum K}) of (\forall x \near F, `|[f x]| >= k%:num * `|[g x]|) : bigOmega_spec F g f. @@ -1330,7 +1328,7 @@ Arguments bigOmega {_ _ _ _}. Section big_omega_in_R. -Variable pT : pointedType. +Variable (pT : pointedType) (R : realFieldType). Lemma addOmega (F : filter_on pT) (f g h : _ -> R^o) (f_nonneg : forall x, 0 <= f x) (g_nonneg : forall x, 0 <= g x) : @@ -1338,7 +1336,7 @@ Lemma addOmega (F : filter_on pT) (f g h : _ -> R^o) Proof. rewrite 2!eqOmegaE !eqOmegaO => /eqOP hOf; apply/eqOP. apply: filter_app hOf; near=> k; apply: filter_app; near=> x => /ler_trans. -by apply; rewrite ler_pmul2l // ![`|[_]|]absRE !ger0_norm // ?addr_ge0 // +by apply; rewrite ler_pmul2l // ![`|[_]|]ger0_norm // ?addr_ge0 // ler_addl. Unshelve. end_near. Grab Existential Variables. end_near. Qed. @@ -1347,7 +1345,7 @@ Lemma mulOmega (F : filter_on pT) (h1 h2 f g : pT -> R^o) : Proof. rewrite eqOmegaE eqOmegaO [in RHS]bigOE //. have [W1 k1 ?] := bigOmega; have [W2 k2 ?] := bigOmega. -near=> k; near=> x; rewrite [`|[_]|]absrM. +near=> k; near=> x; rewrite [`|[_]|]normrM. rewrite (@ler_trans _ ((k2%:num * k1%:num)^-1 * `|[(W1 * W2) x]|)) //. rewrite invrM ?unitfE ?gtr_eqF // -mulrA ler_pdivl_mull //. rewrite ler_pdivl_mull // (mulrA k1%:num) mulrCA [`|[_]|]normrM. @@ -1360,7 +1358,7 @@ End big_omega_in_R. Section big_theta. -Context {K : absRingType} {T : Type} {V : normedModType K}. +Context {K : realFieldType} {T : Type} {V : normedModType K}. Implicit Types W : normedModType K. Let bigTheta_def W (F : set (set T)) (f : T -> V) (g : T -> W) := @@ -1424,7 +1422,7 @@ Canonical the_bigTheta_bigTheta (tag : unit) (F : filter_on T) (phF : phantom (set (set T)) F) f h := [bigTheta of @the_bigTheta tag F phF f h]. Variant bigTheta_spec {W} (F : set (set T)) (g : T -> W) : (T -> V) -> Prop := - BigThetaSpec f (k1 : posreal) (k2 : posreal) of + BigThetaSpec f (k1 : {posnum K}) (k2 : {posnum K}) of (\forall x \near F, k1%:num * `|[g x]| <= `|[f x]|) & (\forall x \near F, `|[f x]| <= k2%:num * `|[g x]|) : bigTheta_spec F g f. @@ -1490,7 +1488,7 @@ Notation "f '=Theta_' F h" := (f%function = mkbigTheta the_tag F f h). Section big_theta_in_R. -Variable pT : pointedType. +Variable (pT : pointedType) (R : rcfType). Lemma addTheta (F : filter_on pT) (f g h : _ -> R^o) (f0 : forall x, 0 <= f x) (g0 : forall x, 0 <= g x) (h0 : forall x, 0 <= h x) : @@ -1512,8 +1510,8 @@ rewrite eqThetaE bigThetaE; split. by rewrite (eqThetaO _ f) (eqThetaO _ g) mulO. rewrite eqOmegaO [in RHS]bigOE //. have [T1 k1 l1 P1 ?] := bigTheta; have [T2 k2 l2 P2 ?] := bigTheta. -near=> k; first near=> x. -rewrite [`|[_]|]absrM (@ler_trans _ ((k2%:num * k1%:num)^-1 * `|[(T1 * T2) x]|)) //. +near=> k; first near=> x; rewrite [`|[_]|]normrM. +rewrite (@ler_trans _ ((k2%:num * k1%:num)^-1 * `|[(T1 * T2) x]|)) //. rewrite invrM ?unitfE ?gtr_eqF // -mulrA ler_pdivl_mull //. rewrite ler_pdivl_mull // (mulrA k1%:num) mulrCA [`|[_]|]normrM ler_pmul //; by [rewrite mulr_ge0 // ltrW|near: x]. diff --git a/misc/uniform_bigO.v b/misc/uniform_bigO.v index 967ace0175..abbc0dc1c5 100644 --- a/misc/uniform_bigO.v +++ b/misc/uniform_bigO.v @@ -2,7 +2,7 @@ Require Import Reals. From Coq Require Import ssreflect ssrfun ssrbool. From mathcomp Require Import ssrnat eqtype choice fintype bigop ssralg ssrnum. -Require Import boolp reals Rstruct Rbar. +Require Import boolp reals Rstruct. Require Import classical_sets posnum topology hierarchy landau. Set Implicit Arguments. @@ -42,7 +42,7 @@ Definition OuPex (f : A -> R * R -> R^o) (g : R * R -> R^o) := `|[f X dX]| <= C * `|[g dX]|. Lemma ler_norm2 (x : R^o * R^o) : - `|[x]| <= sqrt (Rsqr (fst x) + Rsqr (snd x)) <= Num.sqrt 2 * `|[x]|. + `|[x]| <= sqrt (Rsqr (fst x) + Rsqr (snd x))%R <= Num.sqrt 2 * `|[x]|. Proof. rewrite RsqrtE; last by rewrite addr_ge0 //; apply/RleP/Rle_0_sqr. rewrite !Rsqr_pow2 !RpowE; apply/andP; split. @@ -51,7 +51,7 @@ rewrite !Rsqr_pow2 !RpowE; apply/andP; split. wlog lex12 : x / (`|x.1| <= `|x.2|). move=> ler_norm; case: (lerP `|x.1| `|x.2|) => [/ler_norm|/ltrW lex21] //. by rewrite addrC [`|[_]|]maxrC (ler_norm (x.2, x.1)). -rewrite [`|[_]|]maxr_r // [`|[_]|]absRE -[X in X * _]ger0_norm // -normrM. +rewrite [`|[_]|]maxr_r // -[X in X * _]ger0_norm // -normrM. rewrite -sqrtr_sqr ler_wsqrtr // exprMn sqr_sqrtr // mulr_natl mulr2n ler_add2r. rewrite -[_ ^+ 2]ger0_norm ?sqr_ge0 // -[X in _ <=X]ger0_norm ?sqr_ge0 //. by rewrite !normrX ler_expn2r // nnegrE normr_ge0. @@ -76,6 +76,15 @@ Qed. (* then we replace the epsilon/delta definition with bigO *) +Canonical R_filteredType := + [filteredType R of R for realField_filteredType R_realFieldType]. +Canonical R_topologicalType := + [topologicalType of R for realField_topologicalType R_realFieldType]. +Canonical R_uniformType := + [uniformType of R for realField_uniformType R_realFieldType]. +Canonical R_normedModType := + [normedModType R of R for realField_normedModType R_realFieldType]. + Definition OuO (f : A -> R * R -> R^o) (g : R * R -> R^o) := (fun x => f x.1 x.2) =O_ (filter_prod [set setT] (within P [filter of 0 : R^o * R^o])) (fun x => g x.2). @@ -86,23 +95,19 @@ move=> /OuP_to_ex [_/posnumP[a] [_/posnumP[C] fOg]]. apply/eqOP; near=> k; near=> x; apply: ler_trans (fOg _ _ _ _) _; last 2 first. - by near: x; exists (setT, P); [split=> //=; apply: withinT|move=> ? []]. - by rewrite ler_pmul => //; near: k; exists C%:num => ? /ltrW. -- near: x; exists (setT, ball (0 : R^o * R^o) a%:num). +- near: x; exists (setT, ball norm (0 : R^o * R^o) a%:num). by split=> //=; rewrite /within; near=> x =>_; near: x; apply: locally_ball. - move=> x [_ [/=]]; rewrite -ball_normE /= normmB subr0 normmB subr0. - by move=> ??; rewrite ltr_maxl; apply/andP. + by move=> x [_ /=]; rewrite normmB subr0. Grab Existential Variables. all: end_near. Qed. Lemma OuO_to_P f g : OuO f g -> OuP f g. Proof. move=> fOg; apply/Ouex_to_P; move: fOg => /eqOP [k hk]. have /hk [Q [->]] : k < maxr 1 (k + 1) by rewrite ltr_maxr ltr_addl orbC ltr01. -move=> [R [[_/posnumP[e1] Re1] [_/posnumP[e2] Re2]] sRQ] fOg. -exists (minr e1%:num e2%:num) => //. -exists (maxr 1 (k + 1)); first by rewrite ltr_maxr ltr01. +rewrite /= -filter_from_norm_locally => -[_/posnumP[e] seQ] fOg. +exists e%:num => //; exists (maxr 1 (k + 1)); first by rewrite ltr_maxr ltr01. move=> x dx dxe Pdx; apply: (fOg (x, dx)); split=> //=. -move: dxe; rewrite ltr_maxl !ltr_minr => /andP[/andP [dxe11 _] /andP [_ dxe22]]. -by apply/sRQ => //; split; [apply/Re1|apply/Re2]; - rewrite /AbsRing_ball /= absrB subr0. +by apply: seQ => //; rewrite /ball normmB subr0. Qed. End UniformBigO. \ No newline at end of file diff --git a/posnum.v b/posnum.v index e807ffeec0..76f1d0a888 100644 --- a/posnum.v +++ b/posnum.v @@ -1,5 +1,4 @@ (* mathcomp analysis (c) 2017 Inria and AIST. License: CeCILL-C. *) -Require Import Reals. From Coq Require Import ssreflect ssrfun ssrbool. From mathcomp Require Import ssrnat eqtype choice ssralg ssrnum. Require Import boolp reals. @@ -36,12 +35,9 @@ Unset Strict Implicit. Unset Printing Implicit Defensive. Import GRing.Theory Num.Def Num.Theory. -Delimit Scope R_scope with coqR. Delimit Scope real_scope with real. -Close Scope R_scope. Open Scope ring_scope. Open Scope real_scope. -Bind Scope ring_scope with R. (* Enrico's trick for tc resolution in have *) Notation "!! x" := (ltac:(refine x)) (at level 100, only parsing). @@ -73,7 +69,6 @@ Notation "x %:pos" := (pos_of_num (Phantom _ x)) (at level 0, format "x %:pos") : ring_scope. Notation "x %:num" := (num_of_pos x) (at level 0, format "x %:num") : ring_scope. -Notation posreal := {posnum R}. Notation "2" := 2%:R : ring_scope. Section PosNum. @@ -140,5 +135,5 @@ Qed. Hint Resolve posnum_gt0. Hint Resolve posnum_ge0. Hint Resolve posnum_neq0. -Notation "[gt0 'of' x ]" := (posnum_gt0_def (Phantom R x)) +Notation "[gt0 'of' x ]" := (posnum_gt0_def (Phantom _ x)) (format "[gt0 'of' x ]"). diff --git a/summability.v b/summability.v index 08cc68b1ea..a2e0bd580c 100644 --- a/summability.v +++ b/summability.v @@ -1,10 +1,8 @@ (* mathcomp analysis (c) 2017 Inria and AIST. License: CeCILL-C. *) -Require Import Reals. From mathcomp Require Import ssreflect ssrfun ssrbool ssrnat eqtype choice. From mathcomp Require Import seq fintype bigop ssralg ssrint ssrnum finmap. From mathcomp Require Import matrix interval zmodp. -Require Import boolp reals. -Require Import Rstruct Rbar classical_sets posnum topology hierarchy. +Require Import boolp reals classical_sets posnum topology hierarchy. Set Implicit Arguments. Unset Strict Implicit. @@ -15,11 +13,8 @@ Local Open Scope classical_set_scope. (** For Pierre-Yves : definition of sums *) -From mathcomp Require fintype bigop finmap. - Section totally. -Import fintype bigop finmap. Local Open Scope fset_scope. (* :TODO: when eventually is generalized to any lattice *) (* totally can just be replaced by eventually *) @@ -39,10 +34,10 @@ Qed. Definition partial_sum {I : choiceType} {R : zmodType} (x : I -> R) (A : {fset I}) : R := \sum_(i : A) x (val i). -Definition sum (I : choiceType) {K : absRingType} {R : normedModType K} +Definition sum (I : choiceType) {K : numDomainType} {R : normedModType K} (x : I -> R) : R := lim (partial_sum x). -Definition summable (I : choiceType) {K : absRingType} {R : normedModType K} +Definition summable (I : choiceType) {K : realFieldType} {R : normedModType K} (x : I -> R) := \forall M \near +oo, \forall J \near totally, partial_sum (fun i => `|[x i]|) J <= M. diff --git a/topology.v b/topology.v index 49d982e79a..c8048a6668 100644 --- a/topology.v +++ b/topology.v @@ -1,8 +1,7 @@ (* mathcomp analysis (c) 2017 Inria and AIST. License: CeCILL-C. *) From mathcomp Require Import ssreflect ssrfun ssrbool ssrnat eqtype choice. -From mathcomp Require Import seq fintype ssralg finmap matrix. -Require Import boolp. -Require Import classical_sets posnum. +From mathcomp Require Import seq fintype ssralg finmap matrix ssrnum. +Require Import boolp classical_sets posnum. (******************************************************************************) (* This file develops tools for the manipulation of filters and basic *) @@ -1304,6 +1303,37 @@ move=> PQ xP; near=> y; apply: (near PQ y) => //; by apply: (near (near_join xP) y). Grab Existential Variables. all: end_near. Qed. +(* limit composition *) + +Lemma lim_cont {T V U : topologicalType} + (F : set (set T)) (FF : Filter F) + (f : T -> V) (h : V -> U) (a : V) : + {for a, continuous h} -> + f @ F --> a -> (h \o f) @ F --> h a. +Proof. +move=> h_cont fa fb; apply: (flim_trans _ h_cont). +exact: (@flim_comp _ _ _ _ h _ _ _ fa). +Qed. + +Lemma lim_cont2 {T V W U : topologicalType} + (F : set (set T)) (FF : Filter F) + (f : T -> V) (g : T -> W) (h : V -> W -> U) (a : V) (b : W) : + h z.1 z.2 @[z --> (a, b)] --> h a b -> + f @ F --> a -> g @ F --> b -> (fun x => h (f x) (g x)) @ F --> h a b. +Proof. +move=> h_cont fa fb; apply: (flim_trans _ h_cont). +exact: (@flim_comp _ _ _ _ (fun x => h x.1 x.2) _ _ _ (flim_pair fa fb)). +Qed. + +Lemma cst_continuous {T T' : topologicalType} (k : T') + (F : set (set T)) {FF : Filter F} : + (fun _ : T => k) @ F --> k. +Proof. +by move=> x; rewrite !near_simpl => /locally_singleton ?; apply: filterE. +Qed. +Arguments cst_continuous {T T'} k F {FF}. +Hint Resolve cst_continuous. + (** ** Topology defined by a filter *) Section TopologyOfFilter. @@ -1637,6 +1667,23 @@ Definition product_topologicalType := End Product_Topology. +(** * The topology on natural numbers *) + +(* :TODO: ultimately nat could be replaced by any lattice *) +Definition eventually := filter_from setT (fun N => [set n | (N <= n)%N]). +Notation "'\oo'" := eventually : classical_set_scope. + +Canonical eventually_filter_source X := + @Filtered.Source X _ nat (fun f => f @ \oo). + +Global Instance eventually_filter : ProperFilter eventually. +Proof. +eapply @filter_from_proper; last by move=> i _; exists i. +apply: filter_fromT_filter; first by exists 0%N. +move=> i j; exists (maxn i j) => n //=. +by rewrite geq_max => /andP[ltin ltjn]. +Qed. + (** locally' *) (* Should have a generic ^' operator *) @@ -2139,3 +2186,763 @@ End Covers. Definition connected (T : topologicalType) (A : set T) := forall B : set T, B !=set0 -> (exists2 C, open C & B = A `&` C) -> (exists2 C, closed C & B = A `&` C) -> B = A. + +(** * Uniform spaces defined using entourages *) + +Definition comp_rel (R S T : Type) (A : set (R * S)) (B : set (S * T)) := + [set xy | exists2 z, A (xy.1, z) & B (z, xy.2)]. +Notation "B \o A" := (comp_rel A B) : classical_set_scope. + +Definition inv_rel (T T' : Type) (A : set (T * T')) := + [set xy | A (xy.2, xy.1)]. +Notation "A ^-1" := (inv_rel A) : classical_set_scope. + +Definition set_of_rel (T T' : Type) (A : set (T * T')) (x : T) := + [set y | A (x, y)]. + +Definition locally_ {T T'} (ent : set (set (T * T'))) (x : T) := + filter_from ent (fun A => set_of_rel A x). + +Lemma locally_E {T T'} (ent : set (set (T * T'))) x : + locally_ ent x = filter_from ent (fun A => set_of_rel A x). +Proof. by []. Qed. + +Module Uniform. + +Record mixin_of (M : Type) (locally : M -> set (set M)) := Mixin { + entourages : (M * M -> Prop) -> Prop ; + ax1 : Filter entourages ; + ax2 : forall A, entourages A -> [set xy | xy.1 = xy.2] `<=` A ; + ax3 : forall A, entourages A -> entourages A^-1 ; + ax4 : forall A, entourages A -> exists2 B, entourages B & B \o B `<=` A ; + ax5 : locally = locally_ entourages +}. + +Record class_of (M : Type) := Class { + base : Topological.class_of M; + mixin : mixin_of (Filtered.locally_op base) +}. + +Section ClassDef. + +Structure type := Pack { sort; _ : class_of sort ; _ : Type }. +Local Coercion sort : type >-> Sortclass. +Variables (T : Type) (cT : type). +Definition class := let: Pack _ c _ := cT return class_of cT in c. + +Definition clone c of phant_id class c := @Pack T c T. +Let xT := let: Pack T _ _ := cT in T. +Notation xclass := (class : class_of xT). +Local Coercion base : class_of >-> Topological.class_of. +Local Coercion mixin : class_of >-> mixin_of. + +Definition pack loc (m : @mixin_of T loc) := + fun bT (b : Topological.class_of T) of phant_id (@Topological.class bT) b => + fun m' of phant_id m (m' : @mixin_of T (Filtered.locally_op b)) => + @Pack T (@Class _ b m') T. + +Definition eqType := @Equality.Pack cT xclass xT. +Definition choiceType := @Choice.Pack cT xclass xT. +Definition pointedType := @Pointed.Pack cT xclass xT. +Definition filteredType := @Filtered.Pack cT cT xclass xT. +Definition topologicalType := @Topological.Pack cT xclass xT. + +End ClassDef. + +Module Exports. + +Coercion sort : type >-> Sortclass. +Coercion base : class_of >-> Topological.class_of. +Coercion mixin : class_of >-> mixin_of. +Coercion eqType : type >-> Equality.type. +Canonical eqType. +Coercion choiceType : type >-> Choice.type. +Canonical choiceType. +Coercion pointedType : type >-> Pointed.type. +Canonical pointedType. +Coercion filteredType : type >-> Filtered.type. +Canonical filteredType. +Coercion topologicalType : type >-> Topological.type. +Canonical topologicalType. +Notation uniformType := type. +Notation UniformType T m := (@pack T _ m _ _ idfun _ idfun). +Notation UniformMixin := Mixin. +Notation "[ 'uniformType' 'of' T 'for' cT ]" := (@clone T cT _ idfun) + (at level 0, format "[ 'uniformType' 'of' T 'for' cT ]") : form_scope. +Notation "[ 'uniformType' 'of' T ]" := (@clone T _ _ id) + (at level 0, format "[ 'uniformType' 'of' T ]") : form_scope. + +End Exports. + +End Uniform. + +Export Uniform.Exports. + +Section UniformTopology. + +Program Definition topologyOfEntourageMixin (T : Type) + (loc : T -> set (set T)) (m : Uniform.mixin_of loc) : + Topological.mixin_of loc := topologyOfFilterMixin _ _ _. +Next Obligation. +rewrite (Uniform.ax5 m) locally_E; apply filter_from_proper; last first. + by move=> A entA; exists p; apply: Uniform.ax2 entA _ _. +apply: filter_from_filter. + by exists setT; apply: @filterT (Uniform.ax1 m). +move=> A B entA entB; exists (A `&` B) => //. +exact: (@filterI _ _ (Uniform.ax1 m)). +Qed. +Next Obligation. +move: H; rewrite (Uniform.ax5 m) locally_E => - [B entB sBpA]. +by apply: sBpA; apply: Uniform.ax2 entB _ _. +Qed. +Next Obligation. +move: H; rewrite (Uniform.ax5 m) locally_E => - [B entB sBpA]. +have /Uniform.ax4 [C entC sC2B] := entB. +exists C => // q Cpq; rewrite locally_E; exists C => // r Cqr. +by apply/sBpA/sC2B; exists q. +Qed. + +End UniformTopology. + +Definition entourages {M : uniformType} := Uniform.entourages (Uniform.class M). + +Lemma locally_entourageE {M : uniformType} : locally_ (@entourages M) = locally. +Proof. by case: M=> [?[?[]]]. Qed. + +Lemma filter_from_entourageE {M : uniformType} x : + filter_from (@entourages M) (fun A => set_of_rel A x) = locally x. +Proof. by rewrite -locally_entourageE. Qed. + +Module Export LocallyEntourage. +Definition locally_simpl := + (locally_simpl,@filter_from_entourageE,@locally_entourageE). +End LocallyEntourage. + +Lemma locallyP {M : uniformType} (x : M) P : + locally x P <-> locally_ entourages x P. +Proof. by rewrite locally_simpl. Qed. + +Section uniformType1. +Context {M : uniformType}. + +Lemma entourage_refl (A : set (M * M)) x : + entourages A -> A (x, x). +Proof. by move=> entA; apply: Uniform.ax2 entA _ _. Qed. + +Global Instance entourage_filter : ProperFilter (@entourages M). +Proof. +apply Build_ProperFilter; last exact: Uniform.ax1. +by move=> A entA; exists (point, point); apply: entourage_refl. +Qed. + +Lemma entourageT : entourages (@setT (M * M)). +Proof. exact: filterT. Qed. + +Lemma entourage_inv (A : set (M * M)) : entourages A -> entourages A^-1. +Proof. exact: Uniform.ax3. Qed. + +Lemma entourage_split (A : set (M * M)) : + entourages A -> exists2 B, entourages B & B \o B `<=` A. +Proof. exact: Uniform.ax4. Qed. + +Definition split_ent (A : set (M * M)) := + get (entourages `&` [set B | B \o B `<=` A]). + +Lemma split_entP (A : set (M * M)) : entourages A -> + entourages (split_ent A) /\ split_ent A \o split_ent A `<=` A. +Proof. by move/entourage_split/exists2P/getPex. Qed. + +Lemma entourages_split_ent (A : set (M * M)) : entourages A -> + entourages (split_ent A). +Proof. by move=> /split_entP []. Qed. +Hint Extern 0 (entourages (split_ent _)) => exact: entourages_split_ent : core. +Hint Extern 0 (entourages (get _)) => exact: entourages_split_ent : core. + +Lemma subset_split_ent (A : set (M * M)) : entourages A -> + split_ent A \o split_ent A `<=` A. +Proof. by move=> /split_entP []. Qed. + +Lemma entourages_split (z x y : M) A : entourages A -> + split_ent A (x,z) -> split_ent A (z,y) -> A (x,y). +Proof. by move=> /subset_split_ent sA ??; apply: sA; exists z. Qed. +Arguments entourages_split z {x y A}. + +Definition close (x y : M) : Prop := forall A, entourages A -> A (x,y). + +Lemma close_refl (x : M) : close x x. Proof. by move=> ? /entourage_refl. Qed. + +Lemma close_sym (x y : M) : close x y -> close y x. +Proof. by move=> cxy ? /entourage_inv /cxy. Qed. + +Lemma close_trans (x y z : M) : close x y -> close y z -> close x z. +Proof. +by move=> cxy cyz A entA; apply: (entourages_split y) => //; + [apply: cxy|apply: cyz]. +Qed. + +Lemma close_limxx (x y : M) : close x y -> x --> y. +Proof. +move=> cxy A /= /locallyP [B entB sBA]. +apply/locallyP; exists (split_ent B) => // z hByz; apply/sBA. +by apply: subset_split_ent => //; exists x => //=; apply: (close_sym cxy). +Qed. + +Lemma locally_entourage (x : M) A : entourages A -> locally x (set_of_rel A x). +Proof. by move=> ?; apply/locallyP; exists A. Qed. +Hint Resolve locally_entourage. + +Lemma flim_close {F} {FF : ProperFilter F} (x y : M) : + F --> x -> F --> y -> close x y. +Proof. +move=> Fx Fy A entA; apply: subset_split_ent => //. +near F => z; exists z => /=; near: z; first exact/Fx/locally_entourage. +by apply/Fy; apply: locally_entourage (entourage_inv _). +Grab Existential Variables. all: end_near. Qed. + +Lemma flimx_close (x y : M) : x --> y -> close x y. +Proof. exact: flim_close. Qed. + +Lemma flim_entouragesP F (FF : Filter F) (p : M) : + F --> p <-> forall A, entourages A -> \forall q \near F, A (p, q). +Proof. by rewrite -filter_fromP !locally_simpl. Qed. + +Lemma flim_entourages {F} {FF : Filter F} (y : M) : + F --> y -> forall A, entourages A -> \forall y' \near F, A (y,y'). +Proof. by move/flim_entouragesP. Qed. + +Lemma app_flim_entouragesP T (f : T -> M) F (FF : Filter F) p : + f @ F --> p <-> forall A, entourages A -> \forall t \near F, A (p, f t). +Proof. exact: flim_entouragesP. Qed. + +Lemma flimi_entouragesP T {F} {FF : Filter F} (f : T -> M -> Prop) y : + f `@ F --> y <-> + forall A, entourages A -> \forall x \near F, exists z, f x z /\ A (y,z). +Proof. +split=> [Fy A entA|Fy A] /=; first exact/Fy/locally_entourage. +move=> /locallyP [B entB sBA]; rewrite near_simpl near_mapi; near=> x. +by have [//|z [fxz Byz]]:= near (Fy _ entB) x; exists z; split=> //; apply: sBA. +Unshelve. all: end_near. Qed. +Definition flimi_locally := @flimi_entouragesP. + +Lemma flimi_entourages T {F} {FF : Filter F} (f : T -> M -> Prop) y : + f `@ F --> y -> + forall A, entourages A -> F [set x | exists z, f x z /\ A (y,z)]. +Proof. by move/flimi_entouragesP. Qed. + +Lemma flimi_close T {F} {FF: ProperFilter F} (f : T -> set M) (l l' : M) : + {near F, is_fun f} -> f `@ F --> l -> f `@ F --> l' -> close l l'. +Proof. +move=> f_prop fFl fFl'. +suff f_totalfun: infer {near F, is_totalfun f} by exact: flim_close fFl fFl'. +apply: filter_app f_prop; near=> x; split=> //=. +by have [//|y []] := near (flimi_entourages fFl entourageT) x; exists y. +Grab Existential Variables. all: end_near. Qed. +Definition flimi_locally_close := @flimi_close. + +Lemma flim_const {T} {F : set (set T)} + {FF : Filter F} (a : M) : a @[_ --> F] --> a. +Proof. +move=> A /locallyP [B entB sBA]; rewrite !near_simpl /=. +by apply: filterE => _; apply/sBA; apply: entourage_refl. +Qed. + +Lemma close_lim (F1 F2 : set (set M)) {FF2 : ProperFilter F2} : + F1 --> F2 -> F2 --> F1 -> close (lim F1) (lim F2). +Proof. +move=> F12 F21; have [/(flim_trans F21) F2l|dvgF1] := pselect (cvg F1). + by apply: (@flim_close F2) => //; apply: cvgP F2l. +have [/(flim_trans F12)/cvgP//|dvgF2] := pselect (cvg F2). +by rewrite dvgP // dvgP //; apply: close_refl. +Qed. + +Lemma flim_closeP (F : set (set M)) (l : M) : ProperFilter F -> + F --> l <-> ([cvg F in M] /\ close (lim F) l). +Proof. +move=> FF; split=> [Fl|[cvF]Cl]. + by have /cvgP := Fl; split=> //; apply: flim_close. +by apply: flim_trans (close_limxx Cl). +Qed. + +End uniformType1. + +Hint Extern 0 (entourages (split_ent _)) => exact: entourages_split_ent : core. +Hint Extern 0 (entourages (get _)) => exact: entourages_split_ent : core. +Arguments entourages_split {M} z {x y A}. +Hint Resolve locally_entourage. +Hint Resolve close_refl. +Arguments flim_const {M T F FF} a. +Arguments close_lim {M} F1 F2 {FF2} _. + +Definition unif_cont (U V : uniformType) (f : U -> V) := + (fun xy => (f xy.1, f xy.2)) @ entourages --> entourages. + +(** product of two uniform spaces *) + +Section prod_Uniform. + +Context {U V : uniformType}. +Implicit Types A : set ((U * V) * (U * V)). + +Definition prod_ent := + [set A : set ((U * V) * (U * V)) | + filter_prod (@entourages U) (@entourages V) + [set ((xy.1.1,xy.2.1),(xy.1.2,xy.2.2)) | xy in A]]. + +Lemma prod_ent_filter : Filter prod_ent. +Proof. +have prodF := filter_prod_filter (@entourage_filter U) (@entourage_filter V). +split; rewrite /prod_ent; last 1 first. +- by move=> A B sAB; apply: filterS => xy [zt /sAB Bzt ztexy]; exists zt. +- apply: filterS filterT => xy _. + exists ((xy.1.1,xy.2.1),(xy.1.2,xy.2.2)) => //=. + by rewrite -!surjective_pairing. +move=> A B entA entB; apply: filterS (filterI entA entB) => xy []. +move=> [zt Azt ztexy] [zt' Bzt' zt'exy]; exists zt => //; split=> //. +suff -> : zt = zt' by []. +move/eqP: ztexy; rewrite [xy]surjective_pairing. +rewrite [xy.1]surjective_pairing [xy.2]surjective_pairing !xpair_eqE. +rewrite andbACA -!xpair_eqE -!surjective_pairing => /eqP->. +move/eqP: zt'exy; rewrite [xy]surjective_pairing. +rewrite [xy.1]surjective_pairing [xy.2]surjective_pairing !xpair_eqE. +by rewrite andbACA -!xpair_eqE -!surjective_pairing => /eqP->. +Qed. + +Lemma prod_ent_refl A : prod_ent A -> [set xy | xy.1 = xy.2] `<=` A. +Proof. +move=> [B [entB1 entB2] sBA] xy /eqP. +rewrite [_.1]surjective_pairing [xy.2]surjective_pairing xpair_eqE. +move=> /andP [/eqP xy1e /eqP xy2e]. +have /sBA : (B.1 `*` B.2) ((xy.1.1, xy.2.1), (xy.1.2, xy.2.2)). + by rewrite xy1e xy2e; split=> /=; apply: entourage_refl. +move=> [zt Azt /eqP]; rewrite !xpair_eqE. +by rewrite andbACA -!xpair_eqE -!surjective_pairing => /eqP<-. +Qed. + +Lemma prod_ent_inv A : prod_ent A -> prod_ent A^-1. +Proof. +move=> [B [/entourage_inv entB1 /entourage_inv entB2] sBA]. +exists (B.1^-1,B.2^-1) => //= xy []. +rewrite [xy.1]surjective_pairing [xy.2]surjective_pairing /inv_rel /=. +move=> Bxy1 Bxy2. +have /sBA : (B.1 `*` B.2) ((xy.1.2, xy.1.1),(xy.2.2, xy.2.1)) by []. +move=> [zt Azt /eqP]; rewrite !xpair_eqE. +rewrite [X in X && _]andbC [X in _ && X]andbC -!xpair_eqE -!surjective_pairing. +by move/eqP<-; exists (zt.2, zt.1) => //=; rewrite -surjective_pairing. +Qed. + +Lemma prod_ent_split A : prod_ent A -> exists2 B, prod_ent B & B \o B `<=` A. +Proof. +move=> [B [entB1 entB2]] sBA; exists [set xy | split_ent B.1 (xy.1.1,xy.2.1) /\ + split_ent B.2 (xy.1.2,xy.2.2)]. + exists (split_ent B.1,split_ent B.2) => //= xy hBxy. + by exists ((xy.1.1, xy.2.1),(xy.1.2,xy.2.2)); rewrite -!surjective_pairing. +move=> xy [uv /= [hB1xyuv1 hB2xyuv1] [hB1xyuv2 hB2xyuv2]]. +have /sBA : (B.1 `*` B.2) ((xy.1.1, xy.2.1),(xy.1.2,xy.2.2)). + by split=> /=; apply: subset_split_ent => //; [exists uv.1|exists uv.2]. +move=> [zt Azt /eqP]; rewrite !xpair_eqE andbACA -!xpair_eqE. +by rewrite -!surjective_pairing => /eqP<-. +Qed. + +Lemma prod_ent_locallyE : locally = locally_ prod_ent. +Proof. +rewrite predeq2E => xy A; split. + move=> [B []]; rewrite -!locally_entourageE. + move=> [C1 entC1 sCB1] [C2 entC2 sCB2] sBA. + exists [set xy | C1 (xy.1.1, xy.2.1) /\ C2 (xy.1.2, xy.2.2)]. + exists (C1, C2) => //= zt [Czt1 Czt2]. + by exists ((zt.1.1, zt.2.1),(zt.1.2,zt.2.2)); rewrite -!surjective_pairing. + by move=> uv [/= /sCB1 Buv1 /sCB2 /(conj Buv1) /sBA]. +move=> [B [C [entC1 entC2] sCB] sBA]. +exists (set_of_rel C.1 xy.1,set_of_rel C.2 xy.2). + by rewrite -!locally_entourageE; split; [exists C.1|exists C.2]. +move=> uv [/= Cxyuv1 Cxyuv2]; apply: sBA. +have /sCB : (C.1 `*` C.2) ((xy.1,uv.1),(xy.2,uv.2)) by []. +move=> [zt Bzt /eqP]; rewrite /set_of_rel !xpair_eqE andbACA -!xpair_eqE. +by rewrite -!surjective_pairing => /eqP<-. +Qed. + +Definition prod_uniformType_mixin := + Uniform.Mixin prod_ent_filter prod_ent_refl prod_ent_inv prod_ent_split + prod_ent_locallyE. + +End prod_Uniform. + +Canonical prod_uniformType (U V : uniformType) := + UniformType (U * V) (@prod_uniformType_mixin U V). + +Lemma flim_entourages2P (U V : uniformType) {F : set (set U)} {G : set (set V)} + {FF : Filter F} {FG : Filter G} (y : U) (z : V): + (F, G) --> (y, z) <-> + forall A, entourages A -> + \forall y' \near F & z' \near G, A ((y,z),(y',z')). +Proof. +split=> [/flim_entouragesP FGyz A /FGyz|FGyz]. + by rewrite near_simpl -near2_pair. +apply/flim_entouragesP => A /FGyz. +by rewrite (near2_pair _ _ [set p | A ((y,z),p)]). +Qed. + +(** matrices *) + +Section matrix_Uniform. + +Variables (m n : nat) (T : uniformType). + +Implicit Types A : set ('M[T]_(m, n) * 'M[T]_(m, n)). + +Definition mx_ent := + filter_from + [set P : 'I_m -> 'I_n -> set (T * T) | forall i j, entourages (P i j)] + (fun P => [set MN : 'M[T]_(m, n) * 'M[T]_(m, n) | + forall i j, P i j (MN.1 i j, MN.2 i j)]). + +Lemma mx_ent_filter : Filter mx_ent. +Proof. +apply: filter_from_filter => [|A B entA entB]. + by exists (fun _ _ => setT) => _ _; apply: filterT. +exists (fun i j => A i j `&` B i j); first by move=> ??; apply: filterI. +by move=> MN ABMN; split=> i j; have [] := ABMN i j. +Qed. + +Lemma mx_ent_refl A : mx_ent A -> [set MN | MN.1 = MN.2] `<=` A. +Proof. +move=> [B entB sBA] MN MN1e2; apply: sBA => i j. +by rewrite MN1e2; apply: entourage_refl. +Qed. + +Lemma mx_ent_inv A : mx_ent A -> mx_ent A^-1. +Proof. +move=> [B entB sBA]; exists (fun i j => (B i j)^-1). + by move=> i j; apply: entourage_inv. +by move=> MN BMN; apply: sBA. +Qed. + +Lemma mx_ent_split A : mx_ent A -> exists2 B, mx_ent B & B \o B `<=` A. +Proof. +move=> [B entB sBA]. +have Bsplit : forall i j, exists C, entourages C /\ C \o C `<=` B i j. + by move=> ??; apply/exists2P/entourage_split. +exists [set MN : 'M[T]_(m, n) * 'M[T]_(m, n) | + forall i j, get [set C | entourages C /\ C \o C `<=` B i j] + (MN.1 i j, MN.2 i j)]. + by exists (fun i j => get [set C | entourages C /\ C \o C `<=` B i j]). +move=> MN [P CMN1P CPMN2]; apply/sBA => i j. +have /getPex [_] := Bsplit i j; apply; exists (P i j); first exact: CMN1P. +exact: CPMN2. +Qed. + +Lemma mx_ent_locallyE : locally = locally_ mx_ent. +Proof. +rewrite predeq2E => M A; split. + move=> [B]; rewrite -locally_entourageE => M_B sBA. + have {M_B} M_B : forall i j, exists C, entourages C /\ + set_of_rel C (M i j) `<=` B i j by move=> i j; apply/exists2P/M_B. + exists [set MN : 'M[T]_(m, n) * 'M[T]_(m, n) | + forall i j, get [set C | entourages C /\ set_of_rel C (M i j) `<=` B i j] + (MN.1 i j, MN.2 i j)]. + exists (fun i j => get [set C | entourages C /\ + set_of_rel C (M i j) `<=` B i j]) => // i j. + by have /getPex [] := M_B i j. + move=> N CMN; apply/sBA => i j; have /getPex [_] := M_B i j; apply. + exact/CMN. +move=> [B [C entC sCB] sBA]; exists (fun i j => set_of_rel (C i j) (M i j)). + by rewrite -locally_entourageE => i j; exists (C i j). +by move=> N CMN; apply/sBA/sCB. +Qed. + +Definition matrix_uniformType_mixin := + Uniform.Mixin mx_ent_filter mx_ent_refl mx_ent_inv mx_ent_split + mx_ent_locallyE. + +Canonical matrix_uniformType := + UniformType 'M[T]_(m, n) matrix_uniformType_mixin. + +End matrix_Uniform. + +Lemma flim_mx_entouragesP (T : uniformType) m n (F : set (set 'M[T]_(m,n))) + (FF : Filter F) (M : 'M[T]_(m,n)) : + F --> M <-> + forall A, entourages A -> \forall N \near F, + forall i j, A (M i j, (N : 'M[T]_(m,n)) i j). +Proof. +split. + rewrite filter_fromP => FM A entA. + apply: (FM (fun i j => set_of_rel A (M i j))). + by move=> ??; apply: locally_entourage. +move=> FM; apply/flim_entouragesP => A [P entP sPA]; near=> N. +apply: sPA => /=; near: N; set Q := \bigcap_ij P ij.1 ij.2. +apply: filterS (FM Q _); first by move=> N QN i j; apply: (QN _ _ (i, j)). +have -> : Q = + \bigcap_(ij in [set k | k \in [fset x in predT]%fset]) P ij.1 ij.2. + by rewrite predeqE => t; split=> Qt ij _; apply: Qt => //; rewrite !inE. +by apply: filter_bigI => ??; apply: entP. +Grab Existential Variables. all: end_near. Qed. + +(** Functional metric spaces *) + +Section fct_Uniform. + +Variable (T : choiceType) (U : uniformType). + +Definition fct_ent := + filter_from + [set P : T -> set (U * U) | forall t, entourages (P t)] + (fun P => [set fg | forall t, P t (fg.1 t, fg.2 t)]). + +Lemma fct_ent_filter : Filter fct_ent. +Proof. +apply: filter_from_filter; first by exists (fun=> setT) => _; apply: filterT. +move=> A B entA entB. +exists (fun t => A t `&` B t); first by move=> ?; apply: filterI. +by move=> fg ABfg; split=> t; have [] := ABfg t. +Qed. + +Lemma fct_ent_refl A : fct_ent A -> [set fg | fg.1 =fg.2] `<=` A. +Proof. +move=> [B entB sBA] fg feg; apply/sBA => t; rewrite feg. +exact: entourage_refl. +Qed. + +Lemma fct_ent_inv A : fct_ent A -> fct_ent A^-1. +Proof. +move=> [B entB sBA]; exists (fun t => (B t)^-1). + by move=> ?; apply: entourage_inv. +by move=> fg Bgf; apply/sBA. +Qed. + +Lemma fct_ent_split A : fct_ent A -> exists2 B, fct_ent B & B \o B `<=` A. +Proof. +move=> [B entB sBA]. +have Bsplit t : exists C, entourages C /\ C \o C `<=` B t. + exact/exists2P/entourage_split. +exists [set fg | forall t, + get [set C | entourages C /\ C \o C `<=` B t] (fg.1 t, fg.2 t)]. + by exists (fun t => get [set C | entourages C /\ C \o C `<=` B t]). +move=> fg [h Cfh Chg]; apply/sBA => t; have /getPex [_] := Bsplit t; apply. +by exists (h t); [apply: Cfh|apply: Chg]. +Qed. + +Definition fct_uniformType_mixin := + UniformMixin fct_ent_filter fct_ent_refl fct_ent_inv fct_ent_split erefl. + +Definition fct_topologicalTypeMixin := + topologyOfEntourageMixin fct_uniformType_mixin. + +Canonical generic_source_filter := @Filtered.Source _ _ _ (locally_ fct_ent). +Canonical fct_topologicalType := + TopologicalType (T -> U) fct_topologicalTypeMixin. +Canonical fct_uniformType := UniformType (T -> U) fct_uniformType_mixin. + +End fct_Uniform. + +(* TODO: is it possible to remove A's dependency in t? *) +Lemma flim_fct_entouragesP (T : choiceType) (U : uniformType) + (F : set (set (T -> U))) (FF : Filter F) (f : T -> U) : + F --> f <-> + forall A, (forall t, entourages (A t)) -> + \forall g \near F, forall t, A t (f t, g t). +Proof. +split. + move=> /flim_entouragesP Ff A entA. + apply: (Ff [set fg | forall t : T, A t (fg.1 t, fg.2 t)]). + by exists (fun t => A t). +move=> Ff; apply/flim_entouragesP => A [P entP sPA]; near=> g. +by apply: sPA => /=; near: g; apply: Ff. +Grab Existential Variables. all: end_near. Qed. + +(** ** Complete uniform spaces *) + +Definition cauchy {T : uniformType} (F : set (set T)) := (F, F) --> entourages. + +Lemma cvg_cauchy {T : uniformType} (F : set (set T)) : Filter F -> + [cvg F in T] -> cauchy F. +Proof. +move=> FF cvF A entA; have /entourage_split [B entB sB2A] := entA. +exists (set_of_rel B^-1 (lim F), set_of_rel B (lim F)). + split=> /=; apply: cvF; rewrite /= -locally_entourageE; last by exists B. + by exists B^-1 => //; apply: entourage_inv. +by move=> ab [/= Balima Blimb]; apply: sB2A; exists (lim F). +Qed. + +Module Complete. + +Definition axiom (T : uniformType) := + forall (F : set (set T)), ProperFilter F -> cauchy F -> F --> lim F. + +Section ClassDef. + +Record class_of (T : Type) := Class { + base : Uniform.class_of T ; + mixin : axiom (Uniform.Pack base T) +}. +Local Coercion base : class_of >-> Uniform.class_of. +Local Coercion mixin : class_of >-> Complete.axiom. + +Structure type := Pack { sort; _ : class_of sort ; _ : Type }. +Local Coercion sort : type >-> Sortclass. + +Variables (T : Type) (cT : type). + +Definition class := let: Pack _ c _ := cT return class_of cT in c. + +Definition clone c of phant_id class c := @Pack T c T. +Let xT := let: Pack T _ _ := cT in T. +Notation xclass := (class : class_of xT). + +Definition pack b0 (m0 : axiom (@Uniform.Pack T b0 T)) := + fun bT b of phant_id (Uniform.class bT) b => + fun m of phant_id m m0 => @Pack T (@Class T b m) T. + +Definition eqType := @Equality.Pack cT xclass xT. +Definition choiceType := @Choice.Pack cT xclass xT. +Definition pointedType := @Pointed.Pack cT xclass xT. +Definition filteredType := @Filtered.Pack cT cT xclass xT. +Definition topologicalType := @Topological.Pack cT xclass xT. +Definition uniformType := @Uniform.Pack cT xclass xT. + +End ClassDef. + +Module Exports. + +Coercion base : class_of >-> Uniform.class_of. +Coercion mixin : class_of >-> axiom. +Coercion sort : type >-> Sortclass. +Coercion eqType : type >-> Equality.type. +Canonical eqType. +Coercion choiceType : type >-> Choice.type. +Canonical choiceType. +Coercion pointedType : type >-> Pointed.type. +Canonical pointedType. +Coercion filteredType : type >-> Filtered.type. +Canonical filteredType. +Coercion topologicalType : type >-> Topological.type. +Canonical topologicalType. +Coercion uniformType : type >-> Uniform.type. +Canonical uniformType. +Notation completeType := type. +Notation "[ 'completeType' 'of' T 'for' cT ]" := (@clone T cT _ idfun) + (at level 0, format "[ 'completeType' 'of' T 'for' cT ]") : form_scope. +Notation "[ 'completeType' 'of' T ]" := (@clone T _ _ id) + (at level 0, format "[ 'completeType' 'of' T ]") : form_scope. +Notation CompleteType T m := (@pack T _ m _ _ idfun _ idfun). + +End Exports. + +End Complete. + +Export Complete.Exports. + +Section completeType1. + +Context {T : completeType}. + +Lemma complete_cauchy (F : set (set T)) (FF : ProperFilter F) : + cauchy F -> F --> lim F. +Proof. by case: T F FF => [? [?]]. Qed. + +End completeType1. +Arguments complete_cauchy {T} F {FF} _. + +Section matrix_Complete. + +Variables (T : completeType) (m n : nat). + +Lemma mx_complete (F : set (set 'M[T]_(m, n))) : + ProperFilter F -> cauchy F -> cvg F. +Proof. +move=> FF Fc. +have /(_ _ _) /complete_cauchy /app_flim_entouragesP cvF : + cauchy ((fun M : 'M[T]_(m, n) => M _ _) @ F). + move=> i j A /= entA; rewrite near_simpl -near2E near_map2. + by apply: Fc; exists (fun _ _ => A). +apply/cvg_ex. +set Mlim := \matrix_(i, j) (lim ((fun M : 'M[T]_(m, n) => M i j) @ F) : T). +exists Mlim; apply/flim_mx_entouragesP => A entA; near=> M => i j; near F => M'. +apply: subset_split_ent => //; exists (M' i j) => /=. + by near: M'; rewrite mxE; apply: cvF. +move: (i) (j); near: M'; near: M; apply: nearP_dep; apply: Fc. +by exists (fun _ _ => (split_ent A)^-1) => ?? //; apply: entourage_inv. +Grab Existential Variables. all: end_near. Qed. + +Canonical matrix_completeType := CompleteType 'M[T]_(m, n) mx_complete. + +End matrix_Complete. + +Section fun_Complete. + +Context {T : choiceType} {U : completeType}. + +Lemma fun_complete (F : set (set (T -> U))) + {FF : ProperFilter F} : cauchy F -> cvg F. +Proof. +move=> Fc. +have /(_ _) /complete_cauchy /app_flim_entouragesP cvF : cauchy (@^~_ @ F). + move=> t A /= entA; rewrite near_simpl -near2E near_map2. + by apply: Fc; exists (fun=> A). +apply/cvg_ex; exists (fun t => lim (@^~t @ F)). +apply/flim_fct_entouragesP => A entA; near=> f => t; near F => g. +apply: (entourages_split (g t)) => //; first by near: g; apply: cvF. +move: (t); near: g; near: f; apply: nearP_dep; apply: Fc. +by exists (fun t => (split_ent (A t))^-1) => ? //; apply: entourage_inv. +Grab Existential Variables. all: end_near. Qed. + +Canonical fun_completeType := CompleteType (T -> U) fun_complete. + +End fun_Complete. + +(** ** Limit switching *) + +Section Flim_switch. + +Context {T1 T2 : choiceType}. + +Lemma flim_switch_1 {U : uniformType} + F1 {FF1 : ProperFilter F1} F2 {FF2 : Filter F2} + (f : T1 -> T2 -> U) (g : T2 -> U) (h : T1 -> U) (l : U) : + f @ F1 --> g -> (forall x1, f x1 @ F2 --> h x1) -> h @ F1 --> l -> + g @ F2 --> l. +Proof. +move=> fg fh hl; apply/app_flim_entouragesP => A entA. +near F1 => x1; near=> x2; apply: (entourages_split (h x1)) => //. + by near: x1; apply/(hl (set_of_rel _ l))/locally_entourage. +apply: (entourages_split (f x1 x2)) => //. + by near: x2; apply/(fh x1 (set_of_rel _ _))/locally_entourage. +move: (x2); near: x1; have /flim_fct_entouragesP /(_ (fun=> _^-1)):= fg; apply. +by move=> _; apply: entourage_inv. +Grab Existential Variables. all: end_near. Qed. + +Lemma flim_switch_2 {U : completeType} + F1 {FF1 : ProperFilter F1} F2 {FF2 : ProperFilter F2} + (f : T1 -> T2 -> U) (g : T2 -> U) (h : T1 -> U) : + f @ F1 --> g -> (forall x, f x @ F2 --> h x) -> + [cvg h @ F1 in U]. +Proof. +move=> fg fh; apply: complete_cauchy => A entA. +rewrite !near_simpl -near2_pair near_map2; near=> x1 y1 => /=; near F2 => x2. +apply: (entourages_split (f x1 x2)) => //. + by near: x2; apply/(fh _ (set_of_rel _ _))/locally_entourage. +apply: (entourages_split (f y1 x2)) => //; last first. + by near: x2; apply/(fh _ (set_of_rel _^-1 _))/locally_entourage/entourage_inv. +apply: (entourages_split (g x2)) => //; move: (x2); [near: x1|near: y1]. + have /flim_fct_entouragesP /(_ (fun=> _^-1)) := fg; apply=> _. + exact: entourage_inv. +by have /flim_fct_entouragesP /(_ (fun=> _)) := fg; apply. +Grab Existential Variables. all: end_near. Qed. + +(* Alternative version *) +(* Lemma flim_switch {U : completeType} *) +(* F1 (FF1 : ProperFilter F1) F2 (FF2 : ProperFilter F2) (f : T1 -> T2 -> U) (g : T2 -> U) (h : T1 -> U) : *) +(* [cvg f @ F1 in T2 -> U] -> (forall x, [cvg f x @ F2 in U]) -> *) +(* [/\ [cvg [lim f @ F1] @ F2 in U], [cvg (fun x => [lim f x @ F2]) @ F1 in U] *) +(* & [lim [lim f @ F1] @ F2] = [lim (fun x => [lim f x @ F2]) @ F1]]. *) +Lemma flim_switch {U : completeType} + F1 (FF1 : ProperFilter F1) F2 (FF2 : ProperFilter F2) + (f : T1 -> T2 -> U) (g : T2 -> U) (h : T1 -> U) : + f @ F1 --> g -> (forall x1, f x1 @ F2 --> h x1) -> + exists l : U, h @ F1 --> l /\ g @ F2 --> l. +Proof. +move=> Hfg Hfh; have hcv := !! flim_switch_2 Hfg Hfh. +by exists [lim h @ F1 in U]; split=> //; apply: flim_switch_1 Hfg Hfh hcv. +Qed. + +End Flim_switch. \ No newline at end of file From 334a32b5a8770bbba5a10e75ad396da8d5e874af Mon Sep 17 00:00:00 2001 From: damien rouhling Date: Tue, 2 Oct 2018 09:56:16 +0200 Subject: [PATCH 02/11] Shorten proofs --- derive.v | 102 ++++++++++++++++++++-------------------------------- hierarchy.v | 99 +++++++++++++++++++++----------------------------- landau.v | 12 +++---- 3 files changed, 83 insertions(+), 130 deletions(-) diff --git a/derive.v b/derive.v index ed3d5671d5..a9abf0fe85 100644 --- a/derive.v +++ b/derive.v @@ -163,8 +163,7 @@ Proof. move=> /diff_locallyP [dfc]; rewrite -addrA. rewrite (littleo_bigO_eqo (cst (1 : R^o))); last first. apply/eqOP; near=> k; rewrite /cst [`|[_]|]normr1 mulr1. - near=> y; rewrite ltrW //; near: y. - rewrite !near_simpl -locally_nearE -filter_from_norm_locally. + near=> y; rewrite ltrW //; near: y; apply/locally_normP. by exists k; [near: k; exists 0|move=> ? /=; rewrite sub0r normmN]. rewrite addfo; first by move=> /eqolim; rewrite flim_shift add0r. by apply/eqolim0P; apply: (flim_trans (dfc 0)); rewrite linear0. @@ -292,11 +291,9 @@ rewrite /g2. have [/eqP ->|v0] := boolP (v == 0). rewrite (_ : (fun _ => _) = cst 0); first exact: cst_continuous. by rewrite funeqE => ?; rewrite scaler0 /k littleo_lim0 // scaler0. -apply/flim_normP => e e0. -rewrite nearE /= locally_simpl /locally' -filter_from_norm_locally. -have /(littleoP [littleo of k]) : 0 < e / (2 * `|[v]|). +apply/flim_normP => e e0; rewrite /locally' -filter_from_norm_locally. +have /(littleoP [littleo of k]) /locally_normP [i i0 Hi] : 0 < e / (2 * `|[v]|). by rewrite divr_gt0 // pmulr_rgt0 // normm_gt0. -rewrite near_simpl -locally_nearE -filter_from_norm_locally => -[i i0 Hi]. exists (i / `|[v]|); first by rewrite divr_gt0 // normm_gt0. move=> /= j; rewrite normmB subr0 ltr_pdivl_mulr ?normm_gt0 // => jvi j0. rewrite normmB subr0 normmZ -ltr_pdivl_mull ?normr_gt0 ?invr_neq0 //. @@ -371,9 +368,8 @@ rewrite funeqE => x; apply/eqP; case: (ler0P `|[x]|) => [|xn0]. by rewrite normm_le0 => /eqP ->; rewrite linear0. rewrite -normm_le0 -(mul0r `|[x]|) -ler_pdivr_mulr //. apply/ler0_addgt0P => _ /posnumP[e]; rewrite ler_pdivr_mulr //. -have /oid := posnum_gt0 e. -rewrite !near_simpl -locally_nearE -filter_from_norm_locally. -move=> [_ /posnumP[d] dfe]; set k := ((d%:num / 2) / (PosNum xn0)%:num)^-1. +have /oid /locally_normP [_/posnumP[d] dfe] := [gt0 of e%:num]. +set k := ((d%:num / 2) / (PosNum xn0)%:num)^-1. rewrite -{1}(@scalerKV _ _ k _ x) // linearZZ normmZ. rewrite -ler_pdivl_mull ?gtr0_norm // mulrCA. rewrite (@ler_trans _ (e%:num * `|[k^-1 *: x]|)) //; last first. @@ -613,8 +609,7 @@ Qed. Lemma linear_lipschitz (V' W' : normedModType R) (f : {linear V' -> W'}) : continuous f -> exists2 k, k > 0 & forall x, `|[f x]| <= k * `|[x]|. Proof. -move=> /(_ 0); rewrite linear0 => /(_ _ (locally_ball 0 1%:pos)). -rewrite !near_simpl -locally_nearE -filter_from_norm_locally. +move=> /(_ 0); rewrite linear0 => /(_ _ (locally_ball 0 1%:pos)) /locally_normP. move=> [_ /posnumP[e] he]; exists (2 / e%:num) => // x. case: (lerP `|[x]| 0) => [|xn0]. by rewrite normm_le0 => /eqP->; rewrite linear0 !normm0 mulr0. @@ -623,9 +618,9 @@ have kn0 : k != 0 by []. have abskgt0 : `|k| > 0 by rewrite normr_gt0. rewrite -[x in X in X <= _](scalerKV kn0) linearZZ normmZ -ler_pdivl_mull //. suff /he : ball norm 0 e%:num (k^-1 *: x). - rewrite /ball normmB subr0 => /ltrW /ler_trans; apply. + rewrite /= normmB subr0 => /ltrW /ler_trans; apply. by rewrite ger0_norm // mulVf. -rewrite /ball /= normmB subr0 normmZ normfV ger0_norm // invfM -mulrA mulVf //. +rewrite /= normmB subr0 normmZ normfV ger0_norm // invfM -mulrA mulVf //. by rewrite invf_div mulr1 [X in _ < X]splitr; apply: ltr_spaddr. Qed. @@ -648,15 +643,13 @@ Proof. apply/eqoP => _ /posnumP[e]. have /bigO_exP [_ /posnumP[k]] := bigOP [bigO of [O_ (0 : U) id of f]]. have /(_ (e%:num / k%:num)) := littleoP [littleo of [o_ (0 : V') id of g]]. -rewrite !near_simpl -locally_nearE -filter_from_norm_locally. -move=> /(_ _) [//|_/posnumP[d] hd]. +move=> /(_ _) /locally_normP [//|_/posnumP[d] hd]. apply: filter_app; near=> x => leOxkx; apply: ler_trans (hd _ _) _; last first. rewrite -ler_pdivl_mull //; apply: ler_trans leOxkx _. by rewrite invf_div mulrA -[_ / _ * _]mulrA mulVf // mulr1. -rewrite /ball normmB subr0 (ler_lt_trans leOxkx) //. -rewrite -ltr_pdivl_mull //; near: x. -rewrite -locally_nearE -filter_from_norm_locally. -by exists (k%:num ^-1 * d%:num) => // x; rewrite /ball normmB subr0. +rewrite /= normmB subr0 (ler_lt_trans leOxkx) // -ltr_pdivl_mull //; near: x. +apply/locally_normP; exists (k%:num ^-1 * d%:num) => // x. +by rewrite /= normmB subr0. Grab Existential Variables. all: end_near. Qed. Lemma compoO_eqox (K : realFieldType) (U V' W' : normedModType K) (f : U -> V') @@ -671,14 +664,12 @@ Lemma compOo_eqo (K : realFieldType) (U V' W' : normedModType K) (f : U -> V') Proof. apply/eqoP => _ /posnumP[e]. have /bigO_exP [_ /posnumP[k]] := bigOP [bigO of [O_ (0 : V') id of g]]. -rewrite !near_simpl -locally_nearE -filter_from_norm_locally. -move=> [_/posnumP[d] hd]; have ekgt0 : e%:num / k%:num > 0 by []. +move=> /locally_normP [_/posnumP[d] hd]; have ekgt0 : e%:num / k%:num > 0 by []. have /(_ _ ekgt0) := littleoP [littleo of [o_ (0 : U) id of f]]. apply: filter_app; near=> x => leoxekx; apply: ler_trans (hd _ _) _; last first. by rewrite -ler_pdivl_mull // mulrA [_^-1 * _]mulrC. rewrite /ball normmB subr0; apply: ler_lt_trans leoxekx _. -rewrite -ltr_pdivl_mull //; near: x. -rewrite !near_simpl -locally_nearE -filter_from_norm_locally. +rewrite -ltr_pdivl_mull //; near: x; apply/locally_normP. by exists ((e%:num / k%:num) ^-1 * d%:num)=> // x; rewrite /ball normmB subr0. Grab Existential Variables. all: end_near. Qed. @@ -724,9 +715,9 @@ Lemma bilinear_schwarz (U V' W' : normedModType R) (f : {bilinear U -> V' -> W'}) : continuous (fun p => f p.1 p.2) -> exists2 k, k > 0 & forall u v, `|[f u v]| <= k * `|[u]| * `|[v]|. Proof. -move=> /(_ 0); rewrite linear0r => /(_ _ (locally_ball 0 1%:pos)). -rewrite !near_simpl -locally_nearE -!filter_from_norm_locally. -move=> [[A B] /= [[_/posnumP[eA] sA] [_/posnumP[eB] sB]] sAB]. +move=> /(_ (0, 0)); rewrite linear0r => /(_ _ (locally_ball 0 1%:pos)). +move=> [[A B] [/locally_normP /= [_/posnumP[eA] sA]]]. +move=> /locally_normP [_/posnumP[eB] sB] sAB. set e := minr eA%:num eB%:num; exists ((2 / e) ^+2) => // u v. case: (lerP `|[u]| 0) => [|un0]. by rewrite normm_le0 => /eqP->; rewrite linear0l !normm0 mulr0 mul0r. @@ -760,8 +751,8 @@ apply/eqoP=> _ /posnumP[e]; near=> x; rewrite (ler_trans (fschwarz _ _))//. rewrite ler_pmul ?pmulr_rge0 //; last by rewrite ler_maxr orbC lerr. rewrite -ler_pdivl_mull //. suff : `|[x]| <= k%:num ^-1 * e%:num by apply: ler_trans; rewrite ler_maxr lerr. -near: x; rewrite !near_simpl -locally_nearE -filter_from_norm_locally. -by exists (k%:num ^-1 * e%:num) => // ? /=; rewrite normmB subr0 => /ltrW. +near: x; apply/locally_normP; exists (k%:num ^-1 * e%:num) => //. +by move=> ? /=; rewrite normmB subr0 => /ltrW. Grab Existential Variables. all: end_near. Qed. Fact dbilin (U V' W' : normedModType R) (f : {bilinear U -> V' -> W'}) p : @@ -771,19 +762,13 @@ Fact dbilin (U V' W' : normedModType R) (f : {bilinear U -> V' -> W'}) p : (fun q => f p.1 q.2 + f q.1 p.2) +o_ (0 : U * V') id. Proof. move=> fc; split=> [q|]. - apply: (@continuousD _ _ _ (fun q => f p.1 q.2) (fun q => f q.1 p.2)). - move=> A /(fc (_.1, _.2)). - rewrite !near_simpl -!locally_nearE -!filter_from_norm_locally /=. - move=> [PQ [[_/posnumP[eP] sP] [_/posnumP[eQ] sQ]] sPQ]. - exists (setT, PQ.2); first by split; [exists 1|exists eQ%:num]. - move=> xy [_ /= Qy]; apply: (sPQ (_.1,_.2)); split=> //=. - exact/sP/ball_center. - move=> A /(fc (_.1, _.2)). - rewrite !near_simpl -!locally_nearE -!filter_from_norm_locally /=. - move=> [PQ [[_/posnumP[eP] sP] [_/posnumP[eQ] sQ]] sPQ]. - exists (PQ.1, setT); first by split; [exists eP%:num|exists 1]. - move=> xy [/= Px _]; apply: (sPQ (_.1,_.2)); split=> //=. - exact/sQ/ball_center. + apply: (@continuousD _ _ _ (fun q => f p.1 q.2) (fun q => f q.1 p.2)); + move=> A /(fc (_.1, _.2)) [PQ [/locally_normP [_/posnumP[eP] sP]]]; + move=> /locally_normP [_/posnumP[eQ] sQ] sPQ /=; apply/locally_normP. + exists eQ%:num => // xy; rewrite /= ltr_maxl => /andP [_ /sQ Qy]. + by apply: (sPQ (_,_)); split=> //=; apply/sP. + exists eP%:num => // xy; rewrite /= ltr_maxl => /andP [/sP Px _]. + by apply: (sPQ (_,_)); split=> //=; apply/sQ. apply/eqaddoE; rewrite funeqE => q /=. rewrite linearDl !linearDr addrA addrC. rewrite -[f q.1 _ + _ + _]addrA [f q.1 _ + _]addrC addrA [f q.1 _ + _]addrC. @@ -912,8 +897,7 @@ rewrite -[X in X + _]mulr1 -[X in 1 / _ * X](@mulfVK _ (x ^+ 2)); last first. by rewrite sqrf_eq0. rewrite mulrA mulf_div mulr1. have hDx_neq0 : h + x != 0. - near: h; rewrite !near_simpl -locally_nearE -filter_from_norm_locally. - exists `|x|; first by rewrite normr_gt0. + near: h; apply/locally_normP; exists `|x|; first by rewrite normr_gt0. move=> h /=; rewrite normmB subr0 -subr_gt0 => lthx. rewrite -(normm_gt0 (h + x : R^o)) addrC -[h]opprK. apply: ltr_le_trans (ler_distm_dist _ _). @@ -927,14 +911,13 @@ rewrite div1r normfV [X in _ / X]normrM invfM [X in _ * X]mulrC. rewrite mulrA mulrAC ler_pdivr_mulr ?normr_gt0 ?mulf_neq0 //. rewrite mulrAC ler_pdivr_mulr ?normr_gt0 //. have : `|h * h| <= `|x / 2| * (e%:num * `|x * x| * `|[h : R^o]|). - rewrite !mulrA; near: h. - rewrite !near_simpl -locally_nearE -filter_from_norm_locally. + rewrite !mulrA; near: h; apply/locally_normP. exists (`|x / 2| * e%:num * `|x * x|). by rewrite !pmulr_rgt0 // normr_gt0 mulf_neq0. by move=> h /ltrW; rewrite normmB subr0 [`|h * _|]normrM => /ler_pmul; apply. move=> /ler_trans-> //; rewrite [X in X <= _]mulrC ler_pmul ?mulr_ge0 //. -near: h; rewrite !near_simpl -locally_nearE -filter_from_norm_locally. -exists (`|x| / 2); first by rewrite divr_gt0 ?normr_gt0. +near: h; apply/locally_normP; exists (`|x| / 2). + by rewrite divr_gt0 ?normr_gt0. move=> h; rewrite /ball normmB subr0 => lthhx; rewrite addrC -[h]opprK. apply: ler_trans (ler_distm_dist _ _); rewrite normmN [X in _ <= X]ger0_norm. rewrite ler_subr_addr -ler_subr_addl (splitr `|[x : R^o]|). @@ -1275,15 +1258,10 @@ have : (fun h => - ((1 / f x) * (1 / f (h *: v + x))) *: apply: flim_comp2 (@lim_mult _ _ _) => //=. apply: (@lim_opp _ [normedModType R of R^o]); rewrite expr2. exact/lim_scaler/lim_inv. -apply: flim_trans => A. -rewrite {1}/locally' !near_simpl -locally_nearE -filter_from_norm_locally. -move=> [_/posnumP[e] /= Ae]; move: fn0; apply: filter_app; near=> h => /=. -move=> fhvxn0; have he : ball norm 0 e%:num h. - near: h; rewrite /locally' -locally_nearE -filter_from_norm_locally. - by exists e%:num. -have hn0 : h != 0. - near: h; rewrite /locally' -locally_nearE -filter_from_norm_locally. - by exists e%:num. +apply: flim_trans; rewrite [in X in _ --> X]/locally' -filter_from_norm_locally. +move=> A [_/posnumP[e] Ae]; move: fn0; apply: filter_app; near=> h => /= fhvxn0. +have he : ball norm 0 e%:num h by near: h; apply/locally_normP; exists e%:num. +have hn0 : h != 0 by near: h; apply/locally_normP; exists e%:num. suff <- : - ((1 / f x) * (1 / f (h *: v + x))) *: (h^-1 *: (f (h *: v + x) - f x)) = h^-1 *: (1 / f (h *: v + x) - 1 / f x) by exact: Ae. @@ -1432,10 +1410,8 @@ apply/eqP; rewrite eqr_le; apply/andP; split. move=> A [e egt0 Ae]; exists e => // x xe xgt0; apply: Ae => //. exact/lt0r_neq0. near=> h; apply: mulr_ge0_le0. - rewrite invr_ge0; apply: ltrW; near: h. - by rewrite /at_right -filter_from_norm_locally; exists 1. - rewrite subr_le0 [_%:A]mulr1; apply: cmax; near: h. - rewrite /at_right -filter_from_norm_locally. + by rewrite invr_ge0; apply: ltrW; near: h; apply/locally_normP; exists 1. + rewrite subr_le0 [_%:A]mulr1; apply: cmax; near: h; apply/locally_normP. exists (b - c); first by rewrite subr_gt0 (itvP cab). move=> h; rewrite /ball normmB subr0. move=> /(ler_lt_trans (ler_norm _)); rewrite ltr_subr_addr inE => ->. @@ -1447,10 +1423,8 @@ apply: le0r_flim_map; last first. move=> A [e egt0 Ae]; exists e => // x xe xgt0; apply: Ae => //. exact/ltr0_neq0. near=> h; apply: mulr_le0. - rewrite invr_le0; apply: ltrW; near: h. - by rewrite /at_left -filter_from_norm_locally; exists 1. -rewrite subr_le0 [_%:A]mulr1; apply: cmax; near: h. -rewrite /at_left -filter_from_norm_locally. + by rewrite invr_le0; apply: ltrW; near: h; apply/locally_normP; exists 1. +rewrite subr_le0 [_%:A]mulr1; apply: cmax; near: h; apply/locally_normP. exists (c - a); first by rewrite subr_gt0 (itvP cab). move=> h; rewrite /ball normmB subr0. move=> /ltr_normlP []; rewrite ltr_subr_addl ltr_subl_addl inE => -> _. diff --git a/hierarchy.v b/hierarchy.v index 16736c89b5..d74f88c7ad 100644 --- a/hierarchy.v +++ b/hierarchy.v @@ -332,9 +332,6 @@ by rewrite funeqE => x; rewrite /locally_norm entourages_normE filter_from_entourageE. Qed. -Lemma locally_normP x P : locally x P <-> locally_norm x P. -Proof. by rewrite locally_locally_norm. Qed. - Lemma filter_from_norm_locally x : @filter_from K _ [set x : K | 0 < x] (ball norm x) = locally x. Proof. @@ -345,6 +342,10 @@ rewrite -locally_locally_norm => - [B [_/posnumP[e] seB] sBA]. by exists e%:num => // y xye; apply/sBA/seB. Qed. +Lemma locally_normP x P : + locally x P <-> @filter_from K _ [set x : K | 0 < x] (ball norm x) P. +Proof. by rewrite filter_from_norm_locally. Qed. + Lemma locally_normE (x : V) (P : set V) : locally_norm x P = \near x, P x. Proof. by rewrite locally_locally_norm near_simpl. Qed. @@ -400,6 +401,7 @@ Qed. End NormedModule1. Hint Resolve normm_ge0. +Hint Resolve ball_center. Module Export LocallyNorm. Definition locally_simpl := (locally_simpl,@locally_locally_norm,@filter_from_norm_locally). @@ -427,13 +429,9 @@ Lemma continuous_flim_norm {K : numDomainType} (V W : normedModType K) continuous f -> x --> l -> forall e : {posnum K}, `|[f l - f x]| < e. Proof. move=> cf xl e. -move/flim_norm: (cf l) => /(_ _ (posnum_gt0 e)). -rewrite nearE => /locallyP; rewrite -entourages_normE => -[A [i i0 sA]]. -apply; apply: sA. -have /@flim_norm : Filter [filter of x] by apply: filter_on_Filter. -move/(_ _ xl _ i0). -rewrite nearE => /locallyP; rewrite -entourages_normE => -[B [j j0 sB]]. -by apply; apply: sB => /=; rewrite subrr normm0. +have /flim_norm /(_ _ [gt0 of e%:num]) /locally_normP [_/posnumP[d]] := (cf l). +apply; have /flim_norm /(_ _ [gt0 of d%:num]) := xl. +by move=> /locally_normP [_/posnumP[d']]; apply. Qed. Section NormedModuleField. @@ -656,10 +654,8 @@ suff loc_preim r C : locally `|[p - r]| C -> locally r ((fun r => `|[p - r]|) @^-1` C). have [r []] := clp_q _ _ (loc_preim _ _ pp_B) (loc_preim _ _ pq_A). by exists `|[p - r]|. -move=> [D [_/posnumP[e] seD] sDC]; apply: locally_le_locally_norm. -exists [set pq | `|[pq.1 - pq.2]| < e%:num]. - by rewrite entourages_normE; apply: entourages_ball. -move=> s re_s; apply/sDC/seD => /=; apply: ler_lt_trans (ler_distm_dist _ _) _. +move=> /locally_normP [_/posnumP[e] sC]; apply/locally_normP. +exists e%:num=> // s re_s; apply/sC; apply: ler_lt_trans (ler_distm_dist _ _) _. by rewrite opprB addrC -subr_trans normmB. Qed. @@ -697,11 +693,10 @@ Canonical realFieldBar_filteredType := Global Instance realField_locally'_proper (x : K) : ProperFilter (locally' x). Proof. -apply: Build_ProperFilter => A; rewrite /locally' -locally_locally_norm. -move=> [B [_/posnumP[e] seB] sBA]. -exists (x + e%:num / 2); apply: sBA; last first. +apply: Build_ProperFilter => A; rewrite /locally' -filter_from_norm_locally. +move=> [_/posnumP[e] sA]; exists (x + e%:num / 2); apply: sA; last first. by rewrite eq_sym addrC -subr_eq subrr eq_sym. -apply: seB; rewrite opprD addrA subrr sub0r normmN [ `|[_]| ]ger0_norm //. +rewrite /= opprD addrA subrr sub0r normmN [ `|[_]| ]ger0_norm //. by rewrite {2}(splitr e%:num) ltr_spaddl. Qed. @@ -852,9 +847,8 @@ End matrix_normedMod. Lemma coord_continuous {R : realFieldType} m n i j : continuous (fun M : 'M[R]_(m.+1, n.+1) => M i j). Proof. -move=> M s /=; rewrite -filter_from_norm_locally => -[e e0 es]. -apply/locallyP; rewrite locally_entourageE -filter_from_norm_locally. -exists e => // N MN; apply/es; have /bigmaxr_ltrP MeN := MN. +move=> M A /= /locally_normP [_/posnumP[e] sA]; apply/locally_normP. +exists e%:num => // N MN; apply/sA; have /bigmaxr_ltrP MeN := MN. have /(MeN _) : (index (i, j) (enum [finType of 'I_m.+1 * 'I_n.+1]) < size [seq (`|(M - N) ij.1 ij.2|)%R | ij : 'I_m.+1 * 'I_n.+1])%N. @@ -872,9 +866,9 @@ Context {K : realDomainType} {U V : normedModType K}. Definition prod_norm (x : U * V) := maxr `|[x.1]| `|[x.2]|. -Lemma prod_norm_triangle : forall x y : U * V, prod_norm (x + y) <= prod_norm x + prod_norm y. +Lemma prod_norm_triangle x y : prod_norm (x + y) <= prod_norm x + prod_norm y. Proof. -by move=> [xu xv] [yu yv]; rewrite ler_maxl /=; apply/andP; split; +by rewrite ler_maxl /=; apply/andP; split; apply: ler_trans (ler_normm_add _ _) _; apply: ler_add; rewrite ler_maxr lerr // orbC. Qed. @@ -1063,9 +1057,8 @@ Proof. by move=> fc gc; apply: flim_comp2 fc gc _; apply: lim_mult. Qed. Lemma continuous_norm : continuous (@norm _ V). Proof. -move=> x; apply/flim_normP => _/posnumP[e] /=. -rewrite !near_simpl -locally_nearE -filter_from_norm_locally. -by exists e%:num => // y Hy; apply/(ler_lt_trans (ler_distm_dist _ _)). +move=> x A /= /locally_normP [_/posnumP[e] sA]; apply/locally_normP. +by exists e%:num => // ??; apply/sA; apply/(ler_lt_trans (ler_distm_dist _ _)). Qed. (* :TODO: yet, not used anywhere?! *) @@ -1241,9 +1234,9 @@ Definition at_right x := within (fun u : R => x < u) (locally x). Global Instance at_right_proper_filter (x : R) : ProperFilter (at_right x). Proof. -apply: Build_ProperFilter' => -[A [_/posnumP[d] sA]] /(_ (x + d%:num / 2)). -apply; last by rewrite ltr_addl. -apply: sA; rewrite /= opprD !addrA subrr add0r normrN normf_div !ger0_norm //. +apply: Build_ProperFilter'; rewrite /at_right -filter_from_norm_locally. +move=> [_/posnumP[e] /(_ (x + e%:num / 2))]; apply; last by rewrite ltr_addl. +rewrite /= opprD addrA subrr add0r [ `|[_]|]normrN normf_div !ger0_norm //. by rewrite ltr_pdivr_mulr // ltr_pmulr // (_ : 1 = 1%:R) // ltr_nat. Qed. @@ -1260,19 +1253,17 @@ Typeclasses Opaque at_left at_right. Lemma open_lt (y : R) : open (< y). Proof. -move=> x /=; rewrite -subr_gt0 => yDx_gt0. -rewrite entourages_normE locally_entourageE -filter_from_norm_locally. -exists (y - x) => // z. -by rewrite /ball /= normmB ltr_distl addrCA subrr addr0 => /andP[]. +rewrite openE => x /= ltxy; apply/locally_normP; exists (y - x). + by rewrite subr_gt0. +by move=> ?; rewrite /= normmB ltr_distl addrCA subrr addr0 => /andP[]. Qed. Hint Resolve open_lt. Lemma open_gt (y : R) : open (> y). Proof. -move=> x /=; rewrite -subr_gt0 => xDy_gt0. -rewrite entourages_normE locally_entourageE -filter_from_norm_locally. -exists (x - y) => // z. -by rewrite /ball /= normmB ltr_distl opprB addrCA subrr addr0 => /andP[]. +rewrite openE => x /= gtxy; apply/locally_normP; exists (x - y). + by rewrite subr_gt0. +by move=> ?; rewrite /= normmB ltr_distl opprB addrCA subrr addr0 => /andP[]. Qed. Hint Resolve open_gt. @@ -1372,7 +1363,7 @@ wlog ltyx : a b (* leab *) A y Ay Acl x / y < x. move: Acl => [B Bcl AeabB]. have sAab : A `<=` [set x | x \in `[a, b]] by rewrite AeabB => ? []. move=> /asboolPn; rewrite asbool_and => /nandP [/asboolPn /(_ (sAab _))|] //. -move=> /imply_asboolPn [abx nAx] [C Cop AeabC]. +move=> /imply_asboolPn [abx nAx] [C]; rewrite openE => Cop AeabC. set Altx := fun y => y \in A `&` (< x). have Altxn0 : reals.nonempty Altx by exists y; rewrite in_setE. have xub_Altx : x \in ub Altx. @@ -1386,8 +1377,8 @@ have Az : A z. rewrite AeabB; split. suff : {subset `[y, x] <= `[a, b]} by apply. by apply/subitvP; rewrite /= (itvP abx); have /sAab/itvP-> := Ay. - apply: Bcl => D; rewrite -filter_from_norm_locally => -[_/posnumP[e] ze_D]. - have [t] := sup_adherent Altxsup (posnum_gt0 e). + apply: Bcl => D /locally_normP [_/posnumP[e] ze_D]. + have [t] := sup_adherent Altxsup [gt0 of e%:num]. rewrite in_setE => - [At lttx] ltzet. exists t; split; first by move: At; rewrite AeabB => - []. apply/ze_D; rewrite /ball /= ltr_distl. @@ -1397,9 +1388,7 @@ have Az : A z. have ltzx : 0 < x - z. have : z <= x by rewrite (itvP yxz). by rewrite subr_gt0 ler_eqVlt => /orP [/eqP zex|] //; move: nAx; rewrite -zex. -have := Az; rewrite AeabC => -[_ /Cop]. -rewrite entourages_normE locally_entourageE -filter_from_norm_locally. -move => [_ /posnumP[e] ze_C]. +have := Az; rewrite AeabC => -[_ /Cop /(locally_normP _ C) [_/posnumP[e] ze_C]]. suff [t Altxt] : exists2 t, Altx t & z < t. by rewrite ltrNge => /negP; apply; apply/sup_upper_bound. exists (z + (minr (e%:num / 2) ((PosNum ltzx)%:num / 2))); last first. @@ -1465,9 +1454,7 @@ have /fop := Di; rewrite openE => /(_ _ fx). rewrite /(_^°) -filter_from_norm_locally => - [_ /posnumP[e] xe_fi]. have /clAx [y [[aby [D' sD [sayUf _]]] xe_y]] := locally_ball x e. exists (i |` D')%fset; first by move=> j /fset1UP[->|/sD] //; rewrite in_setE. -split=> [z axz|]; last first. - exists i; first by rewrite !inE eq_refl. - exact/xe_fi/ball_center. +split=> [z axz|]; last by exists i; [rewrite !inE eq_refl|apply/xe_fi]. case: (lerP z y) => [lezy|ltyz]. have /sayUf [j Dj fjz] : z \in `[a, y] by rewrite inE (itvP axz) lezy. by exists j => //; rewrite inE orbC Dj. @@ -1518,9 +1505,8 @@ exists (sup A) => //; have lefsupv : f (sup A) <= v. rewrite lerNgt; apply/negP => ltvfsup. have vltfsup : 0 < f (sup A) - v by rewrite subr_gt0. have /fcont /(_ _ (locally_ball _ (PosNum vltfsup))) := supAab. - rewrite !near_simpl -locally_nearE -filter_from_norm_locally. - move=> [_/posnumP[d] supdfe]. - have [t At supd_t] := sup_adherent supA (posnum_gt0 d). + rewrite locally_simpl => /= /locally_normP [_/posnumP[d] supdfe]. + have [t At supd_t] := sup_adherent supA [gt0 of d%:num]. suff /supdfe : ball norm (sup A) d%:num t. rewrite /ball ltr_norml => /andP [_]. by rewrite ltr_add2l ltr_oppr opprK ltrNge; have /andP [_ ->] := At. @@ -1531,19 +1517,15 @@ apply/eqP; rewrite eqr_le; apply/andP; split=> //. rewrite -subr_le0; apply/ler0_addgt0P => _/posnumP[e]. rewrite ler_subl_addr -ler_subl_addl ltrW //. have /fcont /(_ _ (locally_ball _ e)) := supAab. -rewrite !near_simpl -locally_nearE -filter_from_norm_locally. -move=> [_/posnumP[d] supdfe]. +rewrite locally_simpl /= => /locally_normP [_/posnumP[d] supdfe]. have atrF := at_right_proper_filter (sup A); near (at_right (sup A)) => x. have /supdfe /= : ball norm (sup A) d%:num x. - near: x; rewrite /at_right !near_simpl -locally_nearE. - by rewrite -filter_from_norm_locally; exists d%:num. + by near: x; apply/locally_normP; exists d%:num. move/ltr_distW; apply: ler_lt_trans. rewrite ler_add2r ltrW //; suff : forall t, t \in `](sup A), b] -> v < f t. apply; rewrite inE; apply/andP; split. - near: x; rewrite /at_right !near_simpl -locally_nearE. - by rewrite -filter_from_norm_locally; exists 1. - near: x; rewrite /at_right !near_simpl -locally_nearE. - rewrite -filter_from_norm_locally; exists (b - sup A). + by near: x; apply/locally_normP; exists 1. + near: x; apply/locally_normP; exists (b - sup A). rewrite subr_gt0 ltr_def (itvP supAab) andbT; apply/negP => /eqP besup. by move: lefsupv; rewrite lerNgt -besup ltvfb. move=> t lttb ltsupt; move: lttb; rewrite /ball normmB. @@ -1730,8 +1712,7 @@ have covA : A `<=` \bigcup_(n : int) [set p | `|[p]| < n%:~R]. by rewrite rmorphD /= -floorE floorS_gtr. have /Aco [] := covA. move=> n _; rewrite openE => p; rewrite -subr_gt0 => ltpn. - apply/locallyP; rewrite locally_entourageE -filter_from_norm_locally. - exists (n%:~R - `|[p]|) => // q. + apply/locally_normP; exists (n%:~R - `|[p]|) => // q. rewrite /ball ltr_subr_addr normmB; apply: ler_lt_trans. by rewrite -{1}(subrK p q) ler_normm_add. move=> D _ DcovA. diff --git a/landau.v b/landau.v index 197a3bc08b..4fbb79a720 100644 --- a/landau.v +++ b/landau.v @@ -1089,8 +1089,7 @@ Lemma near_shift {K : realFieldType} {R : normedModType K} (y x : R) (P : set R) : (\near x, P x) = (\forall z \near y, (P \o shift (x - y)) z). Proof. -rewrite propeqE; split; - rewrite -!locally_nearE -!filter_from_norm_locally => -[_/posnumP[e] ye]; +rewrite -!filter_from_normE propeqE; split=> -[_/posnumP[e] ye]; exists e%:num => // t /= et. apply: ye; rewrite /= !opprD addrA addrACA subrr add0r. by rewrite opprK addrC. @@ -1120,17 +1119,16 @@ Hypothesis (normm_s : forall k x, `|[s k x]| = `|k| * `|[x]|). Lemma linear_for_continuous (f: {linear U -> V | GRing.Scale.op s_law}) : (f : _ -> _) =O_ (0 : U) (cst (1 : R^o)) -> continuous f. Proof. -move=> /eqO_exP [_/posnumP[k0] Of1] x. +move=> /eqO_exP [_/posnumP[k0] /locally_normP [_/posnumP[d]]]. +rewrite /cst [X in _ * X]normr1 mulr1 => fk0 x. apply/flim_normP => _/posnumP[e]; rewrite !near_simpl. rewrite (near_shift 0) /= subr0; near=> y => /=. rewrite -linearB opprD addrC addrNK linearN normmN; near: y. suff flip : \forall k \near +oo, forall x, `|[f x]| <= k * `|[x]|. near +oo in R => k; near=> y. rewrite (ler_lt_trans (near flip k _ _)) // -ltr_pdivl_mull //. - near: y; rewrite -locally_nearE -filter_from_norm_locally. + near: y; apply/locally_normP. by eexists; last by move=> ?; rewrite /= sub0r normmN; apply. -move: Of1; rewrite near_simpl -locally_nearE -filter_from_norm_locally. -move=> [_/posnumP[d]]; rewrite /cst [X in _ * X]normr1 mulr1 => fk. near=> k => y; case: (ler0P `|[y]|) => [|y0]. by rewrite normm_le0 => /eqP->; rewrite linear0 !normm0 mulr0. have ky0 : 0 <= k0%:num / (k * `|[y]|). @@ -1139,7 +1137,7 @@ rewrite -[X in _ <= X]mulr1 -ler_pdivr_mull ?pmulr_rgt0 //. rewrite -(ler_pmul2l [gt0 of k0%:num]) mulr1 mulrA -[_ / _]ger0_norm //. rewrite -normm_s. have <- : GRing.Scale.op s_law =2 s by rewrite GRing.Scale.opE. -rewrite -linearZ fk //= normmB subr0 normmZ ger0_norm //. +rewrite -linearZ fk0 //= normmB subr0 normmZ ger0_norm //. rewrite invfM mulrA mulfVK ?lt0r_neq0 // ltr_pdivr_mulr //. by rewrite mulrC -ltr_pdivr_mulr //; near: k; apply: locally_pinfty_gt. Grab Existential Variables. all: end_near. Qed. From 22d8bac9448597d9ac954138ee3ec2ccd684e557 Mon Sep 17 00:00:00 2001 From: damien rouhling Date: Wed, 3 Oct 2018 14:42:26 +0200 Subject: [PATCH 03/11] Use {ereal R} in hierarchy --- altreals/distr.v | 10 ++++---- altreals/realseq.v | 22 ++++++++--------- altreals/realsum.v | 7 +++--- hierarchy.v | 60 +++++++++++++++++++--------------------------- reals.v | 54 +++++++++++++++++++++-------------------- 5 files changed, 72 insertions(+), 81 deletions(-) diff --git a/altreals/distr.v b/altreals/distr.v index 0d9646d13f..8ed390ae40 100644 --- a/altreals/distr.v +++ b/altreals/distr.v @@ -544,7 +544,7 @@ End DLetAlg. (* -------------------------------------------------------------------- *) Definition mlim T (f : nat -> distr T) : T -> R := - fun x => nlim (fun n => f n x). + fun x => (nlim (fun n => f n x) : [numDomainType of R]). Lemma isd_mlim T (f : nat -> distr T) : isdistr (mlim f). Proof. split=> [x|J]; rewrite /mlim. @@ -552,7 +552,7 @@ Proof. split=> [x|J]; rewrite /mlim. by move=> n; apply/ge0_mu. move=> uqJ; pose F j := if `[< iscvg (fun n => f n j) >] then fun n => f n j else 0%:S. -apply/(@ler_trans _ (\sum_(j <- J) (nlim (F j) : R))). +apply/(@ler_trans _ (\sum_(j <- J) (nlim (F j) : [numDomainType of R]))). apply/ler_sum=> j _; rewrite /F; case/boolP: `[< _ >] => //. move/asboolPn=> h; rewrite nlimC; case: nlimP=> //. by case=> // l cf; case: h; exists l. @@ -579,7 +579,7 @@ Definition dlim T (f : nat -> distr T) := Notation "\dlim_ ( n ) E" := (dlim (fun n => E)). Lemma dlimE T (f : nat -> distr T) x : - (\dlim_(n) f n) x = nlim (fun n => f n x). + (\dlim_(n) f n) x = (nlim (fun n => f n x) : [numDomainType of R]). Proof. by unlock dlim. Qed. (* -------------------------------------------------------------------- *) @@ -1192,7 +1192,7 @@ Definition convexon (a b : {ereal R}) (f : R -> R) := forall t, 0 <= t <= 1 -> f (t * x + (1 - t) * y) <= t * (f x) + (1 - t) * (f y). -Notation convex f := (convexon \-inf \+inf f). +Notation convex f := (convexon -oo +oo f). Section Jensen. Context (f : R -> R) (x l : I -> R). @@ -1228,6 +1228,6 @@ Qed. End Jensen. End Jensen. -Notation convex f := (convexon \-inf \+inf f). +Notation convex f := (convexon -oo +oo f). (* -------------------------------------------------------------------- *) diff --git a/altreals/realseq.v b/altreals/realseq.v index df144fa3d0..9c2de5a706 100644 --- a/altreals/realseq.v +++ b/altreals/realseq.v @@ -49,8 +49,8 @@ Context {R : realType}. Inductive nbh : {ereal R} -> predArgType := | NFin (c e : R) of (0 < e) : nbh c%:E -| NPInf (M : R) : nbh \+inf -| NNInf (M : R) : nbh \-inf. +| NPInf (M : R) : nbh +oo +| NNInf (M : R) : nbh -oo. Coercion pred_of_nbh l (v : nbh l) := match v with @@ -73,16 +73,16 @@ by move=> e eE v; case: v eE => // c' e' h [->]. Qed. Lemma nbh_pinfW (P : forall x, nbh x -> Prop) : - (forall M, P _ (@NPInf R M)) -> forall (v : nbh \+inf), P _ v. + (forall M, P _ (@NPInf R M)) -> forall (v : nbh +oo), P _ v. Proof. -move=> ih ; move: {-2}\+inf (erefl (@ERPInf R)). +move=> ih ; move: {-2}+oo (erefl (@ERPInf R)). by move=> e eE v; case: v eE => // c' e' h [->]. Qed. Lemma nbh_ninfW (P : forall x, nbh x -> Prop) : - (forall M, P _ (@NNInf R M)) -> forall (v : nbh \-inf), P _ v. + (forall M, P _ (@NNInf R M)) -> forall (v : nbh -oo), P _ v. Proof. -move=> ih ; move: {-2}\-inf (erefl (@ERNInf R)). +move=> ih ; move: {-2}-oo (erefl (@ERNInf R)). by move=> e eE v; case: v eE => // c' e' h [->]. Qed. End NbhElim. @@ -309,7 +309,7 @@ move=> cu cv; pose a := u \- lu%:S; pose b := v \- lv%:S. have eq: (u \* v) =1 (lu * lv)%:S \+ ((lu%:S \* b) \+ (a \* v)). move=> n; rewrite {}/a {}/b /= [u n+_]addrC [(_+_)*(v n)]mulrDl. rewrite !addrA -[LHS]add0r; congr (_ + _); rewrite mulrDr. - by rewrite !(mulrN, mulNr) [X in X-_]addrCA subrr addr0 subrr. + by rewrite !(mulrN, mulNr) addrCA subrr addr0 subrr. apply/(ncvg_eq eq); rewrite -[X in X%:E]addr0; apply/ncvgD. by apply/ncvgC. rewrite -[X in X%:E]addr0; apply/ncvgD. + apply/ncvgMr; first rewrite -[X in X%:E](subrr lv). @@ -429,7 +429,7 @@ Implicit Types (u v : nat -> R). Definition nlim u : {ereal R} := if @idP `[exists l, `[< ncvg u l >]] is ReflectT Px then - xchooseb Px else \-inf. + xchooseb Px else -oo. Lemma nlim_ncvg u : (exists l, ncvg u l) -> ncvg u (nlim u). Proof. @@ -439,7 +439,7 @@ move=> p; rewrite -[xchooseb _](ncvg_uniq cv_u_l) //. by apply/asboolP/(xchoosebP p). Qed. -Lemma nlim_out u : ~ (exists l, ncvg u l) -> nlim u = \-inf. +Lemma nlim_out u : ~ (exists l, ncvg u l) -> nlim u = -oo. Proof. move=> h; rewrite /nlim; case: {-}_ / idP => // p. by case: h; case/existsbP: p => l /asboolP; exists l. @@ -447,7 +447,7 @@ Qed. CoInductive nlim_spec (u : nat -> R) : er R -> Type := | NLimCvg l : ncvg u l -> nlim_spec u l -| NLimOut : ~ (exists l, ncvg u l) -> nlim_spec u \-inf. +| NLimOut : ~ (exists l, ncvg u l) -> nlim_spec u -oo. Lemma nlimP u : nlim_spec u (nlim u). Proof. @@ -521,7 +521,7 @@ Qed. Lemma nlim_sumR {I : eqType} (u : I -> nat -> R) (r : seq I) : (forall i, i \in r -> iscvg (u i)) -> nlim (fun n => \sum_(i <- r) (u i) n) - = (\sum_(i <- r) (nlim (u i) : R))%:E. + = (\sum_(i <- r) (nlim (u i) : [numDomainType of R]))%:E. Proof. move=> h; rewrite nlim_sum //; elim: r h => [|i r ih] h. by rewrite !big_nil. diff --git a/altreals/realsum.v b/altreals/realsum.v index db9975e1d7..94a3a822ee 100644 --- a/altreals/realsum.v +++ b/altreals/realsum.v @@ -169,12 +169,12 @@ Context {R : realType}. Lemma ncvg_mono (u : nat -> R) : (* {mono u : x y / (x <= y)%N >-> u x <= u y *) (forall x y, (x <= y)%N -> u x <= u y) - -> exists2 l, (\-inf < l)%E & ncvg u l. + -> exists2 l, (-oo < l)%E & ncvg u l. Proof. move=> mono_u; pose E := [pred x | `[exists n, x == u n]]. have nzE: nonempty E by exists (u 0%N); apply/imsetbP; exists 0%N. case/boolP: `[< has_sup E >] => /asboolP; last first. - move/has_supPn=> -/(_ nzE) h; exists \+inf => //; elim/nbh_pinfW => M /=. + move/has_supPn=> -/(_ nzE) h; exists +oo => //; elim/nbh_pinfW => M /=. case/(_ M): h=> x /imsetbP[K -> lt_MuK]; exists K=> n le_Kn; rewrite inE. by apply/(ltr_le_trans lt_MuK)/mono_u. move=> supE; exists (sup E)%:E => //; elim/nbh_finW=>e /= gt0_e. @@ -490,7 +490,8 @@ Hypothesis smS : summable S. Hypothesis homo_P : forall n m, (n <= m)%N -> (P n `<=` P m). Hypothesis cover_P : forall x, S x != 0 -> exists n, x \in P n. -Lemma psum_as_lim : psum S = nlim (fun n => \sum_(j : P n) (S (val j))). +Lemma psum_as_lim : + psum S = (nlim (fun n => \sum_(j : P n) (S (val j))) : [numDomainType of R]). Proof. set v := fun n => _; have hm_v m n: (m <= n)%N -> v m <= v n. by move=> le_mn; apply/big_fset_subset/fsubsetP/homo_P. diff --git a/hierarchy.v b/hierarchy.v index d74f88c7ad..30ea087b21 100644 --- a/hierarchy.v +++ b/hierarchy.v @@ -661,37 +661,29 @@ Qed. End NormedModuleRealField. -(** * Adding infinities to realFieldTypes *) +(** * Filters on extended reals *) -Section RealFieldBar. +Section ExtendedReals. -Context {K : realFieldType}. +Context {R : realFieldType}. -Inductive realFieldBar := | Finite of K | p_infty | m_infty. - -Notation "'+oo'" := p_infty : real_scope. -Notation "'-oo'" := m_infty : real_scope. - -Definition realFieldBar_locally' (a : realFieldBar) (A : set K) := +Definition ereal_locally' (a : {ereal R}) (A : set R) := match a with - | Finite a => locally' a A - | +oo => exists M : K, forall x, M < x -> A x - | -oo => exists M : K, forall x, x < M -> A x + | a%:E => locally' a A + | +oo => exists M : R, forall x, M < x -> A x + | -oo => exists M : R, forall x, x < M -> A x end. -Definition realFieldBar_locally (a : realFieldBar) (A : set K) := +Definition ereal_locally (a : {ereal R}) (A : set R) := match a with - | Finite a => locally a A - | +oo => exists M : K, forall x, M < x -> A x - | -oo => exists M : K, forall x, x < M -> A x + | a%:E => locally a A + | +oo => exists M : R, forall x, M < x -> A x + | -oo => exists M : R, forall x, x < M -> A x end. -Canonical realFieldBar_eqType := EqType realFieldBar gen_eqMixin. -Canonical realFieldBar_choiceType := ChoiceType realFieldBar gen_choiceMixin. -Canonical realFieldBar_pointedType := PointedType realFieldBar (+oo). -Canonical realFieldBar_filteredType := - FilteredType K realFieldBar (realFieldBar_locally). +Canonical ereal_pointedType := PointedType {ereal R} (+oo). +Canonical ereal_filteredType := FilteredType R {ereal R} (ereal_locally). -Global Instance realField_locally'_proper (x : K) : ProperFilter (locally' x). +Global Instance realField_locally'_proper (x : R) : ProperFilter (locally' x). Proof. apply: Build_ProperFilter => A; rewrite /locally' -filter_from_norm_locally. move=> [_/posnumP[e] sA]; exists (x + e%:num / 2); apply: sA; last first. @@ -700,8 +692,8 @@ rewrite /= opprD addrA subrr sub0r normmN [ `|[_]| ]ger0_norm //. by rewrite {2}(splitr e%:num) ltr_spaddl. Qed. -Global Instance realFieldBar_locally'_filter : - forall x, ProperFilter (realFieldBar_locally' x). +Global Instance ereal_locally'_filter : + forall x, ProperFilter (ereal_locally' x). Proof. case=> [x||]; first exact: realField_locally'_proper. apply Build_ProperFilter. @@ -715,18 +707,18 @@ split=> /= [|P Q [MP ltMP] [MQ ltMQ] |P Q sPQ [M ltMP]]; first by exists 0. by exists (minr MP MQ) => ?; rewrite ltr_minr => /andP [/ltMP ? /ltMQ]. by exists M => ? /ltMP /sPQ. Qed. -Typeclasses Opaque realFieldBar_locally'. +Typeclasses Opaque ereal_locally'. -Global Instance realFieldBar_locally_filter : - forall x, ProperFilter (realFieldBar_locally x). +Global Instance ereal_locally_filter : + forall x, ProperFilter (ereal_locally x). Proof. case=> [x||]; first exact/locally_filter. - exact: (realFieldBar_locally'_filter +oo). -exact: (realFieldBar_locally'_filter -oo). + exact: (ereal_locally'_filter +oo). +exact: (ereal_locally'_filter -oo). Qed. -Typeclasses Opaque realFieldBar_locally. +Typeclasses Opaque ereal_locally. -Lemma near_pinfty_div2 (A : set K) : +Lemma near_pinfty_div2 (A : set R) : (\forall k \near +oo, A k) -> (\forall k \near +oo, A (k / 2)). Proof. by move=> [M AM]; exists (M * 2) => x; rewrite -ltr_pdivl_mulr //; apply: AM. @@ -738,12 +730,8 @@ Proof. by exists c. Qed. Lemma locally_pinfty_ge c : \forall x \near +oo, c <= x. Proof. by exists c; apply: ltrW. Qed. -End RealFieldBar. +End ExtendedReals. -Notation "'+oo'" := p_infty : real_scope. -Notation "'-oo'" := m_infty : real_scope. -Notation "'+oo' 'in' K" := (@p_infty K) (at level 0) : real_scope. -Notation "'-oo' 'in' K" := (@m_infty K) (at level 0) : real_scope. Hint Extern 0 (is_true (0 < _)) => match goal with H : ?x \is_near (locally +oo) |- _ => solve[near: x; exists 0 => _/posnumP[x] //] end : core. diff --git a/reals.v b/reals.v index 158da72c5c..9cf17b445a 100644 --- a/reals.v +++ b/reals.v @@ -736,7 +736,7 @@ End Sup. (* -------------------------------------------------------------------- *) Section ExtendedReals. -Variable (R : realType). +Variable (R : numDomainType). Inductive er := ERFin of R | ERPInf | ERNInf. @@ -744,8 +744,10 @@ Coercion real_of_er (x : er) := if x is ERFin v then v else 0. End ExtendedReals. -Notation "\+inf" := (@ERPInf _). -Notation "\-inf" := (@ERNInf _). +Notation "'+oo'" := (@ERPInf _). +Notation "'-oo'" := (@ERNInf _). +Notation "'+oo' 'in' R" := (@ERPInf R) (at level 0). +Notation "'-oo' 'in' R" := (@ERNInf R) (at level 0). Notation "x %:E" := (@ERFin _ x) (at level 2, format "x %:E"). Notation "{ 'ereal' R }" := (er R) (format "{ 'ereal' R }"). @@ -755,20 +757,20 @@ Delimit Scope ereal_scope with E. (* -------------------------------------------------------------------- *) Section ERealCode. -Variable (R : realType). +Variable (R : numDomainType). Definition code (x : {ereal R}) := match x with | x%:E => GenTree.Node 0 [:: GenTree.Leaf x] - | \+inf => GenTree.Node 1 [::] - | \-inf => GenTree.Node 2 [::] + | +oo => GenTree.Node 1 [::] + | -oo => GenTree.Node 2 [::] end. Definition decode (x : GenTree.tree R) : option {ereal R} := match x with | GenTree.Node 0 [:: GenTree.Leaf x] => Some x%:E - | GenTree.Node 1 [::] => Some \+inf - | GenTree.Node 2 [::] => Some \-inf + | GenTree.Node 1 [::] => Some +oo + | GenTree.Node 2 [::] => Some -oo | _ => None end. @@ -781,27 +783,27 @@ Definition ereal_choiceMixin := PcanChoiceMixin codeK. Canonical ereal_choiceType := ChoiceType {ereal R} ereal_choiceMixin. End ERealCode. -Lemma eqe {R : realType} (x1 x2 : R) : +Lemma eqe {R : numDomainType} (x1 x2 : R) : (x1%:E == x2%:E) = (x1 == x2). Proof. by apply/eqP/eqP=> [[]|->]. Qed. (* -------------------------------------------------------------------- *) Section ERealOrder. -Context {R : realType}. +Context {R : numDomainType}. Definition lee (x1 x2 : {ereal R}) := match x1, x2 with - | \-inf, _ | _, \+inf => true - | \+inf, _ | _, \-inf => false + | -oo, _ | _, +oo => true + | +oo, _ | _, -oo => false | x1%:E, x2%:E => (x1 <= x2) end. Definition lte (x1 x2 : {ereal R}) := match x1, x2 with - | \-inf, \-inf | \+inf, \+inf => false - | \-inf, _ | _ , \+inf => true - | \+inf, _ | _ , \-inf => false + | -oo, -oo | +oo, +oo => false + | -oo, _ | _ , +oo => true + | +oo, _ | _ , -oo => false | x1%:E, x2%:E => (x1 < x2) end. @@ -817,24 +819,24 @@ Notation "x < y < z" := ((lte x y) && (lte y z)) : ereal_scope. (* -------------------------------------------------------------------- *) Section ERealArith. -Context {R : realType}. +Context {R : numDomainType}. Implicit Types (x y z : {ereal R}). Definition eadd x y := match x, y with | x%:E , y%:E => (x + y)%:E - | \-inf, _ => \-inf - | _ , \-inf => \-inf - | \+inf, _ => \+inf - | _ , \+inf => \+inf + | -oo, _ => -oo + | _ , -oo => -oo + | +oo, _ => +oo + | _ , +oo => +oo end. Definition eopp x := match x with | x%:E => (-x)%:E - | \-inf => \+inf - | \+inf => \-inf + | -oo => +oo + | +oo => -oo end. End ERealArith. @@ -873,7 +875,7 @@ Local Open Scope ereal_scope. (* -------------------------------------------------------------------- *) Section ERealArithTh. -Context {R : realType}. +Context {R : numDomainType}. Implicit Types (x y z : {ereal R}). @@ -901,7 +903,7 @@ End ERealArithTh. (* -------------------------------------------------------------------- *) Section ERealOrderTheory. -Context {R : realType}. +Context {R : numDomainType}. Local Open Scope ereal_scope. @@ -931,10 +933,10 @@ Proof. by elift ltrW: x y. Qed. Lemma eqe_le x y : (x == y) = (x <= y <= x). Proof. by elift eqr_le: x y. Qed. -Lemma leeNgt x y : (x <= y) = ~~ (y < x). +Lemma leeNgt (K : realDomainType) (x y : {ereal K}) : (x <= y) = ~~ (y < x). Proof. by elift lerNgt: x y. Qed. -Lemma lteNgt x y : (x < y) = ~~ (y <= x). +Lemma lteNgt (K : realDomainType) (x y : {ereal K}) : (x < y) = ~~ (y <= x). Proof. by elift ltrNge: x y. Qed. Lemma lee_eqVlt x y : (x <= y) = ((x == y) || (x < y)). From dc99671d4d898cf766e7dbca7d4948def17bb6ed Mon Sep 17 00:00:00 2001 From: damien rouhling Date: Thu, 4 Oct 2018 13:51:41 +0200 Subject: [PATCH 04/11] Reintroduce lemmas about R that were commented out + translate compatibility lemmas about derivative from Coquelicot --- hierarchy.v | 289 +++++++++++++++++++----------------------- misc/Rcompatibility.v | 118 +++++++++++++++++ 2 files changed, 249 insertions(+), 158 deletions(-) create mode 100644 misc/Rcompatibility.v diff --git a/hierarchy.v b/hierarchy.v index 30ea087b21..29cb6181ab 100644 --- a/hierarchy.v +++ b/hierarchy.v @@ -1180,7 +1180,7 @@ Lemma ler_distW (R : realDomainType) (x y e : R): (`|x - y|%R <= e) -> y - e <= x. Proof. by rewrite ler_distl => /andP[]. Qed. -Section real_topology. +Section realComplete. Context {R : realType}. @@ -1216,6 +1216,12 @@ Grab Existential Variables. all: end_near. Qed. Canonical real_completeType := CompleteType R real_complete. Canonical real_completeNormedModType := [completeNormedModType R of R]. +End realComplete. + +Section real_topology. + +Variable (R : realFieldType). + Definition at_left x := within (fun u : R => u < x) (locally x). Definition at_right x := within (fun u : R => x < u) (locally x). (* :TODO: We should have filter notation ^- and ^+ for these *) @@ -1318,7 +1324,59 @@ move=> [y [[z Az oppzey] [t Bt opptey]]]; exists (- y). by split; [rewrite -oppzey opprK|rewrite -opptey opprK]. Qed. -Lemma segment_connected (a b : R) : connected [set x | x \in `[a, b]]. +Lemma segment_closed (a b : R) : closed [set x | x \in `[a, b]]. +Proof. +have -> : [set x | x \in `[a, b]] = (>= a) `&` (<= b). + by rewrite predeqE => ?; rewrite inE; split=> [/andP [] | /= [->]]. +exact: closedI (@closed_ge _) (@closed_le _). +Qed. + +Lemma ler0_addgt0P (x : R) : reflect (forall e, e > 0 -> x <= e) (x <= 0). +Proof. +apply: (iffP idP) => [lex0 e egt0|lex0]. + by apply: ler_trans lex0 _; apply: ltrW. +case: (lerP x 0) => // lt0x. +have /midf_lt [_] := lt0x; rewrite ltrNge -eqbF_neg => /eqP<-. +by rewrite add0r; apply: lex0; rewrite -[x]/((PosNum lt0x)%:num). +Qed. + +(** Local properties in [R] *) + +Lemma lte_locally (a b : {ereal R}) (x : R) : + (a < x%:E < b)%E -> exists e : {posnum R}, + forall y, `|y - x| < e -> (a < y%:E < b)%E. +Proof. +move=> /andP []; move=> [:wlog]; case: a b => [a||] [b||] //= ltax ltxb. +- move: a b ltax ltxb; abstract: wlog. (*BUG*) + move=> {a b} a b ltxa ltxb. + have m_gt0 : minr ((x - a) / 2) ((b - x) / 2) > 0. + by rewrite ltr_minr !divr_gt0 ?subr_gt0. + exists (PosNum m_gt0) => y //=; rewrite ltr_minr !ltr_distl. + move=> /andP[/andP[ay _] /andP[_ yb]]. + rewrite (ltr_trans _ ay) ?(ltr_trans yb) //=. + by rewrite -subr_gt0 opprD addrA {1}[b - x]splitr addrK divr_gt0 ?subr_gt0. + by rewrite -subr_gt0 addrAC {1}[x - a]splitr addrK divr_gt0 ?subr_gt0. +- have [//||d dP] := wlog a (x + 1); rewrite ?ltr_addl //. + by exists d => y /dP /andP[->]. +- have [//||d dP] := wlog (x - 1) b; rewrite ?gtr_addl ?ltrN10 //. + by exists d => y /dP /andP[_ ->]. +- by exists 1%:pos. +Qed. + +Lemma locally_interval (P : R -> Prop) (x : R) (a b : {ereal R}) : + (a < x%:E < b)%E -> (forall (y : R), (a < y%:E < b)%E -> P y) -> locally x P. +Proof. +move=> axb Pab; have /lte_locally [d dab] := axb; apply/locally_normP. +by exists d%:num => // y /=; rewrite normmB => /dab; apply: Pab. +Qed. + +End real_topology. + +Hint Resolve open_lt. +Hint Resolve open_gt. + +Lemma segment_connected (R : realType) (a b : R) : + connected [set x | x \in `[a, b]]. Proof. move=> A [y Ay] Aop Acl. move: Aop; apply: contrapTT; rewrite predeqE => /asboolPn /existsp_asboolPn [x]. @@ -1396,14 +1454,8 @@ apply: ler_trans; rewrite -ler_subr_addl ler_minl; apply/orP; right. by rewrite ler_pdivr_mulr // mulrDr mulr1 ler_addl; apply: ltrW. Qed. -Lemma segment_closed (a b : R) : closed [set x | x \in `[a, b]]. -Proof. -have -> : [set x | x \in `[a, b]] = (>= a) `&` (<= b). - by rewrite predeqE => ?; rewrite inE; split=> [/andP [] | /= [->]]. -exact: closedI (@closed_ge _) (@closed_le _). -Qed. - -Lemma segment_compact (a b : R) : compact [set x | x \in `[a, b]]. +Lemma segment_compact (R : realType) (a b : R) : + compact [set x | x \in `[a, b]]. Proof. case: (lerP a b) => [leab|ltba]; last first. by move=> F FF /filter_ex [x abx]; move: ltba; rewrite (itvP abx). @@ -1453,17 +1505,7 @@ rewrite ltr_subl_addl -ltr_subl_addr; apply: ltr_trans ltyz. by apply: ltr_distW; rewrite distrC. Qed. -Lemma ler0_addgt0P (K : realFieldType) (x : K) : - reflect (forall e, e > 0 -> x <= e) (x <= 0). -Proof. -apply: (iffP idP) => [lex0 e egt0|lex0]. - by apply: ler_trans lex0 _; apply: ltrW. -case: (lerP x 0) => // lt0x. -have /midf_lt [_] := lt0x; rewrite ltrNge -eqbF_neg => /eqP<-. -by rewrite add0r; apply: lex0; rewrite -[x]/((PosNum lt0x)%:num). -Qed. - -Lemma IVT (f : R -> R) (a b v : R) : +Lemma IVT (R : realType) (f : R -> R) (a b v : R) : a <= b -> {in `[a, b], continuous f} -> minr (f a) (f b) <= v <= maxr (f a) (f b) -> exists2 c, c \in `[a, b] & f c = v. @@ -1523,23 +1565,6 @@ move: ltsupt; rewrite ltrNge => /negP; apply; apply: sup_upper_bound => //. by rewrite inE letb leftv. Grab Existential Variables. all: end_near. Qed. -(** Local properties in [R] *) - -(* Lemma locally_interval (P : R -> Prop) (x : R) (a b : Rbar) : *) -(* Rbar_lt a x -> Rbar_lt x b -> *) -(* (forall (y : R), Rbar_lt a y -> Rbar_lt y b -> P y) -> *) -(* locally x P. *) -(* Proof. *) -(* move => Hax Hxb Hp; case: (Rbar_lt_locally _ _ _ Hax Hxb) => d Hd. *) -(* exists d%:num => //= y; rewrite /AbsRing_ball /= absrB. *) -(* by move=> /Hd /andP[??]; apply: Hp. *) -(* Qed. *) - -End real_topology. - -Hint Resolve open_lt. -Hint Resolve open_gt. - (** * Topology on [R]² *) (* Lemma locally_2d_align : *) @@ -1785,123 +1810,86 @@ rewrite -[X in X <= _](nth_index 0 vinseq); apply: bigmaxr_ler. by rewrite index_mem. Qed. -(** Open sets in [Rbar] *) +(** Open sets in extended reals *) -(* Lemma open_Rbar_lt y : open (fun u : R => Rbar_lt u y). *) -(* Proof. *) -(* case: y => [y||] /=. *) -(* exact: open_lt. *) -(* by rewrite trueE; apply: openT. *) -(* by rewrite falseE; apply: open0. *) -(* Qed. *) +Section ErealOpen. -(* Lemma open_Rbar_gt y : open (fun u : R => Rbar_lt y u). *) -(* Proof. *) -(* case: y => [y||] /=. *) -(* exact: open_gt. *) -(* by rewrite falseE; apply: open0. *) -(* by rewrite trueE; apply: openT. *) -(* Qed. *) +Variable (R : realFieldType). -(* Lemma open_Rbar_lt' x y : Rbar_lt x y -> Rbar_locally x (fun u => Rbar_lt u y). *) -(* Proof. *) -(* case: x => [x|//|] xy; first exact: open_Rbar_lt. *) -(* case: y => [y||//] /= in xy *. *) -(* exists y => /= x ? //. *) -(* by exists 0. *) -(* Qed. *) +Local Open Scope ereal_scope. -(* Lemma open_Rbar_gt' x y : Rbar_lt y x -> Rbar_locally x (fun u => Rbar_lt y u). *) -(* Proof. *) -(* case: x => [x||] //=; do ?[exact: open_Rbar_gt]; *) -(* case: y => [y||] //=; do ?by exists 0. *) -(* by exists y => x yx //=. *) -(* Qed. *) +Lemma open_lte y : open (fun u : R => u%:E < y). +Proof. +by case: y => [y||] //=; [apply: open_lt|rewrite trueE; apply: openT]. +Qed. -(* Lemma Rbar_locally'_le x : Rbar_locally' x --> Rbar_locally x. *) -(* Proof. *) -(* move: x => [x P [_/posnumP[e] HP] |x P|x P] //=. *) -(* by exists e%:num => // ???; apply: HP. *) -(* Qed. *) +Lemma open_gte y : open (fun u : R => y < u%:E). +Proof. +by case: y => [y||] //=; [apply: open_gt|rewrite trueE; apply: openT]. +Qed. -(* Lemma Rbar_locally'_le_finite (x : R) : Rbar_locally' x --> locally x. *) -(* Proof. *) -(* by move=> P [_/posnumP[e] HP] //=; exists e%:num => // ???; apply: HP. *) -(* Qed. *) +Lemma open_lte' x y : x < y -> locally x (fun u : R => u%:E < y). +Proof. +by case: x y => [x||] [y||] //= xy; + [apply: open_lt|apply/locally_normP; exists 1|exists y|exists 0]. +Qed. -(** * Some limits on real functions *) +Lemma open_gte' x y : y < x -> locally x (fun u : R => y < u%:E). +Proof. +by case: x y => [x||] [y||] //= yx; + [apply: open_gt|apply/locally_normP; exists 1|exists y|exists 0]. +Qed. -(* Definition Rbar_loc_seq (x : Rbar) (n : nat) := match x with *) -(* | Finite x => x + (INR n + 1)^-1 *) -(* | +oo => INR n *) -(* | -oo => - INR n *) -(* end. *) +Lemma flim_locally'e_locallye (x : {ereal R}) : ereal_locally' x --> locally x. +Proof. +by move: x => [||] x P //=; rewrite locally_simpl /locally /= locallyE' => -[]. +Qed. -(* Lemma flim_Rbar_loc_seq x : Rbar_loc_seq x --> Rbar_locally' x. *) -(* Proof. *) -(* move=> P; rewrite /Rbar_loc_seq. *) -(* case: x => /= [x [_/posnumP[delta] Hp] |[delta Hp] |[delta Hp]]; last 2 first. *) -(* have /ZnatP [N Nfloor] : ifloor (maxr delta 0) \is a Znat. *) -(* by rewrite Znat_def ifloor_ge0 ler_maxr lerr orbC. *) -(* exists N.+1 => // n ltNn; apply: Hp. *) -(* have /ler_lt_trans : delta <= maxr delta 0 by rewrite ler_maxr lerr. *) -(* apply; apply: ltr_le_trans (floorS_gtr _) _; rewrite floorE Nfloor. *) -(* by rewrite -(@natrD [ringType of R] N 1) INRE ler_nat addn1. *) -(* have /ZnatP [N Nfloor] : ifloor (maxr (- delta) 0) \is a Znat. *) -(* by rewrite Znat_def ifloor_ge0 ler_maxr lerr orbC. *) -(* exists N.+1 => // n ltNn; apply: Hp; rewrite ltr_oppl. *) -(* have /ler_lt_trans : - delta <= maxr (- delta) 0 by rewrite ler_maxr lerr. *) -(* apply; apply: ltr_le_trans (floorS_gtr _) _; rewrite floorE Nfloor. *) -(* by rewrite -(@natrD [ringType of R] N 1) INRE ler_nat addn1. *) -(* have /ZnatP [N Nfloor] : ifloor (delta%:num^-1) \is a Znat. *) -(* by rewrite Znat_def ifloor_ge0. *) -(* exists N => // n leNn; have gt0Sn : 0 < INR n + 1. *) -(* by apply: ltr_spaddr => //; apply/RleP/pos_INR. *) -(* apply: Hp; last first. *) -(* by rewrite eq_sym addrC -subr_eq subrr eq_sym; apply/invr_neq0/lt0r_neq0. *) -(* rewrite /AbsRing_ball /= opprD addrA subrr absrB subr0. *) -(* rewrite absRE gtr0_norm; last by rewrite invr_gt0. *) -(* rewrite -[X in X < _]mulr1 ltr_pdivr_mull // -ltr_pdivr_mulr // div1r. *) -(* apply: ltr_le_trans (floorS_gtr _) _; rewrite floorE Nfloor ler_add //. *) -(* by rewrite INRE ler_nat. *) -(* Qed. *) +Lemma flim_locally'e_locally (x : R) : ereal_locally' x%:E --> locally x. +Proof. +move=> P; rewrite locally_simpl => /locally_normP [_/posnumP[e] sP] /=. +by apply/locally_normP; exists e%:num => // ? /sP. +Qed. -(* TODO: express using ball?*) -(* Lemma continuity_pt_locally f x : continuity_pt f x <-> *) -(* forall eps : posreal, locally x (fun u => `|f u - f x| < eps). *) -(* Proof. *) -(* split=> [fcont e|fcont _/RltP/posnumP[e]]; last first. *) -(* have [_/posnumP[d] xd_fxe] := fcont e. *) -(* exists d%:num; split; first by apply/RltP; have := [gt0 of d%:num]. *) -(* by move=> y [_ /RltP yxd]; apply/RltP/xd_fxe; rewrite /AbsRing_ball /= absrB. *) -(* have /RltP egt0 := [gt0 of e%:num]. *) -(* have [_ [/RltP/posnumP[d] dx_fxe]] := fcont e%:num egt0. *) -(* exists d%:num => // y xyd; case: (eqVneq x y) => [->|xney]. *) -(* by rewrite subrr absr0. *) -(* apply/RltP/dx_fxe; split; first by split=> //; apply/eqP. *) -(* by have /RltP := xyd; rewrite absrB. *) -(* Qed. *) +End ErealOpen. -(* Lemma continuity_pt_flim (f : R -> R) (x : R) : *) -(* continuity_pt f x <-> {for x, continuous f}. *) -(* Proof. *) -(* eapply iff_trans; first exact: continuity_pt_locally. *) -(* apply iff_sym. *) -(* have FF : Filter (f @ x) by typeclasses eauto. *) -(* case: (@flim_ballP _ (f @ x) FF (f x)) => {FF}H1 H2. *) -(* (* TODO: in need for lemmas and/or refactoring of already existing lemmas (ball vs. Rabs) *) *) -(* split => [{H2} /H1{H1} H1 eps|{H1} H]. *) -(* - have {H1} [//|_/posnumP[x0] Hx0] := H1 eps%:num. *) -(* exists x0%:num => // Hx0' /Hx0 /=. *) -(* by rewrite ball_absE /= absrB; apply. *) -(* - apply H2 => _ /posnumP[eps]; move: (H eps) => {H} [_ /posnumP[x0] Hx0]. *) -(* exists x0%:num => // y /Hx0 /= {Hx0}Hx0. *) -(* by rewrite ball_absE /= absrB. *) -(* Qed. *) +(** * Some limits on real functions *) -(* Lemma continuity_ptE (f : R -> R) (x : R) : *) -(* continuity_pt f x <-> {for x, continuous f}. *) -(* Proof. exact: continuity_pt_flim. Qed. *) +Definition ereal_loc_seq (R : numDomainType) (x : {ereal R}) (n : nat) := + match x with + | x%:E => x + (n%:R + 1)^-1 + | +oo => n%:R + | -oo => - n%:R + end. + +Lemma flim_ereal_loc_seq (R : realType) (x : {ereal R}) : + ereal_loc_seq x --> ereal_locally' x. +Proof. +move=> P; rewrite /ereal_loc_seq; +case: x=> [x /(locally_normP x)[_/posnumP[d] sP] |[d sP] |[d sP]]; last 2 first. + have /ZnatP [N Nfloor] : ifloor (maxr d 0) \is a Znat. + by rewrite Znat_def ifloor_ge0 ler_maxr lerr orbC. + exists N.+1 => // n ltNn; apply: sP. + have /ler_lt_trans : d <= maxr d 0 by rewrite ler_maxr lerr. + apply; apply: ltr_le_trans (floorS_gtr _) _; rewrite floorE Nfloor. + by rewrite -(@natrD [ringType of R] N 1) ler_nat addn1. + have /ZnatP [N Nfloor] : ifloor (maxr (- d) 0) \is a Znat. + by rewrite Znat_def ifloor_ge0 ler_maxr lerr orbC. + exists N.+1 => // n ltNn; apply: sP; rewrite ltr_oppl. + have /ler_lt_trans : - d <= maxr (- d) 0 by rewrite ler_maxr lerr. + apply; apply: ltr_le_trans (floorS_gtr _) _; rewrite floorE Nfloor. + by rewrite -(@natrD [ringType of R] N 1) ler_nat addn1. +have /ZnatP [N Nfloor] : ifloor (d%:num^-1) \is a Znat. + by rewrite Znat_def ifloor_ge0. +exists N => // n leNn; have gt0Sn : 0 < (n%:R : R) + 1. + by apply/ltr_spaddr => //; rewrite ler0n. +apply: sP; last first. + by rewrite eq_sym addrC -subr_eq subrr eq_sym; apply/invr_neq0/lt0r_neq0. +rewrite /= opprD addrA subrr normmB subr0. +rewrite [ `|[_]|]gtr0_norm; last by rewrite invr_gt0. +rewrite -[X in X < _]mulr1 ltr_pdivr_mull // -ltr_pdivr_mulr // div1r. +by apply: ltr_le_trans (floorS_gtr _) _; rewrite floorE Nfloor ler_add ?ler_nat. +Qed. Lemma continuous_withinNx {U V : uniformType} (f : U -> V) x : @@ -1916,18 +1904,3 @@ rewrite !locally_nearE !near_map !near_locally in fxP *; have /= := cfx P fxP. rewrite !near_simpl near_withinE near_simpl => Pf; near=> y. by have [->|] := eqVneq y x; [by apply: locally_singleton|near: y]. Grab Existential Variables. all: end_near. Qed. - -(* Lemma continuity_pt_flim' f x : *) -(* continuity_pt f x <-> f @ locally' x --> f x. *) -(* Proof. by rewrite continuity_ptE continuous_withinNx. Qed. *) - -(* Lemma continuity_pt_locally' f x : *) -(* continuity_pt f x <-> *) -(* forall eps : R, 0 < eps -> locally' x (fun u => `|f x - f u| < eps)%R. *) -(* Proof. *) -(* by rewrite continuity_pt_flim' (@flim_normP _ [normedModType R of R^o]). *) -(* Qed. *) - -(* Lemma locally_pt_comp (P : R -> Prop) (f : R -> R) (x : R) : *) -(* locally (f x) P -> continuity_pt f x -> \near x, P (f x). *) -(* Proof. by move=> Lf /continuity_pt_flim; apply. Qed. *) diff --git a/misc/Rcompatibility.v b/misc/Rcompatibility.v new file mode 100644 index 0000000000..c8a30a0a66 --- /dev/null +++ b/misc/Rcompatibility.v @@ -0,0 +1,118 @@ +(* mathcomp analysis (c) 2017 Inria and AIST. License: CeCILL-C. *) +Require Import Reals. +From Coq Require Import ssreflect ssrfun ssrbool. +From mathcomp Require Import ssrnat eqtype choice fintype bigop ssralg ssrnum. +Require Import boolp reals Rstruct. +Require Import classical_sets posnum topology hierarchy landau derive. + +Set Implicit Arguments. +Unset Strict Implicit. +Unset Printing Implicit Defensive. +Import GRing.Theory Num.Def Num.Theory. + +(* code duplication that should disappear once Rstruct.v is reworked *) +Canonical R_filteredType := + [filteredType R of R for realField_filteredType R_realFieldType]. +Canonical R_topologicalType := + [topologicalType of R for realField_topologicalType R_realFieldType]. +Canonical R_uniformType := + [uniformType of R for realField_uniformType R_realFieldType]. +Canonical R_normedModType := + [normedModType R of R for realField_normedModType R_realFieldType]. + +Lemma continuity_pt_locally f x : continuity_pt f x <-> + forall eps : {posnum R}, locally x (fun u => `|f u - f x| < eps%:num). +Proof. +split=> [fcont e|fcont _/RltP/posnumP[e]]; last first. + have /locally_normP [_/posnumP[d] xd_fxe] := fcont e. + exists d%:num; split; first by apply/RltP; have := [gt0 of d%:num]. + by move=> y [_ /RltP yxd]; apply/RltP/xd_fxe; rewrite /= normmB. +apply/locally_normP; have /RltP egt0 := [gt0 of e%:num]. +have [_ [/RltP/posnumP[d] dx_fxe]] := fcont e%:num egt0. +exists d%:num => // y xyd; case: (eqVneq x y) => [->|xney]. + by rewrite subrr normr0. +apply/RltP/dx_fxe; split; first by split=> //; apply/eqP. +by have /RltP := xyd; rewrite normmB. +Qed. + +Local Open Scope classical_set_scope. + +Lemma continuity_pt_flim (f : R -> R) (x : R) : + continuity_pt f x <-> {for x, continuous f}. +Proof. +apply: iff_trans (continuity_pt_locally _ _) _; apply: iff_sym. +have FF : Filter (f @ x). +(* (* BUG: this should work *) *) +(* by typeclasses eauto. *) + by apply filtermap_filter; apply: @filter_filter' (locally_filter _). +apply: iff_trans (flim_normP (f x)) _; split=> [fx e|fx _/posnumP[e]]. + have /fx := [gt0 of e%:num]. + by rewrite !near_simpl; apply: filter_app; near=> y; rewrite distrC. +have := fx e; rewrite !near_simpl; apply: filter_app. +by near=> y; rewrite distrC. +Unshelve. all: end_near. Qed. + +Lemma continuity_ptE (f : R -> R) (x : R) : + continuity_pt f x <-> {for x, continuous f}. +Proof. exact: continuity_pt_flim. Qed. + +Lemma continuity_pt_flim' f x : + continuity_pt f x <-> f @ locally' x --> f x. +Proof. by rewrite continuity_ptE continuous_withinNx. Qed. + +Lemma continuity_pt_locally' f x : + continuity_pt f x <-> + forall eps : R, 0 < eps -> locally' x (fun u => `|f x - f u| < eps)%R. +Proof. +by rewrite continuity_pt_flim' (@flim_normP _ [normedModType R of R^o]). +Qed. + +Lemma locally_pt_comp (P : R -> Prop) (f : R -> R) (x : R) : + locally (f x) P -> continuity_pt f x -> \near x, P (f x). +Proof. by move=> Lf /continuity_pt_flim; apply. Qed. + +Lemma is_derive_Reals (f : R^o -> R^o) (x l : R) : + is_derive x 1 f l <-> derivable_pt_lim f x l. +Proof. +split=> f'l. + move=> _/RltP/posnumP[e]. + have /ex_derive/derivable_locallyP/eqaddoP/(_ _ [gt0 of e%:num / 2]) := f'l. + rewrite near_simpl -locally_nearE -filter_from_norm_locally. + move=> [d /RltP lt0d dfE]; exists (mkposreal _ lt0d) => h /eqP hn0 /RltP hd. + apply/RltP; have /= := dfE h; rewrite normmB subr0 => /(_ hd). + rewrite -[(_ - _ : _ -> _) _]/(_ - _) -[(_ + _ : _ -> _) _]/(_ + _) /cst. + rewrite derive_val opprD addrA -ler_pdivr_mulr ?normm_gt0 //. + rewrite mulrC -normfV -normmZ scalerBr scalerA mulVf // scale1r [_ *: _]mulrC. + rewrite [X in f (X + _)]mulr1 RminusE RdivE // RminusE RplusE [h + _]addrC. + by move=> /ler_lt_trans; apply; rewrite [X in _ < X]splitr ltr_addl. +have {f'l} f'l : + (fun h => h^-1 *: ((f (h + x) - f x))) @ locally' (0 : R) --> l. + move=> A /= /locally_normP [_/posnumP[e] sA]; apply/locally_normP. + have /RltP /f'l [[/= _/RltP/posnumP[d] df]] := [gt0 of e%:num]. + exists d%:num => // h /=; rewrite normmB subr0 => /RltP hd hn0. + have /eqP /df /(_ hd) /RltP := hn0. + rewrite RminusE RdivE // RminusE RplusE => dfx. + by apply: sA; rewrite /= normmB [_ *: _]mulrC [h + _]addrC. +have /flim_lim dfE := f'l; apply: DeriveDef; last by rewrite -derive1E . +apply/eqolimP/eqaddoP => _/posnumP[e]; near=> h => /=. +have -> : (fun h => h^-1 *: (f (h%:A + (x : R^o)) - f x)) = + (fun h => h^-1 *: (f (h + x) - f x)). + by rewrite funeqE => ?; rewrite [X in f (X + _)]mulr1. +rewrite -[(_ - _ : _ -> _) _]/(_ - _) {1}/cst /=. +have /eqolimP := f'l; rewrite funeqE => /(_ h) ->. +by rewrite dfE addrC addKr; near: h; case: e => /=; apply: littleoP. +Grab Existential Variables. end_near. Qed. + +Lemma derivable_Reals (f : R -> R) (x : R) : + derivable f x 1 -> derivable_pt f x. +Proof. by move=> /derivableP /is_derive_Reals; exists ('D_1 f x). Qed. + +Lemma Reals_derivable (f : R -> R) (x : R) : + derivable_pt f x -> derivable f x 1. +Proof. by move=> [l /is_derive_Reals]; apply: ex_derive. Qed. + +Lemma derive1_Reals (f : R -> R) (x : R) (pr : derivable_pt f x) : + derive_pt f x pr = f^`() x. +Proof. +by case: pr => /= l /is_derive_Reals /derive_val <-; rewrite derive1E. +Qed. From b644833418ef36d63354c44111b5f62cb521a61f Mon Sep 17 00:00:00 2001 From: damien rouhling Date: Mon, 8 Oct 2018 15:24:22 +0200 Subject: [PATCH 05/11] Add and use a theory of max using bigop --- hierarchy.v | 227 +++++++++++++++++++++++++++++++++++----------------- 1 file changed, 153 insertions(+), 74 deletions(-) diff --git a/hierarchy.v b/hierarchy.v index 29cb6181ab..9614d0cb70 100644 --- a/hierarchy.v +++ b/hierarchy.v @@ -2,7 +2,7 @@ From mathcomp Require Import ssreflect ssrfun ssrbool ssrnat eqtype choice. From mathcomp Require Import seq fintype bigop ssralg ssrint ssrnum finmap. From mathcomp Require Import matrix interval zmodp. -Require Import boolp reals Rstruct classical_sets posnum topology. +Require Import boolp reals classical_sets posnum topology. (******************************************************************************) (* This file extends the topological hierarchy with metric-related notions: *) @@ -748,6 +748,115 @@ by rewrite distrC ler_norm. Qed. Arguments flim_bounded {_ _ F FF}. +Section Bigmaxr. + +Variable (R : realDomainType). + +Lemma bigmaxr_mkcond I r (P : pred I) (F : I -> R) x : + \big[maxr/x]_(i <- r | P i) F i = + \big[maxr/x]_(i <- r) (if P i then F i else x). +Proof. +rewrite unlock; elim: r x => //= i r ihr x. +case P; rewrite ihr // maxr_r //; elim: r {ihr} => //= j r ihr. +by rewrite ler_maxr ihr orbT. +Qed. + +Lemma bigmaxr_split I r (P : pred I) (F1 F2 : I -> R) x : + \big[maxr/x]_(i <- r | P i) (maxr (F1 i) (F2 i)) = + maxr (\big[maxr/x]_(i <- r | P i) F1 i) (\big[maxr/x]_(i <- r | P i) F2 i). +Proof. +by elim/big_rec3: _ => [|i y z _ _ ->]; rewrite ?maxrr // maxrCA -!maxrA maxrCA. +Qed. + +Lemma filter_andb I r (a P : pred I) : + [seq i <- r | P i && a i] = [seq i <- [seq j <- r | P j] | a i]. +Proof. by elim: r => //= i r ->; case P. Qed. + +Lemma bigmaxr_idl I r (P : pred I) (F : I -> R) x : + \big[maxr/x]_(i <- r | P i) F i = maxr x (\big[maxr/x]_(i <- r | P i) F i). +Proof. +rewrite -big_filter; elim: [seq i <- r | P i] => [|i l ihl]. + by rewrite big_nil maxrr. +by rewrite big_cons maxrCA -ihl. +Qed. + +Lemma bigmaxrID I r (a P : pred I) (F : I -> R) x : + \big[maxr/x]_(i <- r | P i) F i = + maxr (\big[maxr/x]_(i <- r | P i && a i) F i) + (\big[maxr/x]_(i <- r | P i && ~~ a i) F i). +Proof. +rewrite -!(big_filter _ (fun _ => _ && _)) !filter_andb !big_filter. +rewrite ![in RHS](bigmaxr_mkcond _ _ F) !big_filter -bigmaxr_split. +have eqmax : forall i, P i -> + maxr (if a i then F i else x) (if ~~ a i then F i else x) = maxr (F i) x. + by move=> i _; case: (a i) => //=; rewrite maxrC. +rewrite [RHS](eq_bigr _ eqmax) -!(big_filter _ P). +elim: [seq j <- r | P j] => [|j l ihl]; first by rewrite !big_nil. +by rewrite !big_cons -maxrA -bigmaxr_idl ihl. +Qed. + +Lemma bigmaxr_seq1 I (i : I) (F : I -> R) x : + \big[maxr/x]_(j <- [:: i]) F j = maxr (F i) x. +Proof. by rewrite unlock /=. Qed. + +Lemma bigmaxr_pred1_eq (I : finType) (i : I) (F : I -> R) x : + \big[maxr/x]_(j | j == i) F j = maxr (F i) x. +Proof. by rewrite -big_filter filter_index_enum enum1 bigmaxr_seq1. Qed. + +Lemma bigmaxr_pred1 (I : finType) i (P : pred I) (F : I -> R) x : + P =1 pred1 i -> \big[maxr/x]_(j | P j) F j = maxr (F i) x. +Proof. by move/(eq_bigl _ _)->; apply: bigmaxr_pred1_eq. Qed. + +Lemma bigmaxrD1 (I : finType) j (P : pred I) (F : I -> R) x : + P j -> \big[maxr/x]_(i | P i) F i + = maxr (F j) (\big[maxr/x]_(i | P i && (i != j)) F i). +Proof. +move=> Pj; rewrite (bigmaxrID _ (pred1 j)) [in RHS]bigmaxr_idl maxrA. +by congr maxr; apply: bigmaxr_pred1 => i; rewrite /= andbC; case: eqP => //->. +Qed. + +Lemma ler_bigmaxr_cond (I : finType) (P : pred I) (F : I -> R) x i0 : + P i0 -> F i0 <= \big[maxr/x]_(i | P i) F i. +Proof. by move=> Pi0; rewrite (bigmaxrD1 _ _ Pi0) ler_maxr lerr. Qed. + +Lemma ler_bigmaxr (I : finType) (F : I -> R) (i0 : I) x : + F i0 <= \big[maxr/x]_i F i. +Proof. exact: ler_bigmaxr_cond. Qed. + +Lemma bigmaxr_lerP (I : finType) (P : pred I) m (F : I -> R) x : + reflect (x <= m /\ forall i, P i -> F i <= m) + (\big[maxr/x]_(i | P i) F i <= m). +Proof. +apply: (iffP idP) => [|[lexm leFm]]; last first. + by elim/big_ind: _ => // ??; rewrite ler_maxl =>->. +rewrite bigmaxr_idl ler_maxl => /andP[-> leFm]; split=> // i Pi. +by apply: ler_trans leFm; apply: ler_bigmaxr_cond. +Qed. + +Lemma bigmaxr_sup (I : finType) i0 (P : pred I) m (F : I -> R) x : + P i0 -> m <= F i0 -> m <= \big[maxr/x]_(i | P i) F i. +Proof. by move=> Pi0 ?; apply: ler_trans (ler_bigmaxr_cond _ _ Pi0). Qed. + +Lemma bigmaxr_ltrP (I : finType) (P : pred I) m (F : I -> R) x : + reflect (x < m /\ forall i, P i -> F i < m) + (\big[maxr/x]_(i | P i) F i < m). +Proof. +apply: (iffP idP) => [|[ltxm ltFm]]; last first. + by elim/big_ind: _ => // ??; rewrite ltr_maxl =>->. +rewrite bigmaxr_idl ltr_maxl => /andP[-> ltFm]; split=> // i Pi. +by apply: ler_lt_trans ltFm; apply: ler_bigmaxr_cond. +Qed. + +End Bigmaxr. + +Arguments bigmaxr_mkcond {R I r}. +Arguments bigmaxrID {R I r}. +Arguments bigmaxr_pred1 {R I} i {P F}. +Arguments bigmaxrD1 {R I} j {P F}. +Arguments ler_bigmaxr_cond {R I P F}. +Arguments ler_bigmaxr {R I F}. +Arguments bigmaxr_sup {R I} i0 {P m F}. + (** ** Matrices *) Section matrix_normedMod. @@ -755,76 +864,53 @@ Section matrix_normedMod. Variables (K : realFieldType) (m n : nat). (* take m.+1,n.+1 because entourages_normE is not provable for m = 0 or n = 0 *) -Definition mx_norm (x : 'M[K]_(m.+1, n.+1)) := - bigmaxr 0 [seq `|x ij.1 ij.2| | ij : 'I_m.+1 * 'I_n.+1]. +Definition mx_norm (x : 'M[K]_(m.+1, n.+1)) := \big[maxr/0]_ij `|x ij.1 ij.2|. Program Definition matrix_NormedModMixin := @NormedModMixin _ _ (@locally _ [filteredType 'M[K]_(m.+1, n.+1) of 'M[K]_(m.+1, n.+1)]) (Uniform.mixin (Uniform.class _)) mx_norm _ _ _ _. Next Obligation. -apply/bigmaxr_lerP=> [|i]; rewrite size_map -cardT mxvec_cast // => ltimn. -rewrite (nth_map (ord0, ord0)); last by rewrite -cardT mxvec_cast. -rewrite mxE; apply: ler_trans (ler_norm_add _ _) _. -do 2 ?[rewrite -(nth_map _ 0 (fun p => `|_ p.1 p.2|)) -?cardT ?mxvec_cast //]. -by apply: ler_add; apply: bigmaxr_ler; rewrite size_map -cardT mxvec_cast. +apply/bigmaxr_lerP; split. + by apply: addr_ge0; apply: (bigmaxr_sup (ord0, ord0)). +move=> ij _; rewrite mxE; apply: ler_trans (ler_norm_add _ _) _. +by apply: ler_add; apply: ler_bigmaxr. Qed. Next Obligation. -apply/bigmaxrP; split=> [|i]. - have : (0 < size [seq `|x ij.1 ij.2|%R | ij : 'I_m.+1 * 'I_n.+1])%N. - by rewrite size_map -cardT mxvec_cast. - move=> /bigmaxr_mem - /(_ 0) /mapP [ij ijP normx]; rewrite [mx_norm _]normx. - by apply/mapP; exists ij => //; rewrite mxE normrM. -rewrite size_map -cardT mxvec_cast // => ltimn. -rewrite (nth_map (ord0, ord0)); last by rewrite -cardT mxvec_cast. -rewrite mxE normrM ler_pmul // -(nth_map _ 0 (fun p => `|x p.1 p.2|)). - by apply: bigmaxr_ler; rewrite size_map -cardT mxvec_cast. -by rewrite -cardT mxvec_cast. +apply/eqP; rewrite eqr_le; apply/andP; split. + apply/bigmaxr_lerP; split. + by apply: mulr_ge0 => //; apply: (bigmaxr_sup (ord0, ord0)). + by move=> ij _; rewrite mxE normrM ler_wpmul2l //; apply: ler_bigmaxr. +case: (ler0P `|l|) => [|ln0]. + rewrite normr_le0 => /eqP->; rewrite scale0r normr0 mul0r. + by apply: (bigmaxr_sup (ord0, ord0)) => //=; rewrite mxE normr0. +rewrite -ler_pdivl_mull //; apply/bigmaxr_lerP; split. + by apply: mulr_ge0; rewrite ?invr_ge0 // (bigmaxr_sup (ord0, ord0)). +by move=> ij _; rewrite ler_pdivl_mull // -normrM (bigmaxr_sup ij) // mxE. Qed. Next Obligation. rewrite predeqE => A; split; last first. move=> [_/posnumP[e] sA]; exists (fun _ _ => [set p | `|p.1 - p.2| < e%:num]). by move=> _ _; exists e%:num. move=> [x y] /= xy; apply: sA => /=. - apply/bigmaxr_ltrP; rewrite size_map -cardT mxvec_cast // => i ltimn. - rewrite (nth_map (ord0, ord0)); last by rewrite -cardT mxvec_cast. - by rewrite !mxE; apply: xy. + by apply/bigmaxr_ltrP; split=> // ij _; rewrite !mxE. move=> [P entP sPA]; have {entP} entP : forall i j, exists e, 0 < e /\ [set pq | `|pq.1 - pq.2| < e] `<=` P i j. by move=> i j; have [e] := entP i j; exists e. set e := fun i j => get [set e | 0 < e /\ [set pq | `|pq.1 - pq.2| < e] `<=` P i j]. -exists (- bigmaxr 0 [seq - e ij.1 ij.2 | ij : 'I_m.+1 * 'I_n.+1]). - rewrite oppr_gt0; apply/bigmaxr_ltrP; rewrite size_map -cardT mxvec_cast //. - move=> i ltimn; rewrite (nth_map (ord0, ord0)); last first. - by rewrite -cardT mxvec_cast. - by rewrite oppr_lt0; set j := _.1; set k := _.2; have /getPex[] := entP j k. -move=> [x y] /= xy; apply: sPA => i j /=. -have s0 : (0 < size [seq `|(x - y) ij.1 ij.2|%R | ij : 'I_m.+1 * 'I_n.+1])%N. - by rewrite size_map -cardT mxvec_cast. -have memij : (i, j) \in enum [finType of 'I_m.+1 * 'I_n.+1]. - by rewrite mem_enum. +exists (- \big[maxr/-1]_ij - e ij.1 ij.2). + rewrite oppr_gt0; apply/bigmaxr_ltrP; split; first by rewrite oppr_lt0. + by move=> ij _; rewrite oppr_lt0; have /getPex [] := entP ij.1 ij.2. +move=> [x y] /= /bigmaxr_ltrP [_ xy]; apply: sPA => i j /=. have /getPex [_] := entP i j; apply=> /=; rewrite -[get _]/(e i j). -move: xy => /bigmaxr_ltrP - /(_ s0 (index (i, j) - (enum [finType of 'I_m.+1 * 'I_n.+1]))). -rewrite size_map index_mem => /(_ memij). -rewrite (nth_map (ord0, ord0)); last by rewrite index_mem. -rewrite !nth_index //= !mxE => /ltr_le_trans; apply. -rewrite ler_oppl; have /(nth_index 0) <- : - - e i j \in [seq - e ij.1 ij.2 | ij : 'I_m.+1 * 'I_n.+1]. - by apply/mapP; exists (i, j). -by apply: bigmaxr_ler; rewrite index_mem; apply/mapP; exists (i, j). +have /ltr_le_trans := xy (i,j) erefl; rewrite !mxE; apply. +by rewrite ler_oppl -[i]/(i,j).1 -{2}[j]/(i,j).2; apply: ler_bigmaxr. Qed. Next Obligation. -apply/matrixP => i j; rewrite mxE; apply/eqP. -rewrite -normr_eq0 eqr_le; apply/andP; split; last exact: normr_ge0. -have /(bigmaxr_ler 0) : - (index (i, j) (enum [finType of 'I_m.+1 * 'I_n.+1]) < - size [seq (`|x ij.1 ij.2|)%R | ij : 'I_m.+1 * 'I_n.+1])%N. - by rewrite size_map index_mem mem_enum. -rewrite -{3}H; apply: ler_trans. -rewrite (nth_map (ord0, ord0)); last by rewrite index_mem mem_enum. -by rewrite nth_index // mem_enum. +have /eqP := H; rewrite eqr_le => /andP [/bigmaxr_lerP [_ x0] _]. +apply/matrixP => i j; rewrite mxE; apply/eqP; rewrite -normr_le0. +exact: (x0 (i,j)). Qed. Canonical matrix_normedModType := @@ -836,14 +922,8 @@ Lemma coord_continuous {R : realFieldType} m n i j : continuous (fun M : 'M[R]_(m.+1, n.+1) => M i j). Proof. move=> M A /= /locally_normP [_/posnumP[e] sA]; apply/locally_normP. -exists e%:num => // N MN; apply/sA; have /bigmaxr_ltrP MeN := MN. -have /(MeN _) : - (index (i, j) (enum [finType of 'I_m.+1 * 'I_n.+1]) < - size [seq (`|(M - N) ij.1 ij.2|)%R | ij : 'I_m.+1 * 'I_n.+1])%N. - by rewrite size_map index_mem mem_enum. -rewrite (nth_map (ord0, ord0)); last by rewrite index_mem mem_enum. -rewrite nth_index ?mem_enum // !mxE; apply. -by rewrite size_map -cardT mxvec_cast. +exists e%:num => // N MN; apply/sA; have /bigmaxr_ltrP [_ MeN] := MN. +by have:= (MeN (i,j)); rewrite !mxE; apply. Qed. (** ** Pairs *) @@ -1716,24 +1796,27 @@ Grab Existential Variables. all: end_near. Qed. Definition bounded (K : realFieldType) (V : normedModType K) (A : set V) := \forall M \near +oo, A `<=` [set x | `|[x]| < M]. +(* TODO: generalize in finmap *) +Lemma big_seq_fsetE (R : Type) (idx : R) (op : R -> R -> R) (I : choiceType) + (X : {fset I}) (P : pred I) (F : I -> R) : + \big[op/idx]_(i <- X | P i) F i = \big[op/idx]_(x : X | P (val x)) F (val x). +Proof. by rewrite enum_fsetE big_map enumT. Qed. + Lemma compact_bounded (K : realType) (V : normedModType K) (A : set V) : compact A -> bounded A. Proof. rewrite compact_cover => Aco. -have covA : A `<=` \bigcup_(n : int) [set p | `|[p]| < n%:~R]. +have : A `<=` \bigcup_(n : int) [set p | `|[p]| < n%:~R]. move=> p Ap; exists (ifloor `|[p]| + 1) => //. by rewrite rmorphD /= -floorE floorS_gtr. -have /Aco [] := covA. - move=> n _; rewrite openE => p; rewrite -subr_gt0 => ltpn. - apply/locally_normP; exists (n%:~R - `|[p]|) => // q. - rewrite /ball ltr_subr_addr normmB; apply: ler_lt_trans. - by rewrite -{1}(subrK p q) ler_normm_add. -move=> D _ DcovA. -exists (bigmaxr 0 [seq n%:~R | n <- enum_fset D]). -move=> x ltmaxx p /DcovA [n Dn /ltr_trans]; apply; apply: ler_lt_trans ltmaxx. -have ltin : (index n (enum_fset D) < size (enum_fset D))%N by rewrite index_mem. -rewrite -(nth_index 0 Dn) -(nth_map _ 0) //; apply: bigmaxr_ler. -by rewrite size_map. +move=> /Aco [n _|D _ DcovA]; last first. + exists (\big[maxr/0]_(n <- D) n%:~R) => x. + rewrite big_seq_fsetE => /bigmaxr_ltrP [_ ltx] p. + by move=> /DcovA [n Dn /ltr_trans]; apply; apply: ltx (FSetSub _) _. +rewrite openE => p; rewrite -subr_gt0 => ltpn. +apply/locally_normP; exists (n%:~R - `|[p]|) => // q. +rewrite /ball ltr_subr_addr normmB; apply: ler_lt_trans. +by rewrite -{1}(subrK p q) ler_normm_add. Qed. Lemma rV_compact (T : topologicalType) n (A : 'I_n.+1 -> set T) : @@ -1803,11 +1886,7 @@ have Mnco : compact by move=> _; apply: segment_compact. apply: subclosed_compact Acl Mnco _ => v /normAltM normvltM i. suff /ltrW : `|[v ord0 i]| < M + 1 by rewrite ler_norml. -apply: ler_lt_trans (normvltM _ _); last by rewrite ltr_addl. -have vinseq : `|v ord0 i| \in [seq `|v ij.1 ij.2| | ij : 'I_1 * 'I_n.+1]. - by apply/mapP; exists (ord0, i) => //=; rewrite mem_enum. -rewrite -[X in X <= _](nth_index 0 vinseq); apply: bigmaxr_ler. -by rewrite index_mem. +by apply: ler_lt_trans (ler_bigmaxr (_,_) _) (normvltM _ _); rewrite ltr_addl. Qed. (** Open sets in extended reals *) From 893bf8856e482939cb3f049af2fdbd0d8ab1666a Mon Sep 17 00:00:00 2001 From: damien rouhling Date: Tue, 9 Oct 2018 10:50:53 +0200 Subject: [PATCH 06/11] Add min using bigop and simplify proofs --- derive.v | 10 +-- hierarchy.v | 196 +++++++++++++++++++++++++++++++++++++++++++--------- 2 files changed, 168 insertions(+), 38 deletions(-) diff --git a/derive.v b/derive.v index a9abf0fe85..8ed53c9cbc 100644 --- a/derive.v +++ b/derive.v @@ -142,7 +142,7 @@ Proof. by move=> ?; apply: DiffDef. Qed. Section jacobian. -Definition jacobian n m (R : realFieldType) (f : 'rV[R]_n.+1 -> 'rV[R]_m.+1) +Definition jacobian n m (R : realFieldType) (f : 'rV[R]_n -> 'rV[R]_m) p := lin1_mx ('d f p). End jacobian. @@ -315,7 +315,7 @@ Section DifferentialR2. Variable (R : realFieldType) (V : normedModType R). -Lemma derivemxE m n (f : 'rV[R]_m.+1 -> 'rV[R]_n.+1) (a v : 'rV[R]_m.+1) : +Lemma derivemxE m n (f : 'rV[R]_m -> 'rV[R]_n) (a v : 'rV[R]_m) : differentiable f a -> 'D_ v f a = v *m jacobian f a. Proof. by move=> /deriveE->; rewrite /jacobian mul_rV_lin1. Qed. @@ -598,10 +598,10 @@ apply: DiffDef; first exact/linear_differentiable/scaler_continuous. by rewrite diff_lin //; apply: scaler_continuous. Qed. -Lemma differentiable_coord m n (M : 'M[R]_(m.+1, n.+1)) i j : - differentiable (fun N : 'M[R]_(m.+1, n.+1) => N i j : R^o) M. +Lemma differentiable_coord m n (M : 'M[R]_(m, n)) i j : + differentiable (fun N : 'M[R]_(m, n) => N i j : R^o) M. Proof. -have @f : {linear 'M[R]_(m.+1, n.+1) -> R^o}. +have @f : {linear 'M[R]_(m, n) -> R^o}. by exists (fun N : 'M[R]_(_, _) => N i j); eexists; move=> ? ?; rewrite !mxE. rewrite (_ : (fun _ => _) = f) //; exact/linear_differentiable/coord_continuous. Qed. diff --git a/hierarchy.v b/hierarchy.v index 9614d0cb70..c2f9fbcbb5 100644 --- a/hierarchy.v +++ b/hierarchy.v @@ -761,6 +761,20 @@ case P; rewrite ihr // maxr_r //; elim: r {ihr} => //= j r ihr. by rewrite ler_maxr ihr orbT. Qed. +Lemma bigminr_maxr I r (P : pred I) (F : I -> R) x : + \big[minr/x]_(i <- r | P i) F i = - \big[maxr/- x]_(i <- r | P i) - F i. +Proof. +by elim/big_rec2: _ => [|i y _ _ ->]; rewrite ?oppr_max opprK. +Qed. + +Lemma bigminr_mkcond I r (P : pred I) (F : I -> R) x : + \big[minr/x]_(i <- r | P i) F i = + \big[minr/x]_(i <- r) (if P i then F i else x). +Proof. +rewrite !bigminr_maxr bigmaxr_mkcond; congr (- _). +by apply: eq_bigr => i _; case P. +Qed. + Lemma bigmaxr_split I r (P : pred I) (F1 F2 : I -> R) x : \big[maxr/x]_(i <- r | P i) (maxr (F1 i) (F2 i)) = maxr (\big[maxr/x]_(i <- r | P i) F1 i) (\big[maxr/x]_(i <- r | P i) F2 i). @@ -768,6 +782,14 @@ Proof. by elim/big_rec3: _ => [|i y z _ _ ->]; rewrite ?maxrr // maxrCA -!maxrA maxrCA. Qed. +Lemma bigminr_split I r (P : pred I) (F1 F2 : I -> R) x : + \big[minr/x]_(i <- r | P i) (minr (F1 i) (F2 i)) = + minr (\big[minr/x]_(i <- r | P i) F1 i) (\big[minr/x]_(i <- r | P i) F2 i). +Proof. +rewrite !bigminr_maxr -oppr_max -bigmaxr_split; congr (- _). +by apply: eq_bigr => i _; rewrite oppr_min. +Qed. + Lemma filter_andb I r (a P : pred I) : [seq i <- r | P i && a i] = [seq i <- [seq j <- r | P j] | a i]. Proof. by elim: r => //= i r ->; case P. Qed. @@ -780,6 +802,10 @@ rewrite -big_filter; elim: [seq i <- r | P i] => [|i l ihl]. by rewrite big_cons maxrCA -ihl. Qed. +Lemma bigminr_idl I r (P : pred I) (F : I -> R) x : + \big[minr/x]_(i <- r | P i) F i = minr x (\big[minr/x]_(i <- r | P i) F i). +Proof. by rewrite !bigminr_maxr {1}bigmaxr_idl oppr_max opprK. Qed. + Lemma bigmaxrID I r (a P : pred I) (F : I -> R) x : \big[maxr/x]_(i <- r | P i) F i = maxr (\big[maxr/x]_(i <- r | P i && a i) F i) @@ -795,18 +821,36 @@ elim: [seq j <- r | P j] => [|j l ihl]; first by rewrite !big_nil. by rewrite !big_cons -maxrA -bigmaxr_idl ihl. Qed. +Lemma bigminrID I r (a P : pred I) (F : I -> R) x : + \big[minr/x]_(i <- r | P i) F i = + minr (\big[minr/x]_(i <- r | P i && a i) F i) + (\big[minr/x]_(i <- r | P i && ~~ a i) F i). +Proof. by rewrite !bigminr_maxr -oppr_max -bigmaxrID. Qed. + Lemma bigmaxr_seq1 I (i : I) (F : I -> R) x : \big[maxr/x]_(j <- [:: i]) F j = maxr (F i) x. Proof. by rewrite unlock /=. Qed. +Lemma bigminr_seq1 I (i : I) (F : I -> R) x : + \big[minr/x]_(j <- [:: i]) F j = minr (F i) x. +Proof. by rewrite unlock /=. Qed. + Lemma bigmaxr_pred1_eq (I : finType) (i : I) (F : I -> R) x : \big[maxr/x]_(j | j == i) F j = maxr (F i) x. Proof. by rewrite -big_filter filter_index_enum enum1 bigmaxr_seq1. Qed. +Lemma bigminr_pred1_eq (I : finType) (i : I) (F : I -> R) x : + \big[minr/x]_(j | j == i) F j = minr (F i) x. +Proof. by rewrite bigminr_maxr bigmaxr_pred1_eq oppr_max !opprK. Qed. + Lemma bigmaxr_pred1 (I : finType) i (P : pred I) (F : I -> R) x : P =1 pred1 i -> \big[maxr/x]_(j | P j) F j = maxr (F i) x. Proof. by move/(eq_bigl _ _)->; apply: bigmaxr_pred1_eq. Qed. +Lemma bigminr_pred1 (I : finType) i (P : pred I) (F : I -> R) x : + P =1 pred1 i -> \big[minr/x]_(j | P j) F j = minr (F i) x. +Proof. by move/(eq_bigl _ _)->; apply: bigminr_pred1_eq. Qed. + Lemma bigmaxrD1 (I : finType) j (P : pred I) (F : I -> R) x : P j -> \big[maxr/x]_(i | P i) F i = maxr (F j) (\big[maxr/x]_(i | P i && (i != j)) F i). @@ -815,14 +859,29 @@ move=> Pj; rewrite (bigmaxrID _ (pred1 j)) [in RHS]bigmaxr_idl maxrA. by congr maxr; apply: bigmaxr_pred1 => i; rewrite /= andbC; case: eqP => //->. Qed. +Lemma bigminrD1 (I : finType) j (P : pred I) (F : I -> R) x : + P j -> \big[minr/x]_(i | P i) F i + = minr (F j) (\big[minr/x]_(i | P i && (i != j)) F i). +Proof. +by move=> Pj; rewrite !bigminr_maxr (bigmaxrD1 _ _ Pj) oppr_max opprK. +Qed. + Lemma ler_bigmaxr_cond (I : finType) (P : pred I) (F : I -> R) x i0 : P i0 -> F i0 <= \big[maxr/x]_(i | P i) F i. Proof. by move=> Pi0; rewrite (bigmaxrD1 _ _ Pi0) ler_maxr lerr. Qed. +Lemma bigminr_ler_cond (I : finType) (P : pred I) (F : I -> R) x i0 : + P i0 -> \big[minr/x]_(i | P i) F i <= F i0. +Proof. by move=> Pi0; rewrite (bigminrD1 _ _ Pi0) ler_minl lerr. Qed. + Lemma ler_bigmaxr (I : finType) (F : I -> R) (i0 : I) x : F i0 <= \big[maxr/x]_i F i. Proof. exact: ler_bigmaxr_cond. Qed. +Lemma bigminr_ler (I : finType) (F : I -> R) (i0 : I) x : + \big[minr/x]_i F i <= F i0. +Proof. exact: bigminr_ler_cond. Qed. + Lemma bigmaxr_lerP (I : finType) (P : pred I) m (F : I -> R) x : reflect (x <= m /\ forall i, P i -> F i <= m) (\big[maxr/x]_(i | P i) F i <= m). @@ -833,10 +892,23 @@ rewrite bigmaxr_idl ler_maxl => /andP[-> leFm]; split=> // i Pi. by apply: ler_trans leFm; apply: ler_bigmaxr_cond. Qed. +Lemma bigminr_gerP (I : finType) (P : pred I) m (F : I -> R) x : + reflect (m <= x /\ forall i, P i -> m <= F i) + (m <= \big[minr/x]_(i | P i) F i). +Proof. +rewrite bigminr_maxr ler_oppr; apply: (iffP idP). + by move=> /bigmaxr_lerP [? lemF]; split=> [|??]; rewrite -ler_opp2 ?lemF. +by move=> [? lemF]; apply/bigmaxr_lerP; split=> [|??]; rewrite ler_opp2 ?lemF. +Qed. + Lemma bigmaxr_sup (I : finType) i0 (P : pred I) m (F : I -> R) x : P i0 -> m <= F i0 -> m <= \big[maxr/x]_(i | P i) F i. Proof. by move=> Pi0 ?; apply: ler_trans (ler_bigmaxr_cond _ _ Pi0). Qed. +Lemma bigminr_inf (I : finType) i0 (P : pred I) m (F : I -> R) x : + P i0 -> F i0 <= m -> \big[minr/x]_(i | P i) F i <= m. +Proof. by move=> Pi0 ?; apply: ler_trans (bigminr_ler_cond _ _ Pi0) _. Qed. + Lemma bigmaxr_ltrP (I : finType) (P : pred I) m (F : I -> R) x : reflect (x < m /\ forall i, P i -> F i < m) (\big[maxr/x]_(i | P i) F i < m). @@ -847,6 +919,63 @@ rewrite bigmaxr_idl ltr_maxl => /andP[-> ltFm]; split=> // i Pi. by apply: ler_lt_trans ltFm; apply: ler_bigmaxr_cond. Qed. +Lemma bigminr_gtrP (I : finType) (P : pred I) m (F : I -> R) x : + reflect (m < x /\ forall i, P i -> m < F i) + (m < \big[minr/x]_(i | P i) F i). +Proof. +rewrite bigminr_maxr ltr_oppr; apply: (iffP idP). + by move=> /bigmaxr_ltrP [? ltmF]; split=> [|??]; rewrite -ltr_opp2 ?ltmF. +by move=> [? ltmF]; apply/bigmaxr_ltrP; split=> [|??]; rewrite ltr_opp2 ?ltmF. +Qed. + +Lemma bigmaxr_gerP (I : finType) (P : pred I) m (F : I -> R) x : + reflect (m <= x \/ exists2 i, P i & m <= F i) + (m <= \big[maxr/x]_(i | P i) F i). +Proof. +apply: (iffP idP) => [|[lemx|[i Pi lemFi]]]; last 2 first. +- by rewrite bigmaxr_idl ler_maxr lemx. +- by rewrite (bigmaxrD1 _ _ Pi) ler_maxr lemFi. +rewrite lerNgt => /bigmaxr_ltrP /asboolPn. +rewrite asbool_and negb_and => /orP [/asboolPn/negP|/existsp_asboolPn [i]]. + by rewrite -lerNgt; left. +by move=> /asboolPn/imply_asboolPn [Pi /negP]; rewrite -lerNgt; right; exists i. +Qed. + +Lemma bigminr_lerP (I : finType) (P : pred I) m (F : I -> R) x : + reflect (x <= m \/ exists2 i, P i & F i <= m) + (\big[minr/x]_(i | P i) F i <= m). +Proof. +rewrite bigminr_maxr ler_oppl; apply: (iffP idP). + by move=> /bigmaxr_gerP [?|[i ??]]; [left|right; exists i => //]; + rewrite -ler_opp2. +by move=> [?|[i ??]]; apply/bigmaxr_gerP; [left|right; exists i => //]; + rewrite ler_opp2. +Qed. + +Lemma bigmaxr_gtrP (I : finType) (P : pred I) m (F : I -> R) x : + reflect (m < x \/ exists2 i, P i & m < F i) + (m < \big[maxr/x]_(i | P i) F i). +Proof. +apply: (iffP idP) => [|[ltmx|[i Pi ltmFi]]]; last 2 first. +- by rewrite bigmaxr_idl ltr_maxr ltmx. +- by rewrite (bigmaxrD1 _ _ Pi) ltr_maxr ltmFi. +rewrite ltrNge => /bigmaxr_lerP /asboolPn. +rewrite asbool_and negb_and => /orP [/asboolPn/negP|/existsp_asboolPn [i]]. + by rewrite -ltrNge; left. +by move=> /asboolPn/imply_asboolPn [Pi /negP]; rewrite -ltrNge; right; exists i. +Qed. + +Lemma bigminr_ltrP (I : finType) (P : pred I) m (F : I -> R) x : + reflect (x < m \/ exists2 i, P i & F i < m) + (\big[minr/x]_(i | P i) F i < m). +Proof. +rewrite bigminr_maxr ltr_oppl; apply: (iffP idP). + by move=> /bigmaxr_gtrP [?|[i ??]]; [left|right; exists i => //]; + rewrite -ltr_opp2. +by move=> [?|[i ??]]; apply/bigmaxr_gtrP; [left|right; exists i => //]; + rewrite ltr_opp2. +Qed. + End Bigmaxr. Arguments bigmaxr_mkcond {R I r}. @@ -856,6 +985,13 @@ Arguments bigmaxrD1 {R I} j {P F}. Arguments ler_bigmaxr_cond {R I P F}. Arguments ler_bigmaxr {R I F}. Arguments bigmaxr_sup {R I} i0 {P m F}. +Arguments bigminr_mkcond {R I r}. +Arguments bigminrID {R I r}. +Arguments bigminr_pred1 {R I} i {P F}. +Arguments bigminrD1 {R I} j {P F}. +Arguments bigminr_ler_cond {R I P F}. +Arguments bigminr_ler {R I F}. +Arguments bigminr_inf {R I} i0 {P m F}. (** ** Matrices *) @@ -863,29 +999,27 @@ Section matrix_normedMod. Variables (K : realFieldType) (m n : nat). -(* take m.+1,n.+1 because entourages_normE is not provable for m = 0 or n = 0 *) -Definition mx_norm (x : 'M[K]_(m.+1, n.+1)) := \big[maxr/0]_ij `|x ij.1 ij.2|. +Definition mx_norm (x : 'M[K]_(m, n)) := \big[maxr/0]_ij `|x ij.1 ij.2|. Program Definition matrix_NormedModMixin := @NormedModMixin _ _ - (@locally _ [filteredType 'M[K]_(m.+1, n.+1) of 'M[K]_(m.+1, n.+1)]) + (@locally _ [filteredType 'M[K]_(m, n) of 'M[K]_(m, n)]) (Uniform.mixin (Uniform.class _)) mx_norm _ _ _ _. Next Obligation. apply/bigmaxr_lerP; split. - by apply: addr_ge0; apply: (bigmaxr_sup (ord0, ord0)). + by apply: addr_ge0; apply/bigmaxr_gerP; left. move=> ij _; rewrite mxE; apply: ler_trans (ler_norm_add _ _) _. by apply: ler_add; apply: ler_bigmaxr. Qed. Next Obligation. apply/eqP; rewrite eqr_le; apply/andP; split. apply/bigmaxr_lerP; split. - by apply: mulr_ge0 => //; apply: (bigmaxr_sup (ord0, ord0)). + by apply: mulr_ge0 => //; apply/bigmaxr_gerP; left. by move=> ij _; rewrite mxE normrM ler_wpmul2l //; apply: ler_bigmaxr. -case: (ler0P `|l|) => [|ln0]. - rewrite normr_le0 => /eqP->; rewrite scale0r normr0 mul0r. - by apply: (bigmaxr_sup (ord0, ord0)) => //=; rewrite mxE normr0. -rewrite -ler_pdivl_mull //; apply/bigmaxr_lerP; split. - by apply: mulr_ge0; rewrite ?invr_ge0 // (bigmaxr_sup (ord0, ord0)). +case: (eqVneq l 0) => [->|]. + by rewrite scale0r normr0 mul0r; apply/bigmaxr_gerP; left. +rewrite -normr_gt0 => l0; rewrite -ler_pdivl_mull //; apply/bigmaxr_lerP; split. + by apply: mulr_ge0; rewrite ?invr_ge0 //; apply/bigmaxr_gerP; left. by move=> ij _; rewrite ler_pdivl_mull // -normrM (bigmaxr_sup ij) // mxE. Qed. Next Obligation. @@ -899,13 +1033,11 @@ move=> [P entP sPA]; have {entP} entP : forall i j, exists e, 0 < e /\ by move=> i j; have [e] := entP i j; exists e. set e := fun i j => get [set e | 0 < e /\ [set pq | `|pq.1 - pq.2| < e] `<=` P i j]. -exists (- \big[maxr/-1]_ij - e ij.1 ij.2). - rewrite oppr_gt0; apply/bigmaxr_ltrP; split; first by rewrite oppr_lt0. - by move=> ij _; rewrite oppr_lt0; have /getPex [] := entP ij.1 ij.2. -move=> [x y] /= /bigmaxr_ltrP [_ xy]; apply: sPA => i j /=. +exists (\big[minr/1]_ij e ij.1 ij.2). + by apply/bigminr_gtrP; split=> // ij _; have /getPex [] := entP ij.1 ij.2. +move=> [x y] /= /bigminr_gtrP [_ xy]; apply: sPA => i j /=. have /getPex [_] := entP i j; apply=> /=; rewrite -[get _]/(e i j). -have /ltr_le_trans := xy (i,j) erefl; rewrite !mxE; apply. -by rewrite ler_oppl -[i]/(i,j).1 -{2}[j]/(i,j).2; apply: ler_bigmaxr. +by have /bigmaxr_ltrP [_ /(_ (i,j))] := xy (i,j) erefl; rewrite !mxE; apply. Qed. Next Obligation. have /eqP := H; rewrite eqr_le => /andP [/bigmaxr_lerP [_ x0] _]. @@ -914,12 +1046,12 @@ exact: (x0 (i,j)). Qed. Canonical matrix_normedModType := - NormedModType K 'M[K]_(m.+1, n.+1) matrix_NormedModMixin. + NormedModType K 'M[K]_(m, n) matrix_NormedModMixin. End matrix_normedMod. Lemma coord_continuous {R : realFieldType} m n i j : - continuous (fun M : 'M[R]_(m.+1, n.+1) => M i j). + continuous (fun M : 'M[R]_(m, n) => M i j). Proof. move=> M A /= /locally_normP [_/posnumP[e] sA]; apply/locally_normP. exists e%:num => // N MN; apply/sA; have /bigmaxr_ltrP [_ MeN] := MN. @@ -1819,23 +1951,23 @@ rewrite /ball ltr_subr_addr normmB; apply: ler_lt_trans. by rewrite -{1}(subrK p q) ler_normm_add. Qed. -Lemma rV_compact (T : topologicalType) n (A : 'I_n.+1 -> set T) : +Lemma rV_compact (T : topologicalType) n (A : 'I_n -> set T) : (forall i, compact (A i)) -> - compact [ set v : 'rV[T]_n.+1 | forall i, A i (v ord0 i)]. + compact [ set v : 'rV[T]_n | forall i, A i (v ord0 i)]. Proof. move=> Aico. have : @compact (product_topologicalType _) [set f | forall i, A i (f i)]. by apply: tychonoff. move=> Aco F FF FA. -set G := [set [set f : 'I_n.+1 -> T | B (\row_j f j)] | B in F]. -have row_simpl (v : 'rV[T]_n.+1) : \row_j (v ord0 j) = v. +set G := [set [set f : 'I_n -> T | B (\row_j f j)] | B in F]. +have row_simpl (v : 'rV[T]_n) : \row_j (v ord0 j) = v. by apply/rowP => ?; rewrite mxE. -have row_simpl' (f : 'I_n.+1 -> T) : (\row_j f j) ord0 = f. +have row_simpl' (f : 'I_n -> T) : (\row_j f j) ord0 = f. by rewrite funeqE=> ?; rewrite mxE. have [f [Af clGf]] : [set f | forall i, A i (f i)] `&` @cluster (product_topologicalType _) G !=set0. suff GF : ProperFilter G. - apply: Aco; exists [set v : 'rV[T]_n.+1 | forall i, A i (v ord0 i)] => //. + apply: Aco; exists [set v : 'rV[T]_n | forall i, A i (v ord0 i)] => //. by rewrite predeqE => f; split => Af i; [have := Af i|]; rewrite row_simpl'. apply Build_ProperFilter. move=> _ [C FC <-]; have /filter_ex [v Cv] := FC. @@ -1843,7 +1975,7 @@ have [f [Af clGf]] : [set f | forall i, A i (f i)] `&` split. - by exists setT => //; apply: filterT. - by move=> _ _ [C FC <-] [D FD <-]; exists (C `&` D) => //; apply: filterI. - move=> C D sCD [E FE EeqC]; exists [set v : 'rV[T]_n.+1 | D (v ord0)]. + move=> C D sCD [E FE EeqC]; exists [set v : 'rV[T]_n | D (v ord0)]. by apply: filterS FE => v Ev; apply/sCD; rewrite -EeqC row_simpl. by rewrite predeqE => ?; rewrite row_simpl'. exists (\row_j f j); split; first by move=> i; rewrite mxE; apply: Af. @@ -1858,12 +1990,10 @@ move=> C D FC f_D; have {f_D} f_D : move=> g Pg; apply: sED => i j; rewrite ord1 row_simpl'. by have /getPex [_ /(_ _ (Pg j))] := exPj j. split; last by move=> j; have /getPex [[]] := exPj j. - exists [set [set g | forall j, get (Pj j) (g j)] | k in 'I_n.+1]; - last first. - rewrite predeqE => g; split; first by move=> [_ [_ _ <-]]. - move=> Pg; exists [set g | forall j, get (Pj j) (g j)] => //. - by exists ord0. - move=> _ [_ _ <-]; set s := [seq (@^~ j) @^-1` (get (Pj j)) | j : 'I_n.+1]. + exists [set [set g | forall j, get (Pj j) (g j)]]; last first. + rewrite predeqE => g; split=> [[? ->] | Pg] //. + by exists [set g | forall j, get (Pj j) (g j)]. + move=> _ ->; set s := [seq (@^~ j) @^-1` (get (Pj j)) | j : 'I_n]. exists [fset x in s]%fset. move=> B'; rewrite in_fset => /mapP [j _ ->]; rewrite inE. apply/asboolP; exists j => //; exists (get (Pj j)) => //. @@ -1876,12 +2006,12 @@ have GC : G [set g | C (\row_j g j)] by exists C. by have [g []] := clGf _ _ GC f_D; exists (\row_j (g j : T)). Qed. -Lemma bounded_closed_compact (R : realType) n (A : set 'rV[R]_n.+1) : +Lemma bounded_closed_compact (R : realType) n (A : set 'rV[R]_n) : bounded A -> closed A -> compact A. Proof. move=> [M normAltM] Acl. have Mnco : compact - [set v : 'rV[R]_n.+1 | (forall i, (v ord0 i) \in `[(- (M + 1)), (M + 1)])]. + [set v : 'rV[R]_n | (forall i, (v ord0 i) \in `[(- (M + 1)), (M + 1)])]. apply: (@rV_compact _ _ (fun _ => [set x | x \in `[(- (M + 1)), (M + 1)]])). by move=> _; apply: segment_compact. apply: subclosed_compact Acl Mnco _ => v /normAltM normvltM i. From a1834b287f26bc4fd14c278666642563d8867f22 Mon Sep 17 00:00:00 2001 From: damien rouhling Date: Tue, 9 Oct 2018 15:50:27 +0200 Subject: [PATCH 07/11] Move and extend Rstruct.v --- _CoqProject | 2 +- misc/Rcompatibility.v | 14 ++------------ Rstruct.v => misc/Rstruct.v | 18 +++++++++++++++++- misc/uniform_bigO.v | 13 ++----------- 4 files changed, 22 insertions(+), 25 deletions(-) rename Rstruct.v => misc/Rstruct.v (96%) diff --git a/_CoqProject b/_CoqProject index 218a9d0ccc..bfeab302aa 100644 --- a/_CoqProject +++ b/_CoqProject @@ -6,7 +6,6 @@ reals.v posnum.v landau.v classical_sets.v -Rstruct.v # Rbar.v topology.v hierarchy.v @@ -18,5 +17,6 @@ altreals/discrete.v altreals/realseq.v altreals/realsum.v altreals/distr.v +misc/Rstruct.v -R . mathcomp.analysis diff --git a/misc/Rcompatibility.v b/misc/Rcompatibility.v index c8a30a0a66..2d63afe0b2 100644 --- a/misc/Rcompatibility.v +++ b/misc/Rcompatibility.v @@ -2,24 +2,14 @@ Require Import Reals. From Coq Require Import ssreflect ssrfun ssrbool. From mathcomp Require Import ssrnat eqtype choice fintype bigop ssralg ssrnum. -Require Import boolp reals Rstruct. -Require Import classical_sets posnum topology hierarchy landau derive. +Require Import boolp reals. +Require Import classical_sets posnum topology hierarchy landau derive Rstruct. Set Implicit Arguments. Unset Strict Implicit. Unset Printing Implicit Defensive. Import GRing.Theory Num.Def Num.Theory. -(* code duplication that should disappear once Rstruct.v is reworked *) -Canonical R_filteredType := - [filteredType R of R for realField_filteredType R_realFieldType]. -Canonical R_topologicalType := - [topologicalType of R for realField_topologicalType R_realFieldType]. -Canonical R_uniformType := - [uniformType of R for realField_uniformType R_realFieldType]. -Canonical R_normedModType := - [normedModType R of R for realField_normedModType R_realFieldType]. - Lemma continuity_pt_locally f x : continuity_pt f x <-> forall eps : {posnum R}, locally x (fun u => `|f u - f x| < eps%:num). Proof. diff --git a/Rstruct.v b/misc/Rstruct.v similarity index 96% rename from Rstruct.v rename to misc/Rstruct.v index 90c1902443..4e5683ac6a 100644 --- a/Rstruct.v +++ b/misc/Rstruct.v @@ -27,6 +27,7 @@ Require Import Rtrigo1 Reals. From mathcomp Require Import ssreflect ssrfun ssrbool eqtype ssrnat seq. From mathcomp Require Import choice bigop ssralg fintype poly. From mathcomp Require Import mxpoly ssrnum finfun. +Require Import classical_sets topology hierarchy. Set Implicit Arguments. Unset Strict Implicit. @@ -345,7 +346,7 @@ case Hab: (a == b). by move=> _; rewrite (eqP Hab) eq_sym Hpb (ltrNge 0) /=; case/andP=> /ltrW ->. rewrite eq_sym Hpb /=; clear=> /RltbP Hab /andP [] /RltbP Hpa /RltbP Hpb. suff Hcp: continuity (fun x => (p.[x])%R). - have [z [[Hza Hzb] /eqP Hz2]]:= IVT _ a b Hcp Hab Hpa Hpb. + have [z [[Hza Hzb] /eqP Hz2]]:= Rsqrt_def.IVT _ a b Hcp Hab Hpa Hpb. by exists z=> //; apply/andP; split; apply/RlebP. rewrite -[p]coefK poly_def. set f := fun _ => _. @@ -632,3 +633,18 @@ Qed. End Max. End ssreal_struct_contd. + +Canonical R_pointedType := + [pointedType of R for zmod_pointedType R_zmodType]. +Canonical R_filteredType := + [filteredType R of R for realField_filteredType R_realFieldType]. +Canonical R_topologicalType := + [topologicalType of R for realField_topologicalType R_realFieldType]. +Canonical R_uniformType := + [uniformType of R for realField_uniformType R_realFieldType]. +Canonical R_completeType := + [completeType of R for @real_completeType real_realType]. +Canonical R_normedModType := + [normedModType R of R for realField_normedModType R_realFieldType]. +Canonical R_completeNormedModType := + [completeNormedModType R of R]. \ No newline at end of file diff --git a/misc/uniform_bigO.v b/misc/uniform_bigO.v index abbc0dc1c5..d26109a94f 100644 --- a/misc/uniform_bigO.v +++ b/misc/uniform_bigO.v @@ -2,8 +2,8 @@ Require Import Reals. From Coq Require Import ssreflect ssrfun ssrbool. From mathcomp Require Import ssrnat eqtype choice fintype bigop ssralg ssrnum. -Require Import boolp reals Rstruct. -Require Import classical_sets posnum topology hierarchy landau. +Require Import boolp reals. +Require Import classical_sets posnum topology hierarchy landau Rstruct. Set Implicit Arguments. Unset Strict Implicit. @@ -76,15 +76,6 @@ Qed. (* then we replace the epsilon/delta definition with bigO *) -Canonical R_filteredType := - [filteredType R of R for realField_filteredType R_realFieldType]. -Canonical R_topologicalType := - [topologicalType of R for realField_topologicalType R_realFieldType]. -Canonical R_uniformType := - [uniformType of R for realField_uniformType R_realFieldType]. -Canonical R_normedModType := - [normedModType R of R for realField_normedModType R_realFieldType]. - Definition OuO (f : A -> R * R -> R^o) (g : R * R -> R^o) := (fun x => f x.1 x.2) =O_ (filter_prod [set setT] (within P [filter of 0 : R^o * R^o])) (fun x => g x.2). From cb4ca5f5cffea73ad67420778fe498239006e885 Mon Sep 17 00:00:00 2001 From: damien rouhling Date: Wed, 10 Oct 2018 10:35:22 +0200 Subject: [PATCH 08/11] Clean up --- classical_sets.v | 3 ++ derive.v | 1 - landau.v | 3 +- topology.v | 99 +++++++++++++++++++++--------------------------- 4 files changed, 47 insertions(+), 59 deletions(-) diff --git a/classical_sets.v b/classical_sets.v index 95081770ba..d0cff7fdea 100644 --- a/classical_sets.v +++ b/classical_sets.v @@ -334,6 +334,9 @@ move=> [i Di]; rewrite predeqE => a; split=> [[Ifa Xa] j Dj|IfIXa]. by split=> [j /IfIXa [] | ] //; have /IfIXa [] := Di. Qed. +Lemma setMT A B : (@setT A) `*` (@setT B) = setT. +Proof. by rewrite predeqE. Qed. + Definition is_prop {A} (X : set A) := forall x y, X x -> X y -> x = y. Definition is_fun {A B} (f : A -> B -> Prop) := all (is_prop \o f). Definition is_total {A B} (f : A -> B -> Prop) := all (nonempty \o f). diff --git a/derive.v b/derive.v index 8ed53c9cbc..5e0325a15f 100644 --- a/derive.v +++ b/derive.v @@ -10,7 +10,6 @@ Unset Printing Implicit Defensive. Import GRing.Theory Num.Def Num.Theory. Local Open Scope classical_set_scope. -Local Open Scope fun_scope. Local Open Scope ring_scope. (******************************************************************************) diff --git a/landau.v b/landau.v index 4fbb79a720..34f7139b57 100644 --- a/landau.v +++ b/landau.v @@ -269,10 +269,9 @@ Reserved Notation "f '=Theta_' F h" format "f '=Theta_' F h"). Delimit Scope real_scope with real. -Local Open Scope classical_set_scope. Local Open Scope ring_scope. Local Open Scope real_scope. -Local Open Scope fun_scope. +Local Open Scope classical_set_scope. Section function_space. diff --git a/topology.v b/topology.v index c8048a6668..6c78deae5d 100644 --- a/topology.v +++ b/topology.v @@ -2189,22 +2189,19 @@ Definition connected (T : topologicalType) (A : set T) := (** * Uniform spaces defined using entourages *) -Definition comp_rel (R S T : Type) (A : set (R * S)) (B : set (S * T)) := - [set xy | exists2 z, A (xy.1, z) & B (z, xy.2)]. -Notation "B \o A" := (comp_rel A B) : classical_set_scope. +Local Notation "B \o A" := + ([set xy | exists2 z, A (xy.1, z) & B (z, xy.2)]) : classical_set_scope. -Definition inv_rel (T T' : Type) (A : set (T * T')) := - [set xy | A (xy.2, xy.1)]. -Notation "A ^-1" := (inv_rel A) : classical_set_scope. +Local Notation "A ^-1" := ([set xy | A (xy.2, xy.1)]) : classical_set_scope. -Definition set_of_rel (T T' : Type) (A : set (T * T')) (x : T) := - [set y | A (x, y)]. +Local Notation "'to_set' A x" := ([set y | A (x, y)]) + (at level 0, A at level 0) : classical_set_scope. Definition locally_ {T T'} (ent : set (set (T * T'))) (x : T) := - filter_from ent (fun A => set_of_rel A x). + filter_from ent (fun A => to_set A x). Lemma locally_E {T T'} (ent : set (set (T * T'))) x : - locally_ ent x = filter_from ent (fun A => set_of_rel A x). + locally_ ent x = filter_from ent (fun A => to_set A x). Proof. by []. Qed. Module Uniform. @@ -2310,7 +2307,7 @@ Lemma locally_entourageE {M : uniformType} : locally_ (@entourages M) = locally. Proof. by case: M=> [?[?[]]]. Qed. Lemma filter_from_entourageE {M : uniformType} x : - filter_from (@entourages M) (fun A => set_of_rel A x) = locally x. + filter_from (@entourages M) (fun A => to_set A x) = locally x. Proof. by rewrite -locally_entourageE. Qed. Module Export LocallyEntourage. @@ -2387,7 +2384,7 @@ apply/locallyP; exists (split_ent B) => // z hByz; apply/sBA. by apply: subset_split_ent => //; exists x => //=; apply: (close_sym cxy). Qed. -Lemma locally_entourage (x : M) A : entourages A -> locally x (set_of_rel A x). +Lemma locally_entourage (x : M) A : entourages A -> locally x (to_set A x). Proof. by move=> ?; apply/locallyP; exists A. Qed. Hint Resolve locally_entourage. @@ -2488,22 +2485,23 @@ Definition prod_ent := filter_prod (@entourages U) (@entourages V) [set ((xy.1.1,xy.2.1),(xy.1.2,xy.2.2)) | xy in A]]. +Lemma prod_entP (A : set (U * U)) (B : set (V * V)) : + entourages A -> entourages B -> + prod_ent [set xy | A (xy.1.1, xy.2.1) /\ B (xy.1.2, xy.2.2)]. +Proof. +move=> entA entB; exists (A,B) => // xy ABxy. +by exists ((xy.1.1, xy.2.1),(xy.1.2,xy.2.2)); rewrite -!surjective_pairing. +Qed. + Lemma prod_ent_filter : Filter prod_ent. Proof. have prodF := filter_prod_filter (@entourage_filter U) (@entourage_filter V). split; rewrite /prod_ent; last 1 first. -- by move=> A B sAB; apply: filterS => xy [zt /sAB Bzt ztexy]; exists zt. -- apply: filterS filterT => xy _. - exists ((xy.1.1,xy.2.1),(xy.1.2,xy.2.2)) => //=. - by rewrite -!surjective_pairing. +- by move=> A B sAB; apply: filterS => ? [xy /sAB ??]; exists xy. +- rewrite -setMT; apply: prod_entP filterT filterT. move=> A B entA entB; apply: filterS (filterI entA entB) => xy []. move=> [zt Azt ztexy] [zt' Bzt' zt'exy]; exists zt => //; split=> //. -suff -> : zt = zt' by []. -move/eqP: ztexy; rewrite [xy]surjective_pairing. -rewrite [xy.1]surjective_pairing [xy.2]surjective_pairing !xpair_eqE. -rewrite andbACA -!xpair_eqE -!surjective_pairing => /eqP->. -move/eqP: zt'exy; rewrite [xy]surjective_pairing. -rewrite [xy.1]surjective_pairing [xy.2]surjective_pairing !xpair_eqE. +move/eqP: ztexy; rewrite -zt'exy !xpair_eqE. by rewrite andbACA -!xpair_eqE -!surjective_pairing => /eqP->. Qed. @@ -2521,21 +2519,16 @@ Qed. Lemma prod_ent_inv A : prod_ent A -> prod_ent A^-1. Proof. move=> [B [/entourage_inv entB1 /entourage_inv entB2] sBA]. -exists (B.1^-1,B.2^-1) => //= xy []. -rewrite [xy.1]surjective_pairing [xy.2]surjective_pairing /inv_rel /=. -move=> Bxy1 Bxy2. -have /sBA : (B.1 `*` B.2) ((xy.1.2, xy.1.1),(xy.2.2, xy.2.1)) by []. -move=> [zt Azt /eqP]; rewrite !xpair_eqE. -rewrite [X in X && _]andbC [X in _ && X]andbC -!xpair_eqE -!surjective_pairing. -by move/eqP<-; exists (zt.2, zt.1) => //=; rewrite -surjective_pairing. +have:= prod_entP entB1 entB2; rewrite /prod_ent; apply: filterS. +move=> _ [p /(sBA (_,_)) [[x y] ? xyE] <-]; exists (y,x) => //=; move/eqP: xyE. +by rewrite !xpair_eqE => /andP[/andP[/eqP-> /eqP->] /andP[/eqP-> /eqP->]]. Qed. Lemma prod_ent_split A : prod_ent A -> exists2 B, prod_ent B & B \o B `<=` A. Proof. move=> [B [entB1 entB2]] sBA; exists [set xy | split_ent B.1 (xy.1.1,xy.2.1) /\ split_ent B.2 (xy.1.2,xy.2.2)]. - exists (split_ent B.1,split_ent B.2) => //= xy hBxy. - by exists ((xy.1.1, xy.2.1),(xy.1.2,xy.2.2)); rewrite -!surjective_pairing. + by apply: prod_entP; apply: entourages_split_ent. move=> xy [uv /= [hB1xyuv1 hB2xyuv1] [hB1xyuv2 hB2xyuv2]]. have /sBA : (B.1 `*` B.2) ((xy.1.1, xy.2.1),(xy.1.2,xy.2.2)). by split=> /=; apply: subset_split_ent => //; [exists uv.1|exists uv.2]. @@ -2545,19 +2538,16 @@ Qed. Lemma prod_ent_locallyE : locally = locally_ prod_ent. Proof. -rewrite predeq2E => xy A; split. - move=> [B []]; rewrite -!locally_entourageE. - move=> [C1 entC1 sCB1] [C2 entC2 sCB2] sBA. +rewrite predeq2E => xy A; split=> [[B []] | [B [C [entC1 entC2] sCB] sBA]]. + rewrite -!locally_entourageE => - [C1 entC1 sCB1] [C2 entC2 sCB2] sBA. exists [set xy | C1 (xy.1.1, xy.2.1) /\ C2 (xy.1.2, xy.2.2)]. - exists (C1, C2) => //= zt [Czt1 Czt2]. - by exists ((zt.1.1, zt.2.1),(zt.1.2,zt.2.2)); rewrite -!surjective_pairing. + exact: prod_entP. by move=> uv [/= /sCB1 Buv1 /sCB2 /(conj Buv1) /sBA]. -move=> [B [C [entC1 entC2] sCB] sBA]. -exists (set_of_rel C.1 xy.1,set_of_rel C.2 xy.2). +exists (to_set (C.1) (xy.1), to_set (C.2) (xy.2)). by rewrite -!locally_entourageE; split; [exists C.1|exists C.2]. move=> uv [/= Cxyuv1 Cxyuv2]; apply: sBA. have /sCB : (C.1 `*` C.2) ((xy.1,uv.1),(xy.2,uv.2)) by []. -move=> [zt Bzt /eqP]; rewrite /set_of_rel !xpair_eqE andbACA -!xpair_eqE. +move=> [zt Bzt /eqP]; rewrite !xpair_eqE andbACA -!xpair_eqE. by rewrite -!surjective_pairing => /eqP<-. Qed. @@ -2578,8 +2568,7 @@ Lemma flim_entourages2P (U V : uniformType) {F : set (set U)} {G : set (set V)} Proof. split=> [/flim_entouragesP FGyz A /FGyz|FGyz]. by rewrite near_simpl -near2_pair. -apply/flim_entouragesP => A /FGyz. -by rewrite (near2_pair _ _ [set p | A ((y,z),p)]). +by apply/flim_entouragesP => A /FGyz; rewrite (near2_pair _ _ (to_set A (y,z))). Qed. (** matrices *) @@ -2635,17 +2624,14 @@ Lemma mx_ent_locallyE : locally = locally_ mx_ent. Proof. rewrite predeq2E => M A; split. move=> [B]; rewrite -locally_entourageE => M_B sBA. - have {M_B} M_B : forall i j, exists C, entourages C /\ - set_of_rel C (M i j) `<=` B i j by move=> i j; apply/exists2P/M_B. - exists [set MN : 'M[T]_(m, n) * 'M[T]_(m, n) | - forall i j, get [set C | entourages C /\ set_of_rel C (M i j) `<=` B i j] - (MN.1 i j, MN.2 i j)]. - exists (fun i j => get [set C | entourages C /\ - set_of_rel C (M i j) `<=` B i j]) => // i j. - by have /getPex [] := M_B i j. + set sB := fun i j => [set C | entourages C /\ to_set C (M i j) `<=` B i j]. + have {M_B} M_B : forall i j, sB i j !=set0 by move=> ??; apply/exists2P/M_B. + exists [set MN : 'M[T]_(m, n) * 'M[T]_(m, n) | forall i j, + get (sB i j) (MN.1 i j, MN.2 i j)]. + by exists (fun i j => get (sB i j)) => // i j; have /getPex [] := M_B i j. move=> N CMN; apply/sBA => i j; have /getPex [_] := M_B i j; apply. exact/CMN. -move=> [B [C entC sCB] sBA]; exists (fun i j => set_of_rel (C i j) (M i j)). +move=> [B [C entC sCB] sBA]; exists (fun i j => to_set (C i j) (M i j)). by rewrite -locally_entourageE => i j; exists (C i j). by move=> N CMN; apply/sBA/sCB. Qed. @@ -2667,7 +2653,7 @@ Lemma flim_mx_entouragesP (T : uniformType) m n (F : set (set 'M[T]_(m,n))) Proof. split. rewrite filter_fromP => FM A entA. - apply: (FM (fun i j => set_of_rel A (M i j))). + apply: (FM (fun i j => to_set A (M i j))). by move=> ??; apply: locally_entourage. move=> FM; apply/flim_entouragesP => A [P entP sPA]; near=> N. apply: sPA => /=; near: N; set Q := \bigcap_ij P ij.1 ij.2. @@ -2758,7 +2744,7 @@ Lemma cvg_cauchy {T : uniformType} (F : set (set T)) : Filter F -> [cvg F in T] -> cauchy F. Proof. move=> FF cvF A entA; have /entourage_split [B entB sB2A] := entA. -exists (set_of_rel B^-1 (lim F), set_of_rel B (lim F)). +exists (to_set (B^-1) (lim F), to_set B (lim F)). split=> /=; apply: cvF; rewrite /= -locally_entourageE; last by exists B. by exists B^-1 => //; apply: entourage_inv. by move=> ab [/= Balima Blimb]; apply: sB2A; exists (lim F). @@ -2904,9 +2890,9 @@ Lemma flim_switch_1 {U : uniformType} Proof. move=> fg fh hl; apply/app_flim_entouragesP => A entA. near F1 => x1; near=> x2; apply: (entourages_split (h x1)) => //. - by near: x1; apply/(hl (set_of_rel _ l))/locally_entourage. + by near: x1; apply/(hl (to_set _ l))/locally_entourage. apply: (entourages_split (f x1 x2)) => //. - by near: x2; apply/(fh x1 (set_of_rel _ _))/locally_entourage. + by near: x2; apply/(fh x1 (to_set _ _))/locally_entourage. move: (x2); near: x1; have /flim_fct_entouragesP /(_ (fun=> _^-1)):= fg; apply. by move=> _; apply: entourage_inv. Grab Existential Variables. all: end_near. Qed. @@ -2920,9 +2906,10 @@ Proof. move=> fg fh; apply: complete_cauchy => A entA. rewrite !near_simpl -near2_pair near_map2; near=> x1 y1 => /=; near F2 => x2. apply: (entourages_split (f x1 x2)) => //. - by near: x2; apply/(fh _ (set_of_rel _ _))/locally_entourage. + by near: x2; apply/(fh _ (to_set _ _))/locally_entourage. apply: (entourages_split (f y1 x2)) => //; last first. - by near: x2; apply/(fh _ (set_of_rel _^-1 _))/locally_entourage/entourage_inv. + near: x2; apply/(fh _ (to_set (_^-1) _)). + exact: locally_entourage (entourage_inv _). apply: (entourages_split (g x2)) => //; move: (x2); [near: x1|near: y1]. have /flim_fct_entouragesP /(_ (fun=> _^-1)) := fg; apply=> _. exact: entourage_inv. From 091772f83078d850bd993f6f281ebced0816f728 Mon Sep 17 00:00:00 2001 From: damien rouhling Date: Fri, 12 Oct 2018 10:08:33 +0200 Subject: [PATCH 09/11] More clean up --- derive.v | 2 +- hierarchy.v | 64 +++++++++--------- topology.v | 192 ++++++++++++++++++++++++++-------------------------- 3 files changed, 128 insertions(+), 130 deletions(-) diff --git a/derive.v b/derive.v index 5e0325a15f..aea1aeda70 100644 --- a/derive.v +++ b/derive.v @@ -283,7 +283,7 @@ pose g1 : R -> W := fun h => (h^-1 * h) *: 'd f a v. pose g2 : R -> W := fun h : R => h^-1 *: k (h *: v ). rewrite (_ : g = g1 + g2) ?funeqE // -(addr0 (_ _ v)); apply: lim_add. rewrite -(scale1r (_ _ v)); apply: lim_scalel. - apply/app_flim_entouragesP => X entX; apply/locallyP. + apply/app_flim_entourageP => X entX; apply/locallyP. rewrite locally_E; exists X => // x _ x0; rewrite mulVf //. exact: entourage_refl. rewrite /g2. diff --git a/hierarchy.v b/hierarchy.v index c2f9fbcbb5..5af3818b64 100644 --- a/hierarchy.v +++ b/hierarchy.v @@ -29,7 +29,7 @@ Require Import boolp reals classical_sets posnum topology. (* ball x e == ball of center x and radius e. *) (* close x y <-> x and y are arbitrarily close w.r.t. to *) (* balls. *) -(* entourages == set of entourages defined by balls. An *) +(* entourage == set of entourages defined by balls. An *) (* entourage can be seen as a *) (* "neighbourhood" of the diagonal set *) (* D = {(x, x) | x in T}. *) @@ -135,7 +135,7 @@ Definition ball (K : numDomainType) (V : zmodType) (norm : V -> K) (x : V) (e : K) := [set y | norm (x - y) < e]. Arguments ball {K V} norm x e%R y /. -Definition entourages_ (K : numDomainType) (V : zmodType) (norm : V -> K) := +Definition entourage_ (K : numDomainType) (V : zmodType) (norm : V -> K) := filter_from [set e : K | e > 0] (fun e => [set xy | norm (xy.1 - xy.2) < e]). @@ -146,7 +146,7 @@ Record mixin_of (K : numDomainType) (V : lmodType K) loc norm : V -> K ; ax1 : forall (x y : V), norm (x + y) <= norm x + norm y ; ax2 : forall (l : K) (x : V), norm (l *: x) = `|l| * norm x; - ax3 : Uniform.entourages m = entourages_ norm; + ax3 : Uniform.entourage m = entourage_ norm; ax4 : forall x : V, norm x = 0 -> x = 0 }. @@ -255,11 +255,11 @@ Proof. exact: NormedModule.ax1. Qed. Lemma normmZ l x : `|[l *: x]| = `|l| * `|[x]|. Proof. exact: NormedModule.ax2. Qed. -Notation entourages_norm := (entourages_ (@norm K V)). +Notation entourage_norm := (entourage_ (@norm K V)). -Notation locally_norm := (locally_ entourages_norm). +Notation locally_norm := (locally_ entourage_norm). -Lemma entourages_normE : entourages_norm = entourages. +Lemma entourage_normE : entourage_norm = entourage. Proof. by rewrite -NormedModule.ax3. Qed. Lemma normm0_eq0 x : `|[x]| = 0 -> x = 0. @@ -310,25 +310,26 @@ rewrite ger0_norm ?subr_ge0 // ler_subl_addr. by rewrite -{1}[x](addrNK y) ler_normm_add. Qed. -Lemma entourages_ball (e : {posnum K}) : - entourages [set xy : V * V | `|[xy.1 - xy.2]| < e%:num]. -Proof. by rewrite -entourages_normE; apply: in_filter_from. Qed. +Lemma entourage_ball (e : {posnum K}) : + entourage [set xy : V * V | `|[xy.1 - xy.2]| < e%:num]. +Proof. by rewrite -entourage_normE; apply: in_filter_from. Qed. +Hint Resolve entourage_ball. Lemma locally_le_locally_norm x : flim (locally x) (locally_norm x). Proof. move=> P [A entA sAB]; apply/locallyP; exists A => //. -by rewrite -entourages_normE. +by rewrite -entourage_normE. Qed. Lemma locally_norm_le_locally x : flim (locally_norm x) (locally x). Proof. -by move=> P /locallyP [A entA sAP]; exists A => //; rewrite entourages_normE. +by move=> P /locallyP [A entA sAP]; exists A => //; rewrite entourage_normE. Qed. (* NB: this lemmas was not here before *) Lemma locally_locally_norm : locally_norm = locally. Proof. -by rewrite funeqE => x; rewrite /locally_norm entourages_normE +by rewrite funeqE => x; rewrite /locally_norm entourage_normE filter_from_entourageE. Qed. @@ -336,8 +337,7 @@ Lemma filter_from_norm_locally x : @filter_from K _ [set x : K | 0 < x] (ball norm x) = locally x. Proof. rewrite predeqE => A; split=> [[_/posnumP[e] sxeA] |]. - by rewrite -locally_entourageE; exists [set xy | `|[xy.1 - xy.2]| < e%:num]; - [apply: entourages_ball|move=> ? /sxeA]. + by rewrite -locally_entourageE; exists [set xy | `|[xy.1 - xy.2]| < e%:num]. rewrite -locally_locally_norm => - [B [_/posnumP[e] seB] sBA]. by exists e%:num => // y xye; apply/sBA/seB. Qed. @@ -394,13 +394,14 @@ Proof. by move=> e1e2 y /ltr_le_trans; apply. Qed. Lemma norm_close x y : close x y = (forall e : {posnum K}, ball norm x e%:num y). Proof. -rewrite propeqE; split; first by move=> xy e; have /xy := entourages_ball e. -by move=> xy A; rewrite -entourages_normE => - [_/posnumP[e]]; apply; apply: xy. +rewrite propeqE; split; first by move=> xy e; have /xy := entourage_ball e. +by move=> xy A; rewrite -entourage_normE => - [_/posnumP[e]]; apply; apply: xy. Qed. End NormedModule1. Hint Resolve normm_ge0. +Hint Resolve entourage_ball. Hint Resolve ball_center. Module Export LocallyNorm. @@ -471,7 +472,7 @@ Proof. rewrite propeqE; split => [cl_xy|->//]; have [//|neq_xy] := eqVneq x y. have dxy_gt0 : `|[x - y]| > 0 by rewrite normm_gt0 subr_eq0. have dxy_ge0 := ltrW dxy_gt0. -have /cl_xy /= := (@entourages_ball _ V ((PosNum dxy_gt0)%:num / 2)%:pos). +have /cl_xy /= := (@entourage_ball _ V ((PosNum dxy_gt0)%:num / 2)%:pos). rewrite -subr_lt0 ler_gtF // -[X in X - _]mulr1 -mulrBr mulr_ge0 //. by rewrite subr_ge0 -(@ler_pmul2r _ 2) // mulVf // mul1r ler1n. Qed. @@ -530,7 +531,7 @@ Hypothesis (ler_dist_add : forall z x y, norm (x - y) <= norm (x - z) + norm (z - y)). Program Definition uniformityOfNormMixin := - @Uniform.Mixin V (locally_ (entourages_ norm)) (entourages_ norm) _ _ _ _ _. + @Uniform.Mixin V (locally_ (entourage_ norm)) (entourage_ norm) _ _ _ _ _. Next Obligation. apply: filter_from_filter; first by exists 1. move=> _ _ /posnumP[e1] /posnumP[e2]; exists (minr e1%:num e2%:num) => // xy. @@ -556,7 +557,7 @@ Definition realField_uniformMixin (R : realFieldType) := uniformityOfNormMixin (@normr0 R) (@distrC _) (@ler_dist_add _). Canonical realField_filteredType (R : realFieldType) := - FilteredType R R (locally_ (entourages_ normr)). + FilteredType R R (locally_ (entourage_ normr)). Canonical realField_topologicalType (R : realFieldType) := TopologicalType R (topologyOfEntourageMixin (realField_uniformMixin R)). Canonical realField_uniformType (R : realFieldType):= @@ -564,7 +565,7 @@ Canonical realField_uniformType (R : realFieldType):= Definition realField_normedModMixin (R : realFieldType) := @NormedModule.Mixin _ (GRing.regular_lmodType R) - (locally_ (entourages_ normr)) (realField_uniformMixin _) normr + (locally_ (entourage_ normr)) (realField_uniformMixin _) normr (@ler_norm_add _) (@normrM _) erefl (@normr0_eq0 _). Canonical realFieldo_normedModType (R : realFieldType) := @@ -1077,16 +1078,15 @@ Lemma prod_norm_scal (l : K) (x : U * V) : prod_norm (l *: x) = `|l| * prod_norm x. Proof. by rewrite /prod_norm !normmZ maxr_pmulr. Qed. -Lemma entourages_prod_normE : entourages = entourages_ prod_norm. +Lemma entourage_prod_normE : entourage = entourage_ prod_norm. Proof. rewrite predeqE => A; split; last first. - move=> [_/posnumP[e] sA]. - exists ([set u | `|[u.1 - u.2]| < e%:num], [set v | `|[v.1 - v.2]| < e%:num]). - by split=> /=; apply: entourages_ball. + move=> [_/posnumP[e] sA]; exists ([set u | `|[u.1 - u.2]| < e%:num], + [set v | `|[v.1 - v.2]| < e%:num]) => //=. move=> /= uv [uv1e uv2e]; exists ((uv.1.1, uv.2.1), (uv.1.2, uv.2.2)). by apply: sA; rewrite ltr_maxl uv1e uv2e. by rewrite /= -!surjective_pairing. -move=> [PQ []]; rewrite -!entourages_normE. +move=> [PQ []]; rewrite -!entourage_normE. move=> [_/posnumP[eP] sP] [_/posnumP[eQ] sQ] sPQA. exists (minr eP%:num eQ%:num) => // xy. rewrite ltr_maxl !ltr_minr => /andP [/andP [xy1P xy1Q] /andP [xy2P xy2Q]]. @@ -1113,7 +1113,7 @@ End prod_NormedModule. Definition prod_NormedModule_mixin (K : realDomainType) (U V : normedModType K) := @NormedModMixin K _ _ _ (@prod_norm K U V) prod_norm_triangle - prod_norm_scal entourages_prod_normE prod_norm_eq0. + prod_norm_scal entourage_prod_normE prod_norm_eq0. Canonical prod_NormedModule (K : realDomainType) (U V : normedModType K) := NormedModType K (U * V) (@prod_NormedModule_mixin K U V). @@ -1374,11 +1374,11 @@ Export CompleteNormedModule.Exports. (** * Extended Types *) -Definition entourages_set (U : uniformType) (A : set ((set U) * (set U))) := - exists2 B, entourages B & forall PQ, A PQ -> forall p q, +Definition entourage_set (U : uniformType) (A : set ((set U) * (set U))) := + exists2 B, entourage B & forall PQ, A PQ -> forall p q, PQ.1 p -> PQ.2 q -> B (p,q). Canonical set_filter_source (U : uniformType) := - @Filtered.Source Prop _ U (fun A => locally_ (@entourages_set U) A). + @Filtered.Source Prop _ U (fun A => locally_ (@entourage_set U) A). (** * The topology on real numbers *) @@ -1400,7 +1400,7 @@ Lemma real_complete (F : set (set R)) : ProperFilter F -> cauchy F -> cvg F. Proof. move=> FF Fc; apply/cvg_ex. pose D := \bigcap_(A in F) (down (mem A)). -have /Fc := @entourages_ball _ [normedModType R of R] 1%:pos. +have /Fc := @entourage_ball _ [normedModType R of R] 1%:pos. rewrite near_simpl -near2_pair => /nearP_dep /filter_ex /= [x0 x01]. have D_has_sup : has_sup (mem D); first split. - exists (x0 - 1); rewrite in_setE => A FA. @@ -1414,7 +1414,7 @@ rewrite ler_distl sup_upper_bound //=. apply: sup_le_ub => //; first by case: D_has_sup. apply/forallbP => y; apply/implyP; rewrite in_setE. move=> /(_ (ball norm x eps%:num) _) /existsbP []. - by near: x; apply: nearP_dep; apply: Fc; apply: entourages_ball. + by near: x; apply: nearP_dep; apply: Fc; apply: entourage_ball. move=> z /andP[]; rewrite in_setE /ball ltr_distl ltr_subl_addr. by move=> /andP [/ltrW /(ler_trans _) le_xeps _ /le_xeps]. rewrite in_setE /D /= => A FA; near F => y. @@ -1422,7 +1422,7 @@ apply/existsbP; exists y; apply/andP; split. by rewrite in_setE; near: y. rewrite ler_subl_addl -ler_subl_addr ltrW //. suff: `|x - y| < eps%:num by rewrite ltr_norml => /andP[_]. -by near: y; near: x; apply: nearP_dep; apply: Fc; apply: entourages_ball. +by near: y; near: x; apply: nearP_dep; apply: Fc; apply: entourage_ball. Grab Existential Variables. all: end_near. Qed. Canonical real_completeType := CompleteType R real_complete. diff --git a/topology.v b/topology.v index 6c78deae5d..1f18621f04 100644 --- a/topology.v +++ b/topology.v @@ -2207,12 +2207,12 @@ Proof. by []. Qed. Module Uniform. Record mixin_of (M : Type) (locally : M -> set (set M)) := Mixin { - entourages : (M * M -> Prop) -> Prop ; - ax1 : Filter entourages ; - ax2 : forall A, entourages A -> [set xy | xy.1 = xy.2] `<=` A ; - ax3 : forall A, entourages A -> entourages A^-1 ; - ax4 : forall A, entourages A -> exists2 B, entourages B & B \o B `<=` A ; - ax5 : locally = locally_ entourages + entourage : (M * M -> Prop) -> Prop ; + ax1 : Filter entourage ; + ax2 : forall A, entourage A -> [set xy | xy.1 = xy.2] `<=` A ; + ax3 : forall A, entourage A -> entourage A^-1 ; + ax4 : forall A, entourage A -> exists2 B, entourage B & B \o B `<=` A ; + ax5 : locally = locally_ entourage }. Record class_of (M : Type) := Class { @@ -2301,13 +2301,13 @@ Qed. End UniformTopology. -Definition entourages {M : uniformType} := Uniform.entourages (Uniform.class M). +Definition entourage {M : uniformType} := Uniform.entourage (Uniform.class M). -Lemma locally_entourageE {M : uniformType} : locally_ (@entourages M) = locally. +Lemma locally_entourageE {M : uniformType} : locally_ (@entourage M) = locally. Proof. by case: M=> [?[?[]]]. Qed. Lemma filter_from_entourageE {M : uniformType} x : - filter_from (@entourages M) (fun A => to_set A x) = locally x. + filter_from (@entourage M) (fun A => to_set A x) = locally x. Proof. by rewrite -locally_entourageE. Qed. Module Export LocallyEntourage. @@ -2316,55 +2316,55 @@ Definition locally_simpl := End LocallyEntourage. Lemma locallyP {M : uniformType} (x : M) P : - locally x P <-> locally_ entourages x P. + locally x P <-> locally_ entourage x P. Proof. by rewrite locally_simpl. Qed. Section uniformType1. Context {M : uniformType}. Lemma entourage_refl (A : set (M * M)) x : - entourages A -> A (x, x). + entourage A -> A (x, x). Proof. by move=> entA; apply: Uniform.ax2 entA _ _. Qed. -Global Instance entourage_filter : ProperFilter (@entourages M). +Global Instance entourage_filter : ProperFilter (@entourage M). Proof. apply Build_ProperFilter; last exact: Uniform.ax1. by move=> A entA; exists (point, point); apply: entourage_refl. Qed. -Lemma entourageT : entourages (@setT (M * M)). +Lemma entourageT : entourage (@setT (M * M)). Proof. exact: filterT. Qed. -Lemma entourage_inv (A : set (M * M)) : entourages A -> entourages A^-1. +Lemma entourage_inv (A : set (M * M)) : entourage A -> entourage A^-1. Proof. exact: Uniform.ax3. Qed. -Lemma entourage_split (A : set (M * M)) : - entourages A -> exists2 B, entourages B & B \o B `<=` A. +Lemma entourage_split_ex (A : set (M * M)) : + entourage A -> exists2 B, entourage B & B \o B `<=` A. Proof. exact: Uniform.ax4. Qed. Definition split_ent (A : set (M * M)) := - get (entourages `&` [set B | B \o B `<=` A]). + get (entourage `&` [set B | B \o B `<=` A]). -Lemma split_entP (A : set (M * M)) : entourages A -> - entourages (split_ent A) /\ split_ent A \o split_ent A `<=` A. -Proof. by move/entourage_split/exists2P/getPex. Qed. +Lemma split_entP (A : set (M * M)) : entourage A -> + entourage (split_ent A) /\ split_ent A \o split_ent A `<=` A. +Proof. by move/entourage_split_ex/exists2P/getPex. Qed. -Lemma entourages_split_ent (A : set (M * M)) : entourages A -> - entourages (split_ent A). +Lemma entourage_split_ent (A : set (M * M)) : entourage A -> + entourage (split_ent A). Proof. by move=> /split_entP []. Qed. -Hint Extern 0 (entourages (split_ent _)) => exact: entourages_split_ent : core. -Hint Extern 0 (entourages (get _)) => exact: entourages_split_ent : core. +Hint Extern 0 (entourage (split_ent _)) => exact: entourage_split_ent : core. +Hint Extern 0 (entourage (get _)) => exact: entourage_split_ent : core. -Lemma subset_split_ent (A : set (M * M)) : entourages A -> +Lemma subset_split_ent (A : set (M * M)) : entourage A -> split_ent A \o split_ent A `<=` A. Proof. by move=> /split_entP []. Qed. -Lemma entourages_split (z x y : M) A : entourages A -> +Lemma entourage_split (z x y : M) A : entourage A -> split_ent A (x,z) -> split_ent A (z,y) -> A (x,y). Proof. by move=> /subset_split_ent sA ??; apply: sA; exists z. Qed. -Arguments entourages_split z {x y A}. +Arguments entourage_split z {x y A}. -Definition close (x y : M) : Prop := forall A, entourages A -> A (x,y). +Definition close (x y : M) : Prop := forall A, entourage A -> A (x,y). Lemma close_refl (x : M) : close x x. Proof. by move=> ? /entourage_refl. Qed. @@ -2373,7 +2373,7 @@ Proof. by move=> cxy ? /entourage_inv /cxy. Qed. Lemma close_trans (x y z : M) : close x y -> close y z -> close x z. Proof. -by move=> cxy cyz A entA; apply: (entourages_split y) => //; +by move=> cxy cyz A entA; apply: (entourage_split y) => //; [apply: cxy|apply: cyz]. Qed. @@ -2384,7 +2384,7 @@ apply/locallyP; exists (split_ent B) => // z hByz; apply/sBA. by apply: subset_split_ent => //; exists x => //=; apply: (close_sym cxy). Qed. -Lemma locally_entourage (x : M) A : entourages A -> locally x (to_set A x). +Lemma locally_entourage (x : M) A : entourage A -> locally x (to_set A x). Proof. by move=> ?; apply/locallyP; exists A. Qed. Hint Resolve locally_entourage. @@ -2399,32 +2399,32 @@ Grab Existential Variables. all: end_near. Qed. Lemma flimx_close (x y : M) : x --> y -> close x y. Proof. exact: flim_close. Qed. -Lemma flim_entouragesP F (FF : Filter F) (p : M) : - F --> p <-> forall A, entourages A -> \forall q \near F, A (p, q). +Lemma flim_entourageP F (FF : Filter F) (p : M) : + F --> p <-> forall A, entourage A -> \forall q \near F, A (p, q). Proof. by rewrite -filter_fromP !locally_simpl. Qed. -Lemma flim_entourages {F} {FF : Filter F} (y : M) : - F --> y -> forall A, entourages A -> \forall y' \near F, A (y,y'). -Proof. by move/flim_entouragesP. Qed. +Lemma flim_entourage {F} {FF : Filter F} (y : M) : + F --> y -> forall A, entourage A -> \forall y' \near F, A (y,y'). +Proof. by move/flim_entourageP. Qed. -Lemma app_flim_entouragesP T (f : T -> M) F (FF : Filter F) p : - f @ F --> p <-> forall A, entourages A -> \forall t \near F, A (p, f t). -Proof. exact: flim_entouragesP. Qed. +Lemma app_flim_entourageP T (f : T -> M) F (FF : Filter F) p : + f @ F --> p <-> forall A, entourage A -> \forall t \near F, A (p, f t). +Proof. exact: flim_entourageP. Qed. -Lemma flimi_entouragesP T {F} {FF : Filter F} (f : T -> M -> Prop) y : +Lemma flimi_entourageP T {F} {FF : Filter F} (f : T -> M -> Prop) y : f `@ F --> y <-> - forall A, entourages A -> \forall x \near F, exists z, f x z /\ A (y,z). + forall A, entourage A -> \forall x \near F, exists z, f x z /\ A (y,z). Proof. split=> [Fy A entA|Fy A] /=; first exact/Fy/locally_entourage. move=> /locallyP [B entB sBA]; rewrite near_simpl near_mapi; near=> x. by have [//|z [fxz Byz]]:= near (Fy _ entB) x; exists z; split=> //; apply: sBA. Unshelve. all: end_near. Qed. -Definition flimi_locally := @flimi_entouragesP. +Definition flimi_locally := @flimi_entourageP. -Lemma flimi_entourages T {F} {FF : Filter F} (f : T -> M -> Prop) y : +Lemma flimi_entourage T {F} {FF : Filter F} (f : T -> M -> Prop) y : f `@ F --> y -> - forall A, entourages A -> F [set x | exists z, f x z /\ A (y,z)]. -Proof. by move/flimi_entouragesP. Qed. + forall A, entourage A -> F [set x | exists z, f x z /\ A (y,z)]. +Proof. by move/flimi_entourageP. Qed. Lemma flimi_close T {F} {FF: ProperFilter F} (f : T -> set M) (l l' : M) : {near F, is_fun f} -> f `@ F --> l -> f `@ F --> l' -> close l l'. @@ -2432,7 +2432,7 @@ Proof. move=> f_prop fFl fFl'. suff f_totalfun: infer {near F, is_totalfun f} by exact: flim_close fFl fFl'. apply: filter_app f_prop; near=> x; split=> //=. -by have [//|y []] := near (flimi_entourages fFl entourageT) x; exists y. +by have [//|y []] := near (flimi_entourage fFl entourageT) x; exists y. Grab Existential Variables. all: end_near. Qed. Definition flimi_locally_close := @flimi_close. @@ -2462,16 +2462,16 @@ Qed. End uniformType1. -Hint Extern 0 (entourages (split_ent _)) => exact: entourages_split_ent : core. -Hint Extern 0 (entourages (get _)) => exact: entourages_split_ent : core. -Arguments entourages_split {M} z {x y A}. -Hint Resolve locally_entourage. +Hint Extern 0 (entourage (split_ent _)) => exact: entourage_split_ent : core. +Hint Extern 0 (entourage (get _)) => exact: entourage_split_ent : core. +Arguments entourage_split {M} z {x y A}. +Hint Extern 0 (locally _ (to_set _ _)) => exact: locally_entourage : core. Hint Resolve close_refl. Arguments flim_const {M T F FF} a. Arguments close_lim {M} F1 F2 {FF2} _. Definition unif_cont (U V : uniformType) (f : U -> V) := - (fun xy => (f xy.1, f xy.2)) @ entourages --> entourages. + (fun xy => (f xy.1, f xy.2)) @ entourage --> entourage. (** product of two uniform spaces *) @@ -2482,11 +2482,11 @@ Implicit Types A : set ((U * V) * (U * V)). Definition prod_ent := [set A : set ((U * V) * (U * V)) | - filter_prod (@entourages U) (@entourages V) + filter_prod (@entourage U) (@entourage V) [set ((xy.1.1,xy.2.1),(xy.1.2,xy.2.2)) | xy in A]]. Lemma prod_entP (A : set (U * U)) (B : set (V * V)) : - entourages A -> entourages B -> + entourage A -> entourage B -> prod_ent [set xy | A (xy.1.1, xy.2.1) /\ B (xy.1.2, xy.2.2)]. Proof. move=> entA entB; exists (A,B) => // xy ABxy. @@ -2528,7 +2528,7 @@ Lemma prod_ent_split A : prod_ent A -> exists2 B, prod_ent B & B \o B `<=` A. Proof. move=> [B [entB1 entB2]] sBA; exists [set xy | split_ent B.1 (xy.1.1,xy.2.1) /\ split_ent B.2 (xy.1.2,xy.2.2)]. - by apply: prod_entP; apply: entourages_split_ent. + by apply: prod_entP; apply: entourage_split_ent. move=> xy [uv /= [hB1xyuv1 hB2xyuv1] [hB1xyuv2 hB2xyuv2]]. have /sBA : (B.1 `*` B.2) ((xy.1.1, xy.2.1),(xy.1.2,xy.2.2)). by split=> /=; apply: subset_split_ent => //; [exists uv.1|exists uv.2]. @@ -2560,15 +2560,15 @@ End prod_Uniform. Canonical prod_uniformType (U V : uniformType) := UniformType (U * V) (@prod_uniformType_mixin U V). -Lemma flim_entourages2P (U V : uniformType) {F : set (set U)} {G : set (set V)} +Lemma flim_entourage2P (U V : uniformType) {F : set (set U)} {G : set (set V)} {FF : Filter F} {FG : Filter G} (y : U) (z : V): (F, G) --> (y, z) <-> - forall A, entourages A -> + forall A, entourage A -> \forall y' \near F & z' \near G, A ((y,z),(y',z')). Proof. -split=> [/flim_entouragesP FGyz A /FGyz|FGyz]. +split=> [/flim_entourageP FGyz A /FGyz|FGyz]. by rewrite near_simpl -near2_pair. -by apply/flim_entouragesP => A /FGyz; rewrite (near2_pair _ _ (to_set A (y,z))). +by apply/flim_entourageP => A /FGyz; rewrite (near2_pair _ _ (to_set A (y,z))). Qed. (** matrices *) @@ -2581,7 +2581,7 @@ Implicit Types A : set ('M[T]_(m, n) * 'M[T]_(m, n)). Definition mx_ent := filter_from - [set P : 'I_m -> 'I_n -> set (T * T) | forall i j, entourages (P i j)] + [set P : 'I_m -> 'I_n -> set (T * T) | forall i j, entourage (P i j)] (fun P => [set MN : 'M[T]_(m, n) * 'M[T]_(m, n) | forall i j, P i j (MN.1 i j, MN.2 i j)]). @@ -2609,12 +2609,12 @@ Qed. Lemma mx_ent_split A : mx_ent A -> exists2 B, mx_ent B & B \o B `<=` A. Proof. move=> [B entB sBA]. -have Bsplit : forall i j, exists C, entourages C /\ C \o C `<=` B i j. - by move=> ??; apply/exists2P/entourage_split. +have Bsplit : forall i j, exists C, entourage C /\ C \o C `<=` B i j. + by move=> ??; apply/exists2P/entourage_split_ex. exists [set MN : 'M[T]_(m, n) * 'M[T]_(m, n) | - forall i j, get [set C | entourages C /\ C \o C `<=` B i j] + forall i j, get [set C | entourage C /\ C \o C `<=` B i j] (MN.1 i j, MN.2 i j)]. - by exists (fun i j => get [set C | entourages C /\ C \o C `<=` B i j]). + by exists (fun i j => get [set C | entourage C /\ C \o C `<=` B i j]). move=> MN [P CMN1P CPMN2]; apply/sBA => i j. have /getPex [_] := Bsplit i j; apply; exists (P i j); first exact: CMN1P. exact: CPMN2. @@ -2624,7 +2624,7 @@ Lemma mx_ent_locallyE : locally = locally_ mx_ent. Proof. rewrite predeq2E => M A; split. move=> [B]; rewrite -locally_entourageE => M_B sBA. - set sB := fun i j => [set C | entourages C /\ to_set C (M i j) `<=` B i j]. + set sB := fun i j => [set C | entourage C /\ to_set C (M i j) `<=` B i j]. have {M_B} M_B : forall i j, sB i j !=set0 by move=> ??; apply/exists2P/M_B. exists [set MN : 'M[T]_(m, n) * 'M[T]_(m, n) | forall i j, get (sB i j) (MN.1 i j, MN.2 i j)]. @@ -2645,17 +2645,15 @@ Canonical matrix_uniformType := End matrix_Uniform. -Lemma flim_mx_entouragesP (T : uniformType) m n (F : set (set 'M[T]_(m,n))) +Lemma flim_mx_entourageP (T : uniformType) m n (F : set (set 'M[T]_(m,n))) (FF : Filter F) (M : 'M[T]_(m,n)) : F --> M <-> - forall A, entourages A -> \forall N \near F, + forall A, entourage A -> \forall N \near F, forall i j, A (M i j, (N : 'M[T]_(m,n)) i j). Proof. split. - rewrite filter_fromP => FM A entA. - apply: (FM (fun i j => to_set A (M i j))). - by move=> ??; apply: locally_entourage. -move=> FM; apply/flim_entouragesP => A [P entP sPA]; near=> N. + by rewrite filter_fromP => FM A ?; apply: (FM (fun i j => to_set A (M i j))). +move=> FM; apply/flim_entourageP => A [P entP sPA]; near=> N. apply: sPA => /=; near: N; set Q := \bigcap_ij P ij.1 ij.2. apply: filterS (FM Q _); first by move=> N QN i j; apply: (QN _ _ (i, j)). have -> : Q = @@ -2672,7 +2670,7 @@ Variable (T : choiceType) (U : uniformType). Definition fct_ent := filter_from - [set P : T -> set (U * U) | forall t, entourages (P t)] + [set P : T -> set (U * U) | forall t, entourage (P t)] (fun P => [set fg | forall t, P t (fg.1 t, fg.2 t)]). Lemma fct_ent_filter : Filter fct_ent. @@ -2699,11 +2697,11 @@ Qed. Lemma fct_ent_split A : fct_ent A -> exists2 B, fct_ent B & B \o B `<=` A. Proof. move=> [B entB sBA]. -have Bsplit t : exists C, entourages C /\ C \o C `<=` B t. - exact/exists2P/entourage_split. +have Bsplit t : exists C, entourage C /\ C \o C `<=` B t. + exact/exists2P/entourage_split_ex. exists [set fg | forall t, - get [set C | entourages C /\ C \o C `<=` B t] (fg.1 t, fg.2 t)]. - by exists (fun t => get [set C | entourages C /\ C \o C `<=` B t]). + get [set C | entourage C /\ C \o C `<=` B t] (fg.1 t, fg.2 t)]. + by exists (fun t => get [set C | entourage C /\ C \o C `<=` B t]). move=> fg [h Cfh Chg]; apply/sBA => t; have /getPex [_] := Bsplit t; apply. by exists (h t); [apply: Cfh|apply: Chg]. Qed. @@ -2722,28 +2720,28 @@ Canonical fct_uniformType := UniformType (T -> U) fct_uniformType_mixin. End fct_Uniform. (* TODO: is it possible to remove A's dependency in t? *) -Lemma flim_fct_entouragesP (T : choiceType) (U : uniformType) +Lemma flim_fct_entourageP (T : choiceType) (U : uniformType) (F : set (set (T -> U))) (FF : Filter F) (f : T -> U) : F --> f <-> - forall A, (forall t, entourages (A t)) -> + forall A, (forall t, entourage (A t)) -> \forall g \near F, forall t, A t (f t, g t). Proof. split. - move=> /flim_entouragesP Ff A entA. + move=> /flim_entourageP Ff A entA. apply: (Ff [set fg | forall t : T, A t (fg.1 t, fg.2 t)]). by exists (fun t => A t). -move=> Ff; apply/flim_entouragesP => A [P entP sPA]; near=> g. +move=> Ff; apply/flim_entourageP => A [P entP sPA]; near=> g. by apply: sPA => /=; near: g; apply: Ff. Grab Existential Variables. all: end_near. Qed. (** ** Complete uniform spaces *) -Definition cauchy {T : uniformType} (F : set (set T)) := (F, F) --> entourages. +Definition cauchy {T : uniformType} (F : set (set T)) := (F, F) --> entourage. Lemma cvg_cauchy {T : uniformType} (F : set (set T)) : Filter F -> [cvg F in T] -> cauchy F. Proof. -move=> FF cvF A entA; have /entourage_split [B entB sB2A] := entA. +move=> FF cvF A entA; have /entourage_split_ex [B entB sB2A] := entA. exists (to_set (B^-1) (lim F), to_set B (lim F)). split=> /=; apply: cvF; rewrite /= -locally_entourageE; last by exists B. by exists B^-1 => //; apply: entourage_inv. @@ -2837,13 +2835,13 @@ Lemma mx_complete (F : set (set 'M[T]_(m, n))) : ProperFilter F -> cauchy F -> cvg F. Proof. move=> FF Fc. -have /(_ _ _) /complete_cauchy /app_flim_entouragesP cvF : +have /(_ _ _) /complete_cauchy /app_flim_entourageP cvF : cauchy ((fun M : 'M[T]_(m, n) => M _ _) @ F). move=> i j A /= entA; rewrite near_simpl -near2E near_map2. by apply: Fc; exists (fun _ _ => A). apply/cvg_ex. set Mlim := \matrix_(i, j) (lim ((fun M : 'M[T]_(m, n) => M i j) @ F) : T). -exists Mlim; apply/flim_mx_entouragesP => A entA; near=> M => i j; near F => M'. +exists Mlim; apply/flim_mx_entourageP => A entA; near=> M => i j; near F => M'. apply: subset_split_ent => //; exists (M' i j) => /=. by near: M'; rewrite mxE; apply: cvF. move: (i) (j); near: M'; near: M; apply: nearP_dep; apply: Fc. @@ -2862,12 +2860,12 @@ Lemma fun_complete (F : set (set (T -> U))) {FF : ProperFilter F} : cauchy F -> cvg F. Proof. move=> Fc. -have /(_ _) /complete_cauchy /app_flim_entouragesP cvF : cauchy (@^~_ @ F). +have /(_ _) /complete_cauchy /app_flim_entourageP cvF : cauchy (@^~_ @ F). move=> t A /= entA; rewrite near_simpl -near2E near_map2. by apply: Fc; exists (fun=> A). apply/cvg_ex; exists (fun t => lim (@^~t @ F)). -apply/flim_fct_entouragesP => A entA; near=> f => t; near F => g. -apply: (entourages_split (g t)) => //; first by near: g; apply: cvF. +apply/flim_fct_entourageP => A entA; near=> f => t; near F => g. +apply: (entourage_split (g t)) => //; first by near: g; apply: cvF. move: (t); near: g; near: f; apply: nearP_dep; apply: Fc. by exists (fun t => (split_ent (A t))^-1) => ? //; apply: entourage_inv. Grab Existential Variables. all: end_near. Qed. @@ -2888,12 +2886,12 @@ Lemma flim_switch_1 {U : uniformType} f @ F1 --> g -> (forall x1, f x1 @ F2 --> h x1) -> h @ F1 --> l -> g @ F2 --> l. Proof. -move=> fg fh hl; apply/app_flim_entouragesP => A entA. -near F1 => x1; near=> x2; apply: (entourages_split (h x1)) => //. - by near: x1; apply/(hl (to_set _ l))/locally_entourage. -apply: (entourages_split (f x1 x2)) => //. - by near: x2; apply/(fh x1 (to_set _ _))/locally_entourage. -move: (x2); near: x1; have /flim_fct_entouragesP /(_ (fun=> _^-1)):= fg; apply. +move=> fg fh hl; apply/app_flim_entourageP => A entA. +near F1 => x1; near=> x2; apply: (entourage_split (h x1)) => //. + by near: x1; apply/(hl (to_set _ l)) => /=. +apply: (entourage_split (f x1 x2)) => //. + by near: x2; apply/(fh x1 (to_set _ _)) => /=. +move: (x2); near: x1; have /flim_fct_entourageP /(_ (fun=> _^-1)):= fg; apply. by move=> _; apply: entourage_inv. Grab Existential Variables. all: end_near. Qed. @@ -2905,15 +2903,15 @@ Lemma flim_switch_2 {U : completeType} Proof. move=> fg fh; apply: complete_cauchy => A entA. rewrite !near_simpl -near2_pair near_map2; near=> x1 y1 => /=; near F2 => x2. -apply: (entourages_split (f x1 x2)) => //. - by near: x2; apply/(fh _ (to_set _ _))/locally_entourage. -apply: (entourages_split (f y1 x2)) => //; last first. +apply: (entourage_split (f x1 x2)) => //. + by near: x2; apply/(fh _ (to_set _ _)) => /=. +apply: (entourage_split (f y1 x2)) => //; last first. near: x2; apply/(fh _ (to_set (_^-1) _)). exact: locally_entourage (entourage_inv _). -apply: (entourages_split (g x2)) => //; move: (x2); [near: x1|near: y1]. - have /flim_fct_entouragesP /(_ (fun=> _^-1)) := fg; apply=> _. +apply: (entourage_split (g x2)) => //; move: (x2); [near: x1|near: y1]. + have /flim_fct_entourageP /(_ (fun=> _^-1)) := fg; apply=> _. exact: entourage_inv. -by have /flim_fct_entouragesP /(_ (fun=> _)) := fg; apply. +by have /flim_fct_entourageP /(_ (fun=> _)) := fg; apply. Grab Existential Variables. all: end_near. Qed. (* Alternative version *) From 3aaf321b7b885c26824db1444ce1fe07a18b121a Mon Sep 17 00:00:00 2001 From: damien rouhling Date: Fri, 12 Oct 2018 13:58:18 +0200 Subject: [PATCH 10/11] Rework documentation and mark missing lemmas --- hierarchy.v | 129 +++++++++------------------------------------------- topology.v | 104 ++++++++++++++++++++++++++++++++++++------ 2 files changed, 112 insertions(+), 121 deletions(-) diff --git a/hierarchy.v b/hierarchy.v index 5af3818b64..74c43ef719 100644 --- a/hierarchy.v +++ b/hierarchy.v @@ -5,81 +5,19 @@ From mathcomp Require Import matrix interval zmodp. Require Import boolp reals classical_sets posnum topology. (******************************************************************************) -(* This file extends the topological hierarchy with metric-related notions: *) -(* balls, absolute values, norms. *) -(* *) -(* * Uniform spaces : *) -(* locally_ ball == neighbourhoods defined using balls *) -(* uniformType == interface type for uniform space *) -(* structure. We use here a pseudo-metric *) -(* definition of uniform space: a type *) -(* equipped with balls. *) -(* UniformMixin brefl bsym btriangle locb == builds the mixin for a *) -(* uniform space from the properties of *) -(* balls and the compatibility between *) -(* balls and locally. *) -(* UniformType T m == packs the uniform space mixin into a *) -(* uniformType. T must have a canonical *) -(* topologicalType structure. *) -(* [uniformType of T for cT] == T-clone of the uniformType structure cT. *) -(* [uniformType of T] == clone of a canonical uniformType *) -(* structure on T. *) -(* topologyOfBallMixin umixin == builds the mixin for a topological space *) -(* from a mixin for a uniform space. *) -(* ball x e == ball of center x and radius e. *) -(* close x y <-> x and y are arbitrarily close w.r.t. to *) -(* balls. *) -(* entourage == set of entourages defined by balls. An *) -(* entourage can be seen as a *) -(* "neighbourhood" of the diagonal set *) -(* D = {(x, x) | x in T}. *) -(* ball_set A e == set A extended with a band of width e *) -(* unif_cont f <-> f is uniformly continuous. *) -(* *) -(* * Rings with absolute value : *) -(* absRingType == interface type for a ring with absolute *) -(* value structure. *) -(* AbsRingMixin abs0 absN1 absD absM abseq0 == builds the mixin for a *) -(* ring with absolute value from the *) -(* algebraic properties of the absolute *) -(* value; the carrier type must have a *) -(* ringType structure. *) -(* [absRingType of T for cT] == T-clone of the absRingType structure cT. *) -(* [absRingType of T] == clone of a canonical absRingType *) -(* structure on T. *) -(* `|x| == the absolute value of x. *) -(* ball_ N == balls defined by the norm/absolute value *) -(* N. *) -(* locally_dist == neighbourhoods defined by a "distance" *) -(* function *) -(* *) -(* * Complete spaces : *) -(* cauchy_ex F <-> the set of sets F is a cauchy filter *) -(* (epsilon-delta definition). *) -(* cauchy F <-> the set of sets F is a cauchy filter *) -(* (using the near notations). *) -(* completeType == interface type for a complete uniform *) -(* space structure. *) -(* CompleteType T cvgCauchy == packs the proof that every proper cauchy *) -(* filter on T converges into a *) -(* completeType structure; T must have a *) -(* canonical uniformType structure. *) -(* [completeType of T for cT] == T-clone of the completeType structure *) -(* cT. *) -(* [completeType of T] == clone of a canonical completeType *) -(* structure on T. *) +(* This file extends the topological hierarchy with norm-related notions. *) (* *) (* * Normed modules : *) +(* ball N == balls defined by the norm N. *) +(* entourage_ N == entourages defined by the norm N. *) (* normedModType K == interface type for a normed module *) -(* structure over the ring with absolute *) -(* value K. *) -(* NormedModMixin normD normZ balln normeq0 == builds the mixin for a *) -(* normed module from the algebraic *) -(* properties of the norm and the *) -(* compatibility between the norm and *) -(* balls; the carrier type must have a *) -(* lmodType K structure for K an *) -(* absRingType. *) +(* structure over the numDomainType K. *) +(* NormedModMixin normD normZ entn normeq0 == builds the mixin for a normed *) +(* module from the algebraic properties of *) +(* the norm and the compatibility between *) +(* the norm and entourages; the carrier *) +(* type must have a lmodType K structure *) +(* for K a numDomaintype. *) (* NormedModType K T m == packs the mixin m to build a *) (* normedModType K; T must have canonical *) (* lmodType K and uniformType structures. *) @@ -88,8 +26,6 @@ Require Import boolp reals classical_sets posnum topology. (* [normedModType K of T] == clone of a canonical normedModType K *) (* structure on T. *) (* `|[x]| == the norm of x. *) -(* ball_norm == balls defined by the norm. *) -(* locally_norm == neighbourhoods defined by the norm. *) (* bounded == set of bounded sets. *) (* *) (* * Complete normed modules : *) @@ -100,20 +36,17 @@ Require Import boolp reals classical_sets posnum topology. (* module structure over K on T. *) (* *) (* * Filters : *) -(* \oo == "eventually" filter on nat: set of *) -(* predicates on natural numbers that are *) -(* eventually true. *) -(* at_left x, at_right x == filters on real numbers for predicates *) -(* that locally hold on the left/right of *) -(* x. *) -(* Rbar_locally' x == filter on extended real numbers that *) -(* corresponds to locally' x if x is a real *) -(* number and to predicates that are *) -(* eventually true if x is +oo/-oo. *) -(* Rbar_locally x == same as Rbar_locally' where locally' is *) -(* replaced with locally. *) -(* Rbar_loc_seq x == sequence that converges to x in the set *) -(* of extended real numbers. *) +(* ereal_locally' x == filter on an extended realFieldType R *) +(* that corresponds to locally' x if x : R *) +(* and to predicates that are eventually *) +(* true if x is +oo/-oo. *) +(* ereal_locally x == same as ereal_locally' where locally' *) +(* is replaced with locally. *) +(* at_left x, at_right x == filters on a realFieldType for *) +(* predicates that locally hold on the *) +(* left/right of x. *) +(* ereal_loc_seq x == sequence that converges to x in an *) +(* extended realType. *) (* *) (* --> We used these definitions to prove the intermediate value theorem and *) (* the Heine-Borel theorem, which states that the compact sets of R^n are *) @@ -1374,12 +1307,6 @@ Export CompleteNormedModule.Exports. (** * Extended Types *) -Definition entourage_set (U : uniformType) (A : set ((set U) * (set U))) := - exists2 B, entourage B & forall PQ, A PQ -> forall p q, - PQ.1 p -> PQ.2 q -> B (p,q). -Canonical set_filter_source (U : uniformType) := - @Filtered.Source Prop _ U (fun A => locally_ (@entourage_set U) A). - (** * The topology on real numbers *) (* :TODO: add to mathcomp *) @@ -2099,17 +2026,3 @@ rewrite [ `|[_]|]gtr0_norm; last by rewrite invr_gt0. rewrite -[X in X < _]mulr1 ltr_pdivr_mull // -ltr_pdivr_mulr // div1r. by apply: ltr_le_trans (floorS_gtr _) _; rewrite floorE Nfloor ler_add ?ler_nat. Qed. - -Lemma continuous_withinNx {U V : uniformType} - (f : U -> V) x : - {for x, continuous f} <-> f @ locally' x --> f x. -Proof. -split=> - cfx P /= fxP. - rewrite /locally' !near_simpl near_withinE. - by rewrite /locally'; apply: flim_within; apply/cfx. - (* :BUG: ssr apply: does not work, - because the type of the filter is not inferred *) -rewrite !locally_nearE !near_map !near_locally in fxP *; have /= := cfx P fxP. -rewrite !near_simpl near_withinE near_simpl => Pf; near=> y. -by have [->|] := eqVneq y x; [by apply: locally_singleton|near: y]. -Grab Existential Variables. all: end_near. Qed. diff --git a/topology.v b/topology.v index 1f18621f04..3af6fc8422 100644 --- a/topology.v +++ b/topology.v @@ -8,7 +8,7 @@ Require Import boolp classical_sets posnum. (* topological notions. *) (* *) (* * Filters : *) -(* filteredType U == interface type for types whose *) +(* filteredType U == interface type for types whose *) (* elements represent sets of sets on U. *) (* These sets are intended to be filters *) (* on U but this is not enforced yet. *) @@ -20,15 +20,6 @@ Require Import boolp classical_sets posnum. (* structure cT. *) (* [filteredType U of T] == clone of a canonical filteredType U *) (* structure on T. *) -(* filter_on_term X Y == structure that records terms x : X *) -(* with a set of sets (filter) in Y. *) -(* Allows to infer the canonical filter *) -(* associated to a term by looking at its *) -(* type. *) -(* filter_term F == if F : filter_on_term X Y, outputs the *) -(* term in X contained in F. *) -(* term_filter F == if F : filter_on_term X Y, outputs the *) -(* filter on Y contained in F. *) (* Filtered.source Y Z == structure that records types X such *) (* that there is a function mapping *) (* functions of type X -> Y to filters on *) @@ -38,7 +29,6 @@ Require Import boolp classical_sets posnum. (* Filtered.Source F == if F : (X -> Y) -> set (set Z), packs *) (* X with F in a Filtered.source Y Z *) (* structure. *) -(* [filter of x] == canonical filter associated to x. *) (* locally p == set of sets associated to p (in a *) (* filtered type). *) (* filter_from D B == set of the supersets of the elements *) @@ -47,6 +37,7 @@ Require Import boolp classical_sets posnum. (* This is a filter if (B_i)_(i in D) *) (* forms a filter base. *) (* filter_prod F G == product of the filters F and G. *) +(* [filter of x] == canonical filter associated to x. *) (* F `=>` G <-> G is included in F; F and G are sets *) (* of sets. *) (* F --> G <-> the canonical filter associated to G *) @@ -68,7 +59,17 @@ Require Import boolp classical_sets posnum. (* ProperFilter F == type class proving that the set of *) (* sets F is a proper filter. *) (* UltraFilter F == type class proving that the set of *) -(* sets F is an ultrafilter *) +(* sets F is an ultrafilter. *) +(* filter_on T == interface type for sets of sets on T *) +(* that are filters. *) +(* FilterType F FF == packs the set of sets F with the proof *) +(* FF of Filter F to build a filter_on T *) +(* structure. *) +(* pfilter_on T == interface type for sets of sets on T *) +(* that are proper filters. *) +(* PFilterPack F FF == packs the set of sets F with the proof *) +(* FF of ProperFilter F to build a *) +(* pfilter_on T structure. *) (* filtermap f F == image of the filter F by the function *) (* f. *) (* E @[x --> F] == image of the canonical filter *) @@ -88,6 +89,13 @@ Require Import boolp classical_sets posnum. (* domain D. *) (* subset_filter F D == similar to within D F, but with *) (* dependent types. *) +(* in_filter F == interface type for the sets that *) +(* belong to the set of sets F. *) +(* InFilter FP == packs a set P with a proof of F P to *) +(* build a in_filter F structure. *) +(* \oo == "eventually" filter on nat: set of *) +(* predicates on natural numbers that are *) +(* eventually true. *) (* *) (* * Near notations and tactics: *) (* --> The purpose of the near notations and tactics is to make the *) @@ -108,6 +116,7 @@ Require Import boolp classical_sets posnum. (* \forall x \near F & y \near G, P x y := {near F & G, forall x y, P x y}. *) (* \forall x & y \near F, P x y == same as before, with G = F. *) (* \near x & y, P x y := \forall z \near x & t \near y, P x y. *) +(* x \is_near F == x belongs to a set P : in_filter F. *) (* --> Tactics: *) (* - near=> x introduces x: *) (* On the goal \forall x \near F, G x, introduces the variable x and an *) @@ -180,10 +189,52 @@ Require Import boolp classical_sets posnum. (* cover-based definition of compactness. *) (* connected A <-> the only non empty subset of A which *) (* is both open and closed in A is A. *) -(* *) (* --> We used these topological notions to prove Tychonoff's Theorem, which *) (* states that any product of compact sets is compact according to the *) (* product topology. *) +(* *) +(* * Uniform spaces : *) +(* locally_ ent == neighbourhoods defined using entourages. *) +(* uniformType == interface type for uniform space *) +(* structure. We use here the entourage *) +(* definition of uniform space: a type *) +(* equipped with a set of entourages, also *) +(* called uniformity. An entourage can be *) +(* seen as a "neighbourhood" of the *) +(* diagonal set {(x, x) | x in T}. *) +(* UniformMixin entF ent_refl ent_inv ent_split loc_ent == builds the mixin *) +(* for a uniform space from the properties *) +(* of entourages and the compatibility *) +(* between entourages and locally. *) +(* UniformType T m == packs the uniform space mixin into a *) +(* uniformType. T must have a canonical *) +(* topologicalType structure. *) +(* [uniformType of T for cT] == T-clone of the uniformType structure cT. *) +(* [uniformType of T] == clone of a canonical uniformType *) +(* structure on T. *) +(* topologyOfEntourageMixin umixin == builds the mixin for a topological *) +(* space from a mixin for a uniform space. *) +(* entourage == set of entourages. *) +(* split_ent A == "half entourage": entourage B such that, *) +(* when seen as a relation, B \o B `<=` A. *) +(* close x y <-> x and y are arbitrarily close, i.e. the *) +(* pair (x,y) is in any entourage. *) +(* unif_cont f <-> f is uniformly continuous. *) +(* entourage_set == set of entourages on the parts of a *) +(* uniform space. *) +(* *) +(* * Complete spaces : *) +(* cauchy F <-> the set of sets F is a cauchy filter *) +(* completeType == interface type for a complete uniform *) +(* space structure. *) +(* CompleteType T cvgCauchy == packs the proof that every proper cauchy *) +(* filter on T converges into a *) +(* completeType structure; T must have a *) +(* canonical uniformType structure. *) +(* [completeType of T for cT] == T-clone of the completeType structure *) +(* cT. *) +(* [completeType of T] == clone of a canonical completeType *) +(* structure on T. *) (******************************************************************************) Set Implicit Arguments. @@ -2399,6 +2450,8 @@ Grab Existential Variables. all: end_near. Qed. Lemma flimx_close (x y : M) : x --> y -> close x y. Proof. exact: flim_close. Qed. +(* TODO: sections locally_fct, locally_fct2 in normed_spaces *) + Lemma flim_entourageP F (FF : Filter F) (p : M) : F --> p <-> forall A, entourage A -> \forall q \near F, A (p, q). Proof. by rewrite -filter_fromP !locally_simpl. Qed. @@ -2470,9 +2523,26 @@ Hint Resolve close_refl. Arguments flim_const {M T F FF} a. Arguments close_lim {M} F1 F2 {FF2} _. +Lemma continuous_withinNx {U V : uniformType} + (f : U -> V) x : + {for x, continuous f} <-> f @ locally' x --> f x. +Proof. +split=> - cfx P /= fxP. + rewrite /locally' !near_simpl near_withinE. + by rewrite /locally'; apply: flim_within; apply/cfx. + (* :BUG: ssr apply: does not work, + because the type of the filter is not inferred *) +rewrite !locally_nearE !near_map !near_locally in fxP *; have /= := cfx P fxP. +rewrite !near_simpl near_withinE near_simpl => Pf; near=> y. +by have [->|] := eqVneq y x; [by apply: locally_singleton|near: y]. +Grab Existential Variables. all: end_near. Qed. + Definition unif_cont (U V : uniformType) (f : U -> V) := (fun xy => (f xy.1, f xy.2)) @ entourage --> entourage. +(* TODO: unif_contP in normed spaces *) +(* Also missing: section locally, filterP_strong *) + (** product of two uniform spaces *) Section prod_Uniform. @@ -2734,10 +2804,18 @@ move=> Ff; apply/flim_entourageP => A [P entP sPA]; near=> g. by apply: sPA => /=; near: g; apply: Ff. Grab Existential Variables. all: end_near. Qed. +Definition entourage_set (U : uniformType) (A : set ((set U) * (set U))) := + exists2 B, entourage B & forall PQ, A PQ -> forall p q, + PQ.1 p -> PQ.2 q -> B (p,q). +Canonical set_filter_source (U : uniformType) := + @Filtered.Source Prop _ U (fun A => locally_ (@entourage_set U) A). + (** ** Complete uniform spaces *) Definition cauchy {T : uniformType} (F : set (set T)) := (F, F) --> entourage. +(* TODO: cauchy_ex and old cauchy in normed spaces *) + Lemma cvg_cauchy {T : uniformType} (F : set (set T)) : Filter F -> [cvg F in T] -> cauchy F. Proof. From 909029df85097c95307922976f6aba213ae6f397 Mon Sep 17 00:00:00 2001 From: damien rouhling Date: Fri, 12 Oct 2018 16:18:55 +0200 Subject: [PATCH 11/11] Reintroduce missing lemmas + more clean up --- derive.v | 31 ++++----- hierarchy.v | 177 +++++++++++++++++++++++++++++++++++++++++++++++++--- topology.v | 14 ++--- 3 files changed, 189 insertions(+), 33 deletions(-) diff --git a/derive.v b/derive.v index aea1aeda70..451fe195be 100644 --- a/derive.v +++ b/derive.v @@ -283,9 +283,8 @@ pose g1 : R -> W := fun h => (h^-1 * h) *: 'd f a v. pose g2 : R -> W := fun h : R => h^-1 *: k (h *: v ). rewrite (_ : g = g1 + g2) ?funeqE // -(addr0 (_ _ v)); apply: lim_add. rewrite -(scale1r (_ _ v)); apply: lim_scalel. - apply/app_flim_entourageP => X entX; apply/locallyP. - rewrite locally_E; exists X => // x _ x0; rewrite mulVf //. - exact: entourage_refl. + apply/app_flim_locally => _/posnumP[e]; apply/locally_normP. + by exists e%:num => // x _ x0; rewrite mulVf. rewrite /g2. have [/eqP ->|v0] := boolP (v == 0). rewrite (_ : (fun _ => _) = cst 0); first exact: cst_continuous. @@ -639,16 +638,14 @@ Lemma compoO_eqo (K : realFieldType) (U V' W' : normedModType K) (f : U -> V') (g : V' -> W') : [o_ (0 : V') id of g] \o [O_ (0 : U) id of f] =o_ (0 : U) id. Proof. -apply/eqoP => _ /posnumP[e]. -have /bigO_exP [_ /posnumP[k]] := bigOP [bigO of [O_ (0 : U) id of f]]. -have /(_ (e%:num / k%:num)) := littleoP [littleo of [o_ (0 : V') id of g]]. -move=> /(_ _) /locally_normP [//|_/posnumP[d] hd]. -apply: filter_app; near=> x => leOxkx; apply: ler_trans (hd _ _) _; last first. - rewrite -ler_pdivl_mull //; apply: ler_trans leOxkx _. - by rewrite invf_div mulrA -[_ / _ * _]mulrA mulVf // mulr1. -rewrite /= normmB subr0 (ler_lt_trans leOxkx) // -ltr_pdivl_mull //; near: x. -apply/locally_normP; exists (k%:num ^-1 * d%:num) => // x. -by rewrite /= normmB subr0. +apply/eqoP => _ /posnumP[e]; have [go] := littleo; have [fO k fOid] := bigO. +move=> /(_ (e%:num / k%:num) _) /locally_normP [//|_/posnumP[d] hd]. +apply: filter_app fOid; near=> x => leOxkx; apply: ler_trans (hd _ _) _. + rewrite /= normmB subr0 (ler_lt_trans leOxkx) // -ltr_pdivl_mull //; near: x. + apply/locally_normP; exists (k%:num ^-1 * d%:num) => // x. + by rewrite /= normmB subr0. +rewrite -ler_pdivl_mull //; apply: ler_trans leOxkx _. +by rewrite invf_div mulrA -[_ / _ * _]mulrA mulVf // mulr1. Grab Existential Variables. all: end_near. Qed. Lemma compoO_eqox (K : realFieldType) (U V' W' : normedModType K) (f : U -> V') @@ -661,11 +658,9 @@ Lemma compOo_eqo (K : realFieldType) (U V' W' : normedModType K) (f : U -> V') (g : V' -> W') : [O_ (0 : V') id of g] \o [o_ (0 : U) id of f] =o_ (0 : U) id. Proof. -apply/eqoP => _ /posnumP[e]. -have /bigO_exP [_ /posnumP[k]] := bigOP [bigO of [O_ (0 : V') id of g]]. -move=> /locally_normP [_/posnumP[d] hd]; have ekgt0 : e%:num / k%:num > 0 by []. -have /(_ _ ekgt0) := littleoP [littleo of [o_ (0 : U) id of f]]. -apply: filter_app; near=> x => leoxekx; apply: ler_trans (hd _ _) _; last first. +apply/eqoP => _/posnumP[e]; have [gO k /locally_normP[_/posnumP[d] hd]] := bigO. +have [fo /(_ _ [gt0 of e%:num / k%:num])] := littleo; apply: filter_app. +near=> x => leoxekx; apply: ler_trans (hd _ _) _; last first. by rewrite -ler_pdivl_mull // mulrA [_^-1 * _]mulrC. rewrite /ball normmB subr0; apply: ler_lt_trans leoxekx _. rewrite -ltr_pdivl_mull //; near: x; apply/locally_normP. diff --git a/hierarchy.v b/hierarchy.v index 74c43ef719..90d6a1c7bb 100644 --- a/hierarchy.v +++ b/hierarchy.v @@ -368,6 +368,118 @@ apply; have /flim_norm /(_ _ [gt0 of d%:num]) := xl. by move=> /locally_normP [_/posnumP[d']]; apply. Qed. +Section Locally. + +Context {K : numDomainType} {V : normedModType K}. + +Lemma forallN {U} (P : set U) : (forall x, ~ P x) = ~ exists x, P x. +Proof. (*boolP*) +rewrite propeqE; split; first by move=> fP [x /fP]. +by move=> nexP x Px; apply: nexP; exists x. +Qed. + +Lemma eqNNP (P : Prop) : (~ ~ P) = P. (*boolP*) +Proof. by rewrite propeqE; split=> [/contrapT|?]. Qed. + +Lemma existsN {U} (P : set U) : (exists x, ~ P x) = ~ forall x, P x. (*boolP*) +Proof. +rewrite propeqE; split=> [[x Px] Nall|Nall]; first exact: Px. +by apply: contrapT; rewrite -forallN => allP; apply: Nall => x; apply: contrapT. +Qed. + +Lemma ex_ball_sig (x : V) (P : set V) : + ~ (forall eps : {posnum K}, ~ (ball norm x eps%:num `<=` ~` P)) -> + {d : {posnum K} | ball norm x d%:num `<=` ~` P}. +Proof. +rewrite forallN eqNNP => exNP. +pose D := [set d : K | d > 0 /\ ball norm x d `<=` ~` P]. +have [|d_gt0] := @getPex _ D; last by exists (PosNum d_gt0). +by move: exNP => [e eP]; exists e%:num. +Qed. + +Lemma locallyC (x : V) (P : set V) : + ~ (forall eps : {posnum K}, ~ (ball norm x eps%:num `<=` ~` P)) -> + locally x (~` P). +Proof. by move=> /ex_ball_sig [e] ?; apply/locally_normP; exists e%:num. Qed. + +Lemma locallyC_ball (x : V) (P : set V) : + locally x (~` P) -> {d : {posnum K} | ball norm x d%:num `<=` ~` P}. +Proof. +move=> /locally_normP xNP; apply: ex_ball_sig. +by have [_ /posnumP[e] eP /(_ _ eP)] := xNP. +Qed. + +Lemma locally_ex (x : V) (P : V -> Prop) : locally x P -> + {d : {posnum K} | forall y, ball norm x d%:num y -> P y}. +Proof. +move=> /locally_normP xP. +pose D := [set d : K | d > 0 /\ forall y, ball norm x d y -> P y]. +have [|d_gt0 dP] := @getPex _ D; last by exists (PosNum d_gt0). +by move: xP => [e bP]; exists e. +Qed. + +End Locally. + +Lemma unif_contP (K K' : numDomainType) (V : normedModType K) + (V' : normedModType K') (f : V -> V') : + unif_cont f <-> + forall e, e > 0 -> exists2 d, d > 0 & + forall x, ball norm x.1 d x.2 -> ball norm (f x.1) e (f x.2). +Proof. +have fappF : Filter ((fun xy => (f xy.1, f xy.2)) @ entourage_ norm). + by rewrite entourage_normE; apply: filtermap_filter. +by rewrite /unif_cont -!entourage_normE filter_fromP. +Qed. + +Section Locally_fct. + +Context {T : Type} {K : numDomainType} {V : normedModType K}. + +Lemma near_ball (y : V) (eps : {posnum K}) : + \forall y' \near y, ball norm y eps%:num y'. +Proof. exact: locally_ball. Qed. + +Lemma flim_ballP {F} {FF : Filter F} (y : V) : + F --> y <-> forall eps, 0 < eps -> \forall y' \near F, ball norm y eps y'. +Proof. exact: flim_normP. Qed. +Definition flim_locally := @flim_ballP. + +Lemma flim_ballPpos {F} {FF : Filter F} (y : V) : + F --> y <-> + forall eps : {posnum K}, \forall y' \near F, ball norm y eps%:num y'. +Proof. +by split => [/flim_ballP|] pos; [case|apply/flim_ballP=> _/posnumP[eps] //]. +Qed. + +Lemma flim_ball {F} {FF : Filter F} (y : V) : + F --> y -> forall eps, 0 < eps -> \forall y' \near F, ball norm y eps y'. +Proof. by move/flim_ballP. Qed. + +Lemma app_flim_locally {F} {FF : Filter F} (f : T -> V) y : + f @ F --> y <-> + forall eps, 0 < eps -> \forall x \near F, ball norm y eps (f x). +Proof. exact: flim_ballP. Qed. + +Lemma flimi_ballP {F} {FF : Filter F} (f : T -> V -> Prop) y : + f `@ F --> y <-> + forall eps, 0 < eps -> + \forall x \near F, exists z, f x z /\ ball norm y eps z. +Proof. +split=> [Fy _/posnumP[eps] |Fy P] /=; first exact/Fy/locally_ball. +move=> /locally_normP[_ /posnumP[eps] subP]. +rewrite near_simpl near_mapi; near=> x. +have [//|z [fxz yz]] := near (Fy _ (posnum_gt0 eps)) x. +by exists z => //; split => //; apply: subP. +Unshelve. all: end_near. Qed. +Definition flimi_locally := @flimi_ballP. + +Lemma flimi_ball {F} {FF : Filter F} (f : T -> V -> Prop) y : + f `@ F --> y -> + forall eps, 0 < eps -> F [set x | exists z, f x z /\ ball norm y eps z]. +Proof. by move/flimi_ballP. Qed. + +End Locally_fct. + Section NormedModuleField. Context {K : numFieldType} {V : normedModType K}. @@ -962,15 +1074,14 @@ rewrite predeqE => A; split; last first. by move=> _ _; exists e%:num. move=> [x y] /= xy; apply: sA => /=. by apply/bigmaxr_ltrP; split=> // ij _; rewrite !mxE. -move=> [P entP sPA]; have {entP} entP : forall i j, exists e, 0 < e /\ - [set pq | `|pq.1 - pq.2| < e] `<=` P i j. - by move=> i j; have [e] := entP i j; exists e. -set e := fun i j => get [set e | 0 < e /\ +move=> [P entP sPA]; set sP := fun i j => [set e | 0 < e /\ [set pq | `|pq.1 - pq.2| < e] `<=` P i j]. -exists (\big[minr/1]_ij e ij.1 ij.2). +have {entP} entP : forall i j, sP i j !=set0. + by move=> i j; have [e] := entP i j; exists e. +exists (\big[minr/1]_ij get (sP ij.1 ij.2)). by apply/bigminr_gtrP; split=> // ij _; have /getPex [] := entP ij.1 ij.2. move=> [x y] /= /bigminr_gtrP [_ xy]; apply: sPA => i j /=. -have /getPex [_] := entP i j; apply=> /=; rewrite -[get _]/(e i j). +have /getPex [_] := entP i j; apply => /=. by have /bigmaxr_ltrP [_ /(_ (i,j))] := xy (i,j) erefl; rewrite !mxE; apply. Qed. Next Obligation. @@ -1053,8 +1164,7 @@ Canonical prod_NormedModule (K : realDomainType) (U V : normedModType K) := Section NormedModule3. -Context {T : Type} {K : realDomainType} {U : normedModType K} - {V : normedModType K}. +Context {T : Type} {K : realDomainType} {U V : normedModType K}. Lemma flim_norm2P {F : set (set U)} {G : set (set V)} {FF : Filter F} {FG : Filter G} (y : U) (z : V): @@ -1081,6 +1191,17 @@ Lemma flim_norm2 {F : set (set U)} {G : set (set V)} \forall y' \near F & z' \near G, `|[(y, z) - (y', z')]| < e. Proof. by rewrite flim_normP. Qed. +Lemma flim_ball2P {F : set (set U)} {G : set (set V)} + {FF : Filter F} {FG : Filter G} (y : U) (z : V): + (F, G) --> (y, z) <-> + forall eps, eps > 0 -> \forall y' \near F & z' \near G, + ball norm y eps y' /\ ball norm z eps z'. +Proof. +by rewrite flim_norm2P; split=> FGyz e /FGyz; + apply: filter_app; apply: filterE => x; rewrite ltr_maxl; + [move=> /andP[] | move=> [-> ->]]. +Qed. + End NormedModule3. Arguments flim_norm2 {_ _ _ F G FF FG}. @@ -1223,6 +1344,46 @@ End limit_composition. (** ** Complete Normed Modules *) +Section Cauchy. + +Context {K : numDomainType} {V : normedModType K}. + +Definition cauchy_ball (F : set (set V)) := + forall e, e > 0 -> \forall x & y \near F, ball norm x e y. + +Lemma cauchy_cauchy_ball (F : set (set V)) : cauchy F -> cauchy_ball F. +Proof. by move=> Fc _/posnumP[e]; apply: Fc (entourage_ball e). Qed. + +Lemma cauchy_ballP (F : set (set V)) : Filter F -> cauchy_ball F <-> cauchy F. +Proof. +move=> FF; split=> [Fc A|/cauchy_cauchy_ball] //. +by rewrite -entourage_normE => -[_/posnumP[e] sA]; apply: filterS sA (Fc _ _). +Qed. + +Definition cauchy_ex (F : set (set V)) := + forall eps, 0 < eps -> exists x, F (ball norm x eps). + +Lemma cvg_cauchy_ex (F : set (set V)) : [cvg F in V] -> cauchy_ex F. +Proof. by move=> Fl _/posnumP[e]; exists (lim F); apply/Fl/locally_ball. Qed. + +End Cauchy. + +Lemma cauchy_exP (K : numFieldType) (V : normedModType K) (F : set (set V)) : + Filter F -> cauchy_ex F -> cauchy F. +Proof. +move=> FF Fc A; rewrite -entourage_normE => -[_/posnumP[e] sA]. +have /Fc [z /= Fze] := [gt0 of e%:num / 2]; near=> x y; apply: sA => /=. +by apply: (@distm_lt_splitr _ _ z); [near: x|near: y]. +Grab Existential Variables. all: end_near. Qed. + +Lemma cauchyP (K : numFieldType) (V : normedModType K) (F : set (set V)) : + ProperFilter F -> cauchy F <-> cauchy_ex F. +Proof. +move=> FF; split=> [Fcauchy _/posnumP[e] |/cauchy_exP//]. +near F => x; exists x; near: x; apply: (@nearP_dep _ _ F F); apply: Fcauchy. +exact: entourage_ball. +Grab Existential Variables. all: end_near. Qed. + Module CompleteNormedModule. Section ClassDef. diff --git a/topology.v b/topology.v index 3af6fc8422..1d899c3bee 100644 --- a/topology.v +++ b/topology.v @@ -509,6 +509,13 @@ Notation ProperFilter := ProperFilter'. Lemma filter_setT (T' : Type) : Filter (@setT (set T')). Proof. by constructor. Qed. +Lemma filterP_strong T (F : set (set T)) {FF : Filter F} (P : set T) : + (exists Q : set T, exists FQ : F Q, forall x : T, Q x -> P x) <-> F P. +Proof. +split; last by exists P. +by move=> [Q [FQ QP]]; apply: (filterS QP). +Qed. + Lemma filter_bigI T (I : choiceType) (D : {fset I}) (f : I -> set T) (F : set (set T)) : Filter F -> (forall i, i \in D -> F (f i)) -> @@ -2450,8 +2457,6 @@ Grab Existential Variables. all: end_near. Qed. Lemma flimx_close (x y : M) : x --> y -> close x y. Proof. exact: flim_close. Qed. -(* TODO: sections locally_fct, locally_fct2 in normed_spaces *) - Lemma flim_entourageP F (FF : Filter F) (p : M) : F --> p <-> forall A, entourage A -> \forall q \near F, A (p, q). Proof. by rewrite -filter_fromP !locally_simpl. Qed. @@ -2540,9 +2545,6 @@ Grab Existential Variables. all: end_near. Qed. Definition unif_cont (U V : uniformType) (f : U -> V) := (fun xy => (f xy.1, f xy.2)) @ entourage --> entourage. -(* TODO: unif_contP in normed spaces *) -(* Also missing: section locally, filterP_strong *) - (** product of two uniform spaces *) Section prod_Uniform. @@ -2814,8 +2816,6 @@ Canonical set_filter_source (U : uniformType) := Definition cauchy {T : uniformType} (F : set (set T)) := (F, F) --> entourage. -(* TODO: cauchy_ex and old cauchy in normed spaces *) - Lemma cvg_cauchy {T : uniformType} (F : set (set T)) : Filter F -> [cvg F in T] -> cauchy F. Proof.