Module mathcomp.analysis.topology_theory.num_topology
From HB Require Import structures.From mathcomp Require Import all_ssreflect all_algebra archimedean.
From mathcomp Require Import all_classical.
From mathcomp Require Import interval_inference reals topology_structure.
From mathcomp Require Import uniform_structure pseudometric_structure.
From mathcomp Require Import order_topology.
# Topological notions for numerical types
We endow `numFieldType` with the types of topological notions (accessible
with `Import numFieldTopology.Exports).
Import Order.TTheory GRing.Theory Num.Theory.
Set Implicit Arguments.
Unset Strict Implicit.
Unset Printing Implicit Defensive.
Local Open Scope classical_set_scope.
Local Open Scope ring_scope.
HB.instance Definition _ (R : numDomainType) := hasNbhs.Build R^o
(nbhs_ball_ (ball_ (fun x => `|x|))).
HB.instance Definition _ (R : numFieldType) :=
Nbhs_isPseudoMetric.Build R R^o
nbhs_ball_normE ball_norm_center ball_norm_symmetric ball_norm_triangle erefl.
Lemma real_order_nbhsE (R : realFieldType) (x : R^o) : nbhs x =
filter_from (fun i => itv_open_ends i /\ x \in i) (fun i => [set` i]).
Proof.
rewrite eqEsubset; split => U.
case => _ /posnumP[e] xeU.
exists (`]x - e%:num, x + e%:num[); first split; first by right.
by rewrite in_itv /= -real_lter_distl subrr // normr0.
apply: (subset_trans _ xeU) => z /=.
by rewrite in_itv /= -real_lter_distl //= distrC.
case => [][[[]l|[]]] [[]r|[]] [[]]//= _.
- move=> xlr lrU; exists (Order.min (x - l) (r - x)).
by rewrite /= lt_min ?lterBDr ?add0r ?(itvP xlr).
apply/(subset_trans _ lrU)/subset_ball_prop_in_itv.
suff : (`]x - Order.min (x - l) (r - x), x + Order.min (x - l) (r - x)[
<= `]l, r[)%O by move/subitvP => H ? ?; exact: H.
rewrite subitvE 2!lteBSide/=.
by rewrite lerBrDl [_ + l]addrC -2!lerBrDl 2!ge_min 2!lexx orbT.
- move=> xl lU; exists (x - l) => /=; first by rewrite lterBDr add0r (itvP xl).
apply/(subset_trans _ lU)/subset_ball_prop_in_itv.
suff : (`]x - (x - l), x + (x - l)[ <= `]l, +oo[)%O.
by move/subitvP => H ?; exact: H.
by rewrite subitvE lteBSide/= subKr lexx.
- move=> xr rU; exists (r - x) => /=; first by rewrite lterBDr add0r (itvP xr).
apply/(subset_trans _ rU)/subset_ball_prop_in_itv.
suff : (`]x - (r - x), x + (r - x)[ <= `]-oo, r[)%O.
by move/subitvP => H ?; exact: H.
by rewrite subitvE lteBSide/= addrC subrK.
- by move=> _; rewrite set_itvE subTset => ->; exists 1 => /=.
Qed.
case => _ /posnumP[e] xeU.
exists (`]x - e%:num, x + e%:num[); first split; first by right.
by rewrite in_itv /= -real_lter_distl subrr // normr0.
apply: (subset_trans _ xeU) => z /=.
by rewrite in_itv /= -real_lter_distl //= distrC.
case => [][[[]l|[]]] [[]r|[]] [[]]//= _.
- move=> xlr lrU; exists (Order.min (x - l) (r - x)).
by rewrite /= lt_min ?lterBDr ?add0r ?(itvP xlr).
apply/(subset_trans _ lrU)/subset_ball_prop_in_itv.
suff : (`]x - Order.min (x - l) (r - x), x + Order.min (x - l) (r - x)[
<= `]l, r[)%O by move/subitvP => H ? ?; exact: H.
rewrite subitvE 2!lteBSide/=.
by rewrite lerBrDl [_ + l]addrC -2!lerBrDl 2!ge_min 2!lexx orbT.
- move=> xl lU; exists (x - l) => /=; first by rewrite lterBDr add0r (itvP xl).
apply/(subset_trans _ lU)/subset_ball_prop_in_itv.
suff : (`]x - (x - l), x + (x - l)[ <= `]l, +oo[)%O.
by move/subitvP => H ?; exact: H.
by rewrite subitvE lteBSide/= subKr lexx.
- move=> xr rU; exists (r - x) => /=; first by rewrite lterBDr add0r (itvP xr).
apply/(subset_trans _ rU)/subset_ball_prop_in_itv.
suff : (`]x - (r - x), x + (r - x)[ <= `]-oo, r[)%O.
by move/subitvP => H ?; exact: H.
by rewrite subitvE lteBSide/= addrC subrK.
- by move=> _; rewrite set_itvE subTset => ->; exists 1 => /=.
Qed.
Module numFieldTopology.
#[export, non_forgetful_inheritance]
HB.instance Definition _ (R : realType) := PseudoPointedMetric.copy R R^o.
#[export, non_forgetful_inheritance]
HB.instance Definition _ (R : rcfType) := PseudoPointedMetric.copy R R^o.
#[export, non_forgetful_inheritance]
HB.instance Definition _ (R : archiFieldType) := PseudoPointedMetric.copy R R^o.
#[export, non_forgetful_inheritance]
HB.instance Definition _ (R : realFieldType) := PseudoPointedMetric.copy R R^o.
#[export, non_forgetful_inheritance]
HB.instance Definition _ (R : numClosedFieldType) :=
PseudoPointedMetric.copy R R^o.
#[export, non_forgetful_inheritance]
HB.instance Definition _ (R : realFieldType) :=
Order_isNbhs.Build _ R (@real_order_nbhsE R).
#[export, non_forgetful_inheritance]
HB.instance Definition _ (R : numFieldType) := PseudoPointedMetric.copy R R^o.
Module Exports. HB.reexport. End Exports.
End numFieldTopology.
Import numFieldTopology.Exports.
Lemma closure_sup (R : realType) (A : set R) :
A !=set0 -> has_ubound A -> closure A (sup A).
Proof.
move=> A0 ?; have [|AsupA] := pselect (A (sup A)); first exact: subset_closure.
rewrite closure_limit_point; right => U /nbhs_ballP[_ /posnumP[e]] supAeU.
suff [x [Ax /andP[sAex xsA]]] : exists x, A x /\ sup A - e%:num < x < sup A.
exists x; split => //; first by rewrite lt_eqF.
apply supAeU; rewrite /ball /= ltr_distl (addrC x e%:num) -ltrBlDl sAex.
by rewrite andbT (le_lt_trans _ xsA) // lerBlDl lerDr.
apply: contrapT => /forallNP Ax.
suff /(sup_le_ub A0) : ubound A (sup A - e%:num).
by rewrite leNgt => /negP; apply; rewrite ltrBlDl ltrDr.
move=> y Ay; have /not_andP[//|/negP] := Ax y.
rewrite negb_and leNgt => /orP[//|]; apply: contra => sAey.
rewrite lt_neqAle sup_upper_bound // andbT.
by apply: contra_not_neq AsupA => <-.
Qed.
rewrite closure_limit_point; right => U /nbhs_ballP[_ /posnumP[e]] supAeU.
suff [x [Ax /andP[sAex xsA]]] : exists x, A x /\ sup A - e%:num < x < sup A.
exists x; split => //; first by rewrite lt_eqF.
apply supAeU; rewrite /ball /= ltr_distl (addrC x e%:num) -ltrBlDl sAex.
by rewrite andbT (le_lt_trans _ xsA) // lerBlDl lerDr.
apply: contrapT => /forallNP Ax.
suff /(sup_le_ub A0) : ubound A (sup A - e%:num).
by rewrite leNgt => /negP; apply; rewrite ltrBlDl ltrDr.
move=> y Ay; have /not_andP[//|/negP] := Ax y.
rewrite negb_and leNgt => /orP[//|]; apply: contra => sAey.
rewrite lt_neqAle sup_upper_bound // andbT.
by apply: contra_not_neq AsupA => <-.
Qed.
Lemma right_bounded_interior (R : realType) (X : set R) :
has_ubound X -> X^° `<=` [set r | r < sup X].
Proof.
move=> uX r Xr; rewrite /mkset ltNge; apply/negP.
rewrite le_eqVlt => /orP[/eqP supXr|]; last first.
by apply/negP; rewrite -leNgt sup_ubound//; exact: interior_subset.
suff : ~ X^° (sup X) by rewrite supXr.
case/nbhs_ballP => _/posnumP[e] supXeX.
have [f XsupXf] : exists f : {posnum R}, X (sup X + f%:num).
exists (e%:num / 2)%:pos; apply supXeX; rewrite /ball /= opprD addNKr normrN.
by rewrite gtr0_norm // ltr_pdivrMr // ltr_pMr // ltr1n.
have : sup X + f%:num <= sup X by exact: sup_ubound.
by apply/negP; rewrite -ltNge; rewrite ltrDl.
Qed.
rewrite le_eqVlt => /orP[/eqP supXr|]; last first.
by apply/negP; rewrite -leNgt sup_ubound//; exact: interior_subset.
suff : ~ X^° (sup X) by rewrite supXr.
case/nbhs_ballP => _/posnumP[e] supXeX.
have [f XsupXf] : exists f : {posnum R}, X (sup X + f%:num).
exists (e%:num / 2)%:pos; apply supXeX; rewrite /ball /= opprD addNKr normrN.
by rewrite gtr0_norm // ltr_pdivrMr // ltr_pMr // ltr1n.
have : sup X + f%:num <= sup X by exact: sup_ubound.
by apply/negP; rewrite -ltNge; rewrite ltrDl.
Qed.
Lemma left_bounded_interior (R : realType) (X : set R) :
has_lbound X -> X^° `<=` [set r | inf X < r].
Proof.
move=> lX r Xr; rewrite /mkset ltNge; apply/negP.
rewrite le_eqVlt => /orP[/eqP rinfX|]; last first.
by apply/negP; rewrite -leNgt inf_lbound//; exact: interior_subset.
suff : ~ X^° (inf X) by rewrite -rinfX.
case/nbhs_ballP => _/posnumP[e] supXeX.
have [f XsupXf] : exists f : {posnum R}, X (inf X - f%:num).
exists (e%:num / 2)%:pos; apply supXeX; rewrite /ball /= opprB addrC subrK.
by rewrite gtr0_norm // ltr_pdivrMr // ltr_pMr // ltr1n.
have : inf X <= inf X - f%:num by exact: inf_lbound.
by apply/negP; rewrite -ltNge; rewrite ltrBlDr ltrDl.
Qed.
rewrite le_eqVlt => /orP[/eqP rinfX|]; last first.
by apply/negP; rewrite -leNgt inf_lbound//; exact: interior_subset.
suff : ~ X^° (inf X) by rewrite -rinfX.
case/nbhs_ballP => _/posnumP[e] supXeX.
have [f XsupXf] : exists f : {posnum R}, X (inf X - f%:num).
exists (e%:num / 2)%:pos; apply supXeX; rewrite /ball /= opprB addrC subrK.
by rewrite gtr0_norm // ltr_pdivrMr // ltr_pMr // ltr1n.
have : inf X <= inf X - f%:num by exact: inf_lbound.
by apply/negP; rewrite -ltNge; rewrite ltrBlDr ltrDl.
Qed.
Lemma nbhsN {R : numFieldType} (x : R) : nbhs (- x) = -%R @ x.
Proof.
Lemma cvg_compNP {T : topologicalType} {R : numFieldType} (f : R -> T) (a : R)
(l : T) :
(f \o -%R) x @[x --> a] --> l <-> f x @[x --> (- a)] --> l.
Proof.
Lemma nbhsNimage {R : numFieldType} (x : R) :
nbhs (- x) = [set -%R @` A | A in nbhs x].
Proof.
Lemma nearN {R : numFieldType} (x : R) (P : R -> Prop) :
(\forall y \near - x, P y) <-> \near x, P (- x).
Proof.
Lemma openN {R : numFieldType} (A : set R) : open A -> open [set - x | x in A].
Proof.
Lemma closedN (R : numFieldType) (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 nbhsNimage; exists B.
move=> [y [[z Az oppzey] [t Bt opptey]]]; exists (- y).
by split; [rewrite -oppzey opprK|rewrite -opptey opprK].
Qed.
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 nbhsNimage; exists B.
move=> [y [[z Az oppzey] [t Bt opptey]]]; exists (- y).
by split; [rewrite -oppzey opprK|rewrite -opptey opprK].
Qed.
Lemma dnbhsN {R : numFieldType} (r : R) :
(- r)%R^' = (fun A => -%R @` A) @` r^'.
Proof.
apply/seteqP; split=> [A [e/= e0 reA]|_/= [A [e/= e0 reA <-]]].
exists (-%R @` A).
exists e => // x/= rxe xr; exists (- x)%R; rewrite ?opprK//.
by apply: reA; rewrite ?eqr_opp//= opprK addrC distrC.
rewrite image_comp (_ : _ \o _ = idfun) ?image_id// funeqE => x/=.
by rewrite opprK.
exists e => //= x/=; rewrite -opprD normrN => axe xa.
exists (- x)%R; rewrite ?opprK//; apply: reA; rewrite ?eqr_oppLR//=.
by rewrite opprK.
Qed.
exists (-%R @` A).
exists e => // x/= rxe xr; exists (- x)%R; rewrite ?opprK//.
by apply: reA; rewrite ?eqr_opp//= opprK addrC distrC.
rewrite image_comp (_ : _ \o _ = idfun) ?image_id// funeqE => x/=.
by rewrite opprK.
exists e => //= x/=; rewrite -opprD normrN => axe xa.
exists (- x)%R; rewrite ?opprK//; apply: reA; rewrite ?eqr_oppLR//=.
by rewrite opprK.
Qed.
Lemma withinN {R : numFieldType} (A : set R) (r : R) :
within A (nbhs (- r)) = - x @[x --> within (-%R @` A) (nbhs r)].
Proof.
Lemma in_continuous_mksetP {T : realFieldType} {U : realFieldType}
(i : interval T) (f : T -> U) :
{in i, continuous f} <-> {in [set` i], continuous f}.
Lemma nbhs0_ltW (R : realFieldType) (x : R) : (0 < x)%R ->
\forall r \near nbhs (0%R:R), (r <= x)%R.
Proof.
Lemma nbhs0_lt (R : realFieldType) (x : R) : (0 < x)%R ->
\forall r \near nbhs (0%R:R), (r < x)%R.
Proof.
Lemma lt_nbhsl {R : realType} (x a : R) : x < a ->
\forall y \near nbhs x, y < a.
Proof.
Lemma Nlt_nbhsl {R : realType} (x a : R) :
- x < a -> \forall y \near nbhs x, - y < a.
Proof.
Section nbhs_lt_le.
Context {R : numFieldType}.
Implicit Types x z : R.
Lemma lt_nbhsl_lt x z : x < z -> \forall y \near x, x <= y -> y < z.
Proof.
Lemma lt_nbhsl_le x z : x < z -> \forall y \near x, x <= y -> y <= z.
Proof.
End nbhs_lt_le.
#[deprecated(since="mathcomp-analysis 1.9.0", note="use `lt_nbhsl_lt` instead")]
Notation nbhs_lt := lt_nbhsl_lt (only parsing).
#[deprecated(since="mathcomp-analysis 1.9.0", note="use `lt_nbhsl_le` instead")]
Notation nbhs_le := lt_nbhsl_le (only parsing).
Lemma lt_nbhsr {R : realFieldType} (a z : R) : a < z ->
\forall x \near z, a < x.
Proof.
Global Instance Proper_dnbhs_regular_numFieldType (R : numFieldType) (x : R^o) :
ProperFilter x^'.
Proof.
Global Instance Proper_dnbhs_numFieldType (R : numFieldType) (x : R) :
ProperFilter x^'.
Proof.
Lemma dense_rat (R : realType) : dense (@ratr R @` setT).
Proof.
move=> A [r Ar]; rewrite openE => /(_ _ Ar)/nbhs_ballP[_/posnumP[e] reA].
have /rat_in_itvoo[q /itvP qre] : r < r + e%:num by rewrite ltrDl.
exists (ratr q) => //; split; last by exists q.
apply: reA; rewrite /ball /= distrC ltr_distl qre andbT.
by rewrite (@le_lt_trans _ _ r)// ?qre// lerBlDl lerDr ltW.
Qed.
have /rat_in_itvoo[q /itvP qre] : r < r + e%:num by rewrite ltrDl.
exists (ratr q) => //; split; last by exists q.
apply: reA; rewrite /ball /= distrC ltr_distl qre andbT.
by rewrite (@le_lt_trans _ _ r)// ?qre// lerBlDl lerDr ltW.
Qed.
Lemma separated_open_countable
{R : realType} (I : Type) (B : I -> set R) (D : set I) :
(forall i, open (B i)) -> (forall i, B i !=set0) ->
trivIset D B -> countable D.
Proof.
move=> oB B0 tB; have [f fB] :
{f : I -> rat & forall i, D i -> B i (ratr (f i))}.
apply: (@choice _ _ (fun x y => D x -> B x (ratr y))) => i.
have [r [Bir [q _ qr]]] := dense_rat (B0 _) (oB i).
by exists q => Di; rewrite qr.
have inj_f : {in D &, injective f}.
move=> i j /[!inE] Di Dj /(congr1 ratr) ratrij.
have ? : (B i `&` B j) (ratr (f i)).
by split => //; [exact: fB|rewrite ratrij; exact: fB].
by apply/(tB _ _ Di Dj); exists (ratr (f i)).
apply/pcard_injP; have /card_bijP/cid[g bijg] := card_rat.
pose nat_of_rat (q : rat) : nat := set_val (g (to_setT q)).
have inj_nat_of_rat : injective nat_of_rat.
rewrite /nat_of_rat; apply: inj_comp => //; apply: inj_comp => //.
exact/bij_inj.
by exists (nat_of_rat \o f) => i j Di Dj /inj_nat_of_rat/inj_f; exact.
Qed.
{f : I -> rat & forall i, D i -> B i (ratr (f i))}.
apply: (@choice _ _ (fun x y => D x -> B x (ratr y))) => i.
have [r [Bir [q _ qr]]] := dense_rat (B0 _) (oB i).
by exists q => Di; rewrite qr.
have inj_f : {in D &, injective f}.
move=> i j /[!inE] Di Dj /(congr1 ratr) ratrij.
have ? : (B i `&` B j) (ratr (f i)).
by split => //; [exact: fB|rewrite ratrij; exact: fB].
by apply/(tB _ _ Di Dj); exists (ratr (f i)).
apply/pcard_injP; have /card_bijP/cid[g bijg] := card_rat.
pose nat_of_rat (q : rat) : nat := set_val (g (to_setT q)).
have inj_nat_of_rat : injective nat_of_rat.
rewrite /nat_of_rat; apply: inj_comp => //; apply: inj_comp => //.
exact/bij_inj.
by exists (nat_of_rat \o f) => i j Di Dj /inj_nat_of_rat/inj_f; exact.
Qed.