Module mathcomp.analysis.reals
# An axiomatization of real numbers $\mathbb{R}$
This file provides a classical axiomatization of real numbers as a
discrete real archimedean field with in particular a theory of floor and
ceil.
```
realType == type of real numbers
The HB class is Real.
sup A == where A : set R with R : realType, the supremum of A when
it exists, 0 otherwise
inf A := - sup (- A)
```
The mixin corresponding to realType extends an archiFieldType with two
properties:
- when sup A exists, it is an upper bound of A (lemma sup_upper_bound)
- when sup A exists, there exists an element x in A such that
sup A - eps < x for any 0 < eps (lemma sup_adherent)
```
Rint == the set of real numbers that can be written as z%:~R,
i.e., as an integer
Rtoint r == r when r is an integer, 0 otherwise
floor_set x := [set y | Rtoint y /\ y <= x]
Rfloor x == the floor of x as a real number
range1 x := [set y |x <= y < x + 1]
Rceil x == the ceil of x as a real number, i.e., - Rfloor (- x)
```
From HB Require Import structures.
From mathcomp Require Import all_ssreflect all_algebra archimedean.
From mathcomp Require Import boolp classical_sets set_interval.
From Coq Require Import Setoid.
Declare Scope real_scope.
Set Implicit Arguments.
Unset Strict Implicit.
Unset Printing Implicit Defensive.
Unset SsrOldRewriteGoalsOrder.
Import Order.TTheory Order.Syntax GRing.Theory Num.Def Num.Theory.
From mathcomp Require Import mathcomp_extra.
Delimit Scope real_scope with real.
Local Open Scope classical_set_scope.
Local Open Scope ring_scope.
Section subr_image.
Variable R : numDomainType.
Implicit Types (E : set R) (x : R).
Lemma has_ub_lbN E : has_ubound E <-> has_lbound (-%R @` E).
Proof.
Lemma has_lbound0 : has_lbound (@set0 R)
Proof.
by exists 0. Qed.
Lemma has_ubound0 : has_ubound (@set0 R)
Proof.
by exists 0. Qed.
Lemma ubound0 : ubound (@set0 R) = setT.
Proof.
Lemma lboundT : lbound [set: R] = set0.
Proof.
End subr_image.
Section has_bound_lemmas.
Variable R : realDomainType.
Implicit Types E : set R.
Implicit Types x : R.
Lemma has_ub_image_norm E : has_ubound (normr @` E) -> has_ubound E.
Proof.
Lemma has_inf_supN E : has_sup (-%R @` E) <-> has_inf E.
Proof.
Lemma has_supPn {E} : E !=set0 ->
~ has_sup E <-> (forall x, exists2 y, E y & x < y).
Proof.
End has_bound_lemmas.
HB.mixin Record ArchimedeanField_isReal R of Num.ArchiField R := {
sup_upper_bound_subdef : forall E : set [the archiFieldType of R],
has_sup E -> ubound E (supremum 0 E) ;
sup_adherent_subdef : forall (E : set [the archiFieldType of R]) (eps : R),
0 < eps -> has_sup E -> exists2 e : R, E e & (supremum 0 E - eps) < e
}.
#[short(type=realType)]
HB.structure Definition Real := {R of ArchimedeanField_isReal R
& Num.ArchiField R & Num.RealClosedField R}.
Bind Scope ring_scope with Real.sort.
Definition sup {R : realType} := @supremum _ R 0.
Definition inf {R : realType} (E : set R) := - sup (-%R @` E).
Lemma sup_upper_bound {R : realType} (E : set R):
has_sup E -> ubound E (sup E).
Proof.
Lemma sup_adherent {R : realType} (E : set R) (eps : R) : 0 < eps ->
has_sup E -> exists2 e : R, E e & (sup E - eps) < e.
Proof.
Section IsInt.
Context {R : realFieldType}.
Definition Rint_pred := fun x : R => `[< exists z, x == z%:~R >].
Arguments Rint_pred _ /.
Definition Rint := [qualify a x | Rint_pred x].
Lemma Rint_def x : (x \is a Rint) = (`[< exists z, x == z%:~R >]).
Proof.
by []. Qed.
Lemma RintP x : reflect (exists z, x = z%:~R) (x \in Rint).
Lemma RintC z : z%:~R \is a Rint.
Proof.
Lemma Rint0 : 0 \is a Rint.
Lemma Rint1 : 1 \is a Rint.
Hint Resolve Rint0 Rint1 RintC : core.
Lemma Rint_subring_closed : subring_closed Rint.
Proof.
HB.instance Definition _ := GRing.isSubringClosed.Build R Rint_pred
Rint_subring_closed.
Lemma Rint_ler_addr1 (x y : R) : x \is a Rint -> y \is a Rint ->
(x + 1 <= y) = (x < y).
Proof.
Lemma Rint_ltr_addr1 (x y : R) : x \is a Rint -> y \is a Rint ->
(x < y + 1) = (x <= y).
Proof.
End IsInt.
Arguments Rint_pred _ _ /.
Section ToInt.
Context {R : realType}.
Implicit Types x y : R.
Definition Rtoint (x : R) : int :=
if insub x : {? x | x \is a Rint} is Some Px then
xchoose (asboolP _ (tagged Px))
else 0.
Lemma RtointK (x : R): x \is a Rint -> (Rtoint x)%:~R = x.
Proof.
Lemma Rtointz (z : int): Rtoint z%:~R = z.
Lemma Rtointn (n : nat): Rtoint n%:R = n%:~R.
Lemma inj_Rtoint : {in Rint &, injective Rtoint}.
Lemma RtointN x : x \is a Rint -> Rtoint (- x) = - Rtoint x.
End ToInt.
Section RealDerivedOps.
Variable R : realType.
Implicit Types x y : R.
Definition floor_set x := [set y : R | (y \is a Rint) && (y <= x)].
Definition Rfloor x : R := (floor x)%:~R.
Definition range1 (x : R) := [set y | x <= y < x + 1].
Definition Rceil x : R := (ceil x)%:~R.
End RealDerivedOps.
Section RealLemmas.
Variables (R : realType).
Implicit Types E : set R.
Implicit Types x : R.
Lemma sup_out E : ~ has_sup E -> sup E = 0
Proof.
Lemma sup0 : sup (@set0 R) = 0
Proof.
Lemma sup1 x : sup [set x] = x
Proof.
Lemma sup_ubound {E} : has_ubound E -> ubound E (sup E).
Proof.
move=> ubE; apply/ubP=> x x_in_E; move: (x) (x_in_E).
by apply/ubP/sup_upper_bound=> //; split; first by exists x.
Qed.
by apply/ubP/sup_upper_bound=> //; split; first by exists x.
Qed.
Lemma sup_ub_strict E : has_ubound E ->
~ E (sup E) -> E `<=` [set r | r < sup E].
Proof.
Lemma sup_total {E} x : has_sup E -> down E x \/ sup E <= x.
Proof.
Lemma sup_le_ub {E} x : E !=set0 -> (ubound E) x -> sup E <= x.
Proof.
move=> hasE leEx; set y := sup E; pose z := (x + y) / 2%:R.
have Dz: 2%:R * z = x + y.
by rewrite mulrCA divff ?mulr1 // pnatr_eq0.
have ubE : has_sup E by split => //; exists x.
have [/downP [t Et lezt] | leyz] := sup_total z ubE.
rewrite -(lerD2l x) -Dz -mulr2n -[leRHS]mulr_natl.
rewrite ler_pM2l ?ltr0Sn //; apply/(le_trans lezt).
by move/ubP : leEx; exact.
rewrite -(lerD2r y) -Dz -mulr2n -[leLHS]mulr_natl.
by rewrite ler_pM2l ?ltr0Sn.
Qed.
have Dz: 2%:R * z = x + y.
by rewrite mulrCA divff ?mulr1 // pnatr_eq0.
have ubE : has_sup E by split => //; exists x.
have [/downP [t Et lezt] | leyz] := sup_total z ubE.
rewrite -(lerD2l x) -Dz -mulr2n -[leRHS]mulr_natl.
rewrite ler_pM2l ?ltr0Sn //; apply/(le_trans lezt).
by move/ubP : leEx; exact.
rewrite -(lerD2r y) -Dz -mulr2n -[leLHS]mulr_natl.
by rewrite ler_pM2l ?ltr0Sn.
Qed.
Lemma sup_setU (A B : set R) : has_sup B ->
(forall a b, A a -> B b -> a <= b) -> sup (A `|` B) = sup B.
Proof.
move=> [B0 [l Bl]] AB; apply/eqP; rewrite eq_le; apply/andP; split.
- apply sup_le_ub => [|x [Ax|]]; first by apply: subset_nonempty B0 => ?; right.
by case: B0 => b Bb; rewrite (le_trans (AB _ _ Ax Bb)) // sup_ubound //; exists l.
- by move=> Bx; rewrite sup_ubound //; exists l.
- apply sup_le_ub => // b Bb; apply: sup_ubound; last by right.
by exists l => x [Ax|Bx]; [rewrite (le_trans (AB _ _ Ax Bb)) // Bl|exact: Bl].
Qed.
- apply sup_le_ub => [|x [Ax|]]; first by apply: subset_nonempty B0 => ?; right.
by case: B0 => b Bb; rewrite (le_trans (AB _ _ Ax Bb)) // sup_ubound //; exists l.
- by move=> Bx; rewrite sup_ubound //; exists l.
- apply sup_le_ub => // b Bb; apply: sup_ubound; last by right.
by exists l => x [Ax|Bx]; [rewrite (le_trans (AB _ _ Ax Bb)) // Bl|exact: Bl].
Qed.
Lemma sup_gt (S : set R) (x : R) : S !=set0 ->
(x < sup S -> exists2 y, S y & x < y)%R.
Proof.
End RealLemmas.
#[deprecated(since="mathcomp-analysis 1.3.0", note="Renamed `sup_ubound`.")]
Notation sup_ub := sup_ubound (only parsing).
Section sup_sum.
Context {R : realType}.
Lemma sup_sumE (A B : set R) :
has_sup A -> has_sup B -> sup [set x + y | x in A & y in B] = sup A + sup B.
Proof.
move=> /[dup] supA [[a Aa] ubA] /[dup] supB [[b Bb] ubB].
have ABsup : has_sup [set x + y | x in A & y in B].
split; first by exists (a + b), a => //; exists b.
case: ubA ubB => p up [q uq]; exists (p + q) => ? [r Ar [s Bs] <-].
by apply: lerD; [exact: up | exact: uq].
apply: le_anti; apply/andP; split.
apply: sup_le_ub; first by case: ABsup.
by move=> ? [p Ap [q Bq] <-]; apply: lerD; exact: sup_ubound.
rewrite leNgt -subr_gt0; apply/negP.
set eps := (_ + _ - _) => epos.
have e2pos : 0 < eps / 2%:R by rewrite divr_gt0// ltr0n.
have [r Ar supBr] := sup_adherent e2pos supA.
have [s Bs supAs] := sup_adherent e2pos supB.
have := ltrD supBr supAs.
rewrite -addrA [-_+_]addrC -addrA -opprD -splitr addrA /= opprD opprK addrA.
rewrite subrr add0r; apply/negP; rewrite -leNgt.
by apply: sup_upper_bound => //; exists r => //; exists s.
Qed.
have ABsup : has_sup [set x + y | x in A & y in B].
split; first by exists (a + b), a => //; exists b.
case: ubA ubB => p up [q uq]; exists (p + q) => ? [r Ar [s Bs] <-].
by apply: lerD; [exact: up | exact: uq].
apply: le_anti; apply/andP; split.
apply: sup_le_ub; first by case: ABsup.
by move=> ? [p Ap [q Bq] <-]; apply: lerD; exact: sup_ubound.
rewrite leNgt -subr_gt0; apply/negP.
set eps := (_ + _ - _) => epos.
have e2pos : 0 < eps / 2%:R by rewrite divr_gt0// ltr0n.
have [r Ar supBr] := sup_adherent e2pos supA.
have [s Bs supAs] := sup_adherent e2pos supB.
have := ltrD supBr supAs.
rewrite -addrA [-_+_]addrC -addrA -opprD -splitr addrA /= opprD opprK addrA.
rewrite subrr add0r; apply/negP; rewrite -leNgt.
by apply: sup_upper_bound => //; exists r => //; exists s.
Qed.
Lemma inf_sumE (A B : set R) :
has_inf A -> has_inf B -> inf [set x + y | x in A & y in B] = inf A + inf B.
Proof.
move/has_inf_supN => ? /has_inf_supN ?; rewrite /inf.
rewrite [X in - sup X = _](_ : _ =
[set x + y | x in [set - x | x in A ] & y in [set - x | x in B]]).
rewrite eqEsubset; split => /= t [] /= x []a Aa.
case => b Bb <- <-; exists (- a); first by exists a.
by exists (- b); [exists b|rewrite opprD].
move=> <- [y] [b Bb] <- <-; exists (a + b); last by rewrite opprD.
by exists a => //; exists b.
by rewrite sup_sumE // -opprD.
Qed.
rewrite [X in - sup X = _](_ : _ =
[set x + y | x in [set - x | x in A ] & y in [set - x | x in B]]).
rewrite eqEsubset; split => /= t [] /= x []a Aa.
case => b Bb <- <-; exists (- a); first by exists a.
by exists (- b); [exists b|rewrite opprD].
move=> <- [y] [b Bb] <- <-; exists (a + b); last by rewrite opprD.
by exists a => //; exists b.
by rewrite sup_sumE // -opprD.
Qed.
End sup_sum.
Section InfTheory.
Variables (R : realType).
Implicit Types E : set R.
Implicit Types x : R.
Lemma inf_adherent E (eps : R) : 0 < eps ->
has_inf E -> exists2 e, E e & e < inf E + eps.
Proof.
move=> + /has_inf_supN supNE => /sup_adherent /(_ supNE)[e NEx egtsup].
exists (- e); first by case: NEx => x Ex <-{}; rewrite opprK.
by rewrite ltrNl -mulN1r mulrDr !mulN1r opprK.
Qed.
exists (- e); first by case: NEx => x Ex <-{}; rewrite opprK.
by rewrite ltrNl -mulN1r mulrDr !mulN1r opprK.
Qed.
Lemma inf_out E : ~ has_inf E -> inf E = 0.
Proof.
move=> ninfE; rewrite -oppr0 -(@sup_out _ (-%R @` E)) => // supNE; apply: ninfE.
exact/has_inf_supN.
Qed.
exact/has_inf_supN.
Qed.
Lemma inf0 : inf (@set0 R) = 0.
Proof.
Lemma inf1 x : inf [set x] = x.
Proof.
Lemma inf_lbound E : has_lbound E -> lbound E (inf E).
Proof.
Lemma inf_lb_strict E : has_lbound E ->
~ E (inf E) -> E `<=` [set r | inf E < r].
Proof.
Lemma lb_le_inf E x : nonempty E -> (lbound E) x -> x <= inf E.
Lemma has_infPn E : nonempty E ->
~ has_inf E <-> (forall x, exists2 y, E y & y < x).
Proof.
Lemma inf_setU (A B : set R) : has_inf A ->
(forall a b, A a -> B b -> a <= b) -> inf (A `|` B) = inf A.
Proof.
move=> hiA AB; congr (- _).
rewrite image_setU setUC sup_setU //; first exact/has_inf_supN.
by move=> _ _ [] b Bb <-{} [] a Aa <-{}; rewrite lerNl opprK; apply AB.
Qed.
rewrite image_setU setUC sup_setU //; first exact/has_inf_supN.
by move=> _ _ [] b Bb <-{} [] a Aa <-{}; rewrite lerNl opprK; apply AB.
Qed.
Lemma inf_lt (S : set R) (x : R) : S !=set0 ->
(inf S < x -> exists2 y, S y & y < x)%R.
Proof.
End InfTheory.
#[deprecated(since="mathcomp-analysis 1.3.0", note="Renamed `inf_lbound`.")]
Notation inf_lb := inf_lbound (only parsing).
Section FloorTheory.
Variable R : realType.
Implicit Types x y : R.
Lemma has_sup_floor_set x : has_sup (floor_set x).
Proof.
split; [exists (- (Num.bound (-x))%:~R) | exists (Num.bound x)%:~R].
rewrite /floor_set/mkset rpredN rpred_int /= lerNl.
case: (ger0P (-x)) => [/archi_boundP/ltW//|].
by move/ltW/le_trans; apply; rewrite ler0z.
apply/ubP=> y /andP[_] /le_trans; apply.
case: (ger0P x)=> [/archi_boundP/ltW|] //.
by move/ltW/le_trans; apply; rewrite ler0z.
Qed.
rewrite /floor_set/mkset rpredN rpred_int /= lerNl.
case: (ger0P (-x)) => [/archi_boundP/ltW//|].
by move/ltW/le_trans; apply; rewrite ler0z.
apply/ubP=> y /andP[_] /le_trans; apply.
case: (ger0P x)=> [/archi_boundP/ltW|] //.
by move/ltW/le_trans; apply; rewrite ler0z.
Qed.
Lemma sup_in_floor_set x : (floor_set x) (sup (floor_set x)).
Proof.
have /(sup_adherent ltr01) [y Fy] := has_sup_floor_set x.
have /sup_upper_bound /ubP /(_ _ Fy) := has_sup_floor_set x.
rewrite le_eqVlt=> /orP[/eqP<-//| lt_yFx].
rewrite ltrBlDr -ltrBlDl => lt1_FxBy.
pose e := sup (floor_set x) - y; have := has_sup_floor_set x.
move/sup_adherent=> -/(_ e) []; first by rewrite subr_gt0.
move=> z Fz; rewrite /e opprB addrCA subrr addr0 => lt_yz.
have /sup_upper_bound /ubP /(_ _ Fz) := has_sup_floor_set x.
rewrite -(lerD2r (-y)) => /le_lt_trans /(_ lt1_FxBy).
case/andP: Fy Fz lt_yz=> /RintP[yi -> _].
case/andP=> /RintP[zi -> _]; rewrite -rmorphB /= ltrz1 ltr_int.
rewrite lt_neqAle => /andP[ne_yz le_yz].
rewrite -[_-_]gez0_abs ?subr_ge0 // ltz_nat ltnS leqn0.
by rewrite absz_eq0 subr_eq0 eq_sym (negbTE ne_yz).
Qed.
have /sup_upper_bound /ubP /(_ _ Fy) := has_sup_floor_set x.
rewrite le_eqVlt=> /orP[/eqP<-//| lt_yFx].
rewrite ltrBlDr -ltrBlDl => lt1_FxBy.
pose e := sup (floor_set x) - y; have := has_sup_floor_set x.
move/sup_adherent=> -/(_ e) []; first by rewrite subr_gt0.
move=> z Fz; rewrite /e opprB addrCA subrr addr0 => lt_yz.
have /sup_upper_bound /ubP /(_ _ Fz) := has_sup_floor_set x.
rewrite -(lerD2r (-y)) => /le_lt_trans /(_ lt1_FxBy).
case/andP: Fy Fz lt_yz=> /RintP[yi -> _].
case/andP=> /RintP[zi -> _]; rewrite -rmorphB /= ltrz1 ltr_int.
rewrite lt_neqAle => /andP[ne_yz le_yz].
rewrite -[_-_]gez0_abs ?subr_ge0 // ltz_nat ltnS leqn0.
by rewrite absz_eq0 subr_eq0 eq_sym (negbTE ne_yz).
Qed.
Lemma isint_Rfloor x : Rfloor x \is a Rint.
Lemma RfloorE x : Rfloor x = (floor x)%:~R.
Proof.
by []. Qed.
Lemma mem_rg1_floor x : (range1 (floor x)%:~R) x.
Proof.
Lemma mem_rg1_Rfloor x : (range1 (Rfloor x)) x.
Proof.
Lemma Rfloor_le x : Rfloor x <= x.
Proof.
#[deprecated(since="mathcomp-analysis 1.3.0", note="use `ge_floor` instead")]
Lemma floor_le x : (floor x)%:~R <= x.
Proof.
Lemma lt_succ_Rfloor x : x < Rfloor x + 1.
Proof.
Lemma range1z_inj x (m1 m2 : int) :
(range1 m1%:~R) x -> (range1 m2%:~R) x -> m1 = m2.
Proof.
Lemma range1rr x : (range1 x) x.
Lemma range1zP (m : int) x : Rfloor x = m%:~R <-> (range1 m%:~R) x.
Proof.
Lemma Rfloor_natz (z : int) : Rfloor z%:~R = z%:~R :> R.
Lemma Rfloor0 : Rfloor 0 = 0 :> R
Lemma Rfloor1 : Rfloor 1 = 1 :> R
Lemma le_Rfloor : {homo (@Rfloor R) : x y / x <= y}.
Lemma Rfloor_ge_int x (n : int) : (n%:~R <= x)= (n%:~R <= Rfloor x).
Proof.
Lemma Rfloor_lt_int x (z : int) : (x < z%:~R) = (Rfloor x < z%:~R).
Proof.
Lemma Rfloor_le0 x : x <= 0 -> Rfloor x <= 0.
Lemma Rfloor_lt0 x : x < 0 -> Rfloor x < 0.
Proof.
#[deprecated(since="mathcomp-analysis 1.3.0", note="use `Num.Theory.floor_le` instead")]
Lemma le_floor : {homo @Num.floor R : x y / x <= y}.
Lemma ltr_add_invr (y x : R) : y < x -> exists k, y + k.+1%:R^-1 < x.
Proof.
End FloorTheory.
Section CeilTheory.
Variable R : realType.
Implicit Types x y : R.
Lemma isint_Rceil x : Rceil x \is a Rint.
Lemma Rceil0 : Rceil 0 = 0 :> R.
Lemma Rceil_ge x : x <= Rceil x.
Lemma le_Rceil : {homo (@Rceil R) : x y / x <= y}.
Lemma Rceil_ge0 x : 0 <= x -> 0 <= Rceil x.
Lemma RceilE x : Rceil x = (ceil x)%:~R.
Proof.
by []. Qed.
#[deprecated(since="mathcomp-analysis 1.3.0", note="use `Num.Theory.ceil_le` instead")]
Lemma le_ceil : {homo @Num.ceil R : x y / x <= y}.
Proof.
End CeilTheory.
Section Sup.
Context {R : realType}.
Implicit Types A B : set R.
Lemma le_down A : A `<=` down A.
Proof.
Lemma downK A : down (down A) = down A.
Proof.
Lemma has_sup_down A : has_sup (down A) <-> has_sup A.
Proof.
split=> -[nzA nzubA].
case: nzA => x /downP[y yS le_xy]; split; first by exists y.
case: nzubA=> u /ubP ubA; exists u; apply/ubP=> z zS.
by apply/ubA; apply/downP; exists z.
case: nzA => x xA; split; first by exists x; apply/le_down.
case: nzubA => u /ubP ubA; exists u; apply/ubP=> y /downP [].
by move=> z zA /le_trans; apply; apply/ubA.
Qed.
case: nzA => x /downP[y yS le_xy]; split; first by exists y.
case: nzubA=> u /ubP ubA; exists u; apply/ubP=> z zS.
by apply/ubA; apply/downP; exists z.
case: nzA => x xA; split; first by exists x; apply/le_down.
case: nzubA => u /ubP ubA; exists u; apply/ubP=> y /downP [].
by move=> z zA /le_trans; apply; apply/ubA.
Qed.
Lemma le_sup A B : A `<=` down B -> nonempty A -> has_sup B ->
sup A <= sup B.
Proof.
move=> le_AB nz_A hs_B; have hs_A: has_sup A.
split=> //; case: hs_B => _ [x ubx].
exists x; apply/ubP=> y /le_AB /downP[z zB le_yz].
by apply/(le_trans le_yz); move/ubP: ubx; apply.
rewrite leNgt -subr_gt0; apply/negP => lt_sup.
case: (sup_adherent lt_sup hs_A )=> x /le_AB xdB.
rewrite subKr => lt_Bx; case/downP: xdB => z zB.
move/(lt_le_trans lt_Bx); rewrite ltNge.
by move/ubP : (sup_upper_bound hs_B) => ->.
Qed.
split=> //; case: hs_B => _ [x ubx].
exists x; apply/ubP=> y /le_AB /downP[z zB le_yz].
by apply/(le_trans le_yz); move/ubP: ubx; apply.
rewrite leNgt -subr_gt0; apply/negP => lt_sup.
case: (sup_adherent lt_sup hs_A )=> x /le_AB xdB.
rewrite subKr => lt_Bx; case/downP: xdB => z zB.
move/(lt_le_trans lt_Bx); rewrite ltNge.
by move/ubP : (sup_upper_bound hs_B) => ->.
Qed.
Lemma le_inf A B : -%R @` B `<=` down (-%R @` A) -> nonempty B -> has_inf A ->
inf A <= inf B.
Proof.
Lemma sup_down A : sup (down A) = sup A.
Proof.
Lemma lt_sup_imfset {T : Type} (F : T -> R) l :
has_sup [set y | exists x, y = F x] ->
l < sup [set y | exists x, y = F x] ->
exists2 x, l < F x & F x <= sup [set y | exists x, y = F x].
Proof.
set P := [set y | _] => hs; rewrite -subr_gt0.
move=> /sup_adherent/(_ hs)[_ [x ->]]; rewrite subKr=> lt_lFx.
by exists x => //; move/ubP : (sup_upper_bound hs) => -> //; exists x.
Qed.
move=> /sup_adherent/(_ hs)[_ [x ->]]; rewrite subKr=> lt_lFx.
by exists x => //; move/ubP : (sup_upper_bound hs) => -> //; exists x.
Qed.
Lemma lt_inf_imfset {T : Type} (F : T -> R) l :
has_inf [set y | exists x, y = F x] ->
inf [set y | exists x, y = F x] < l ->
exists2 x, F x < l & inf [set y | exists x, y = F x] <= F x.
Proof.
End Sup.
Lemma int_lbound_has_minimum (B : set int) i : B !=set0 -> lbound B i ->
exists j, B j /\ forall k, B k -> j <= k.
Proof.
move=> [i0 Bi0] lbBi; have [n i0in] : exists n, i0 - i = n%:Z.
by exists `|i0 - i|%N; rewrite gez0_abs // subr_ge0; exact: lbBi.
elim: n i lbBi i0in => [i lbBi /eqP|n ih i lbBi i0in1].
by rewrite subr_eq0 => /eqP i0i; exists i0; split =>// k Bk; rewrite i0i lbBi.
have i0i1n : i0 - (i + 1) = n by rewrite opprD addrA i0in1 -addn1 PoszD addrK.
have [?|/not_forallP] := pselect (lbound B (i + 1)); first exact: (ih (i + 1)).
move=> /contrapT[x /not_implyP[Bx i1x]]; exists x; split => // k Bk.
rewrite (le_trans _ (lbBi _ Bk)) //.
by move/negP : i1x; rewrite -ltNge ltzD1.
Qed.
by exists `|i0 - i|%N; rewrite gez0_abs // subr_ge0; exact: lbBi.
elim: n i lbBi i0in => [i lbBi /eqP|n ih i lbBi i0in1].
by rewrite subr_eq0 => /eqP i0i; exists i0; split =>// k Bk; rewrite i0i lbBi.
have i0i1n : i0 - (i + 1) = n by rewrite opprD addrA i0in1 -addn1 PoszD addrK.
have [?|/not_forallP] := pselect (lbound B (i + 1)); first exact: (ih (i + 1)).
move=> /contrapT[x /not_implyP[Bx i1x]]; exists x; split => // k Bk.
rewrite (le_trans _ (lbBi _ Bk)) //.
by move/negP : i1x; rewrite -ltNge ltzD1.
Qed.
Section rat_in_itvoo.
Let bound_div (R : archiFieldType) (x y : R) : nat :=
if y < 0 then 0%N else Num.bound (y / x).
Let archi_bound_divP (R : archiFieldType) (x y : R) :
0 < x -> y < x *+ bound_div x y.
Proof.
move=> x0; have [y0|y0] := leP 0 y; last by rewrite /bound_div y0 mulr0n.
rewrite /bound_div (ltNge y 0) y0/= -mulr_natl -ltr_pdivrMr//.
by rewrite archi_boundP// (divr_ge0 _(ltW _)).
Qed.
rewrite /bound_div (ltNge y 0) y0/= -mulr_natl -ltr_pdivrMr//.
by rewrite archi_boundP// (divr_ge0 _(ltW _)).
Qed.
Lemma rat_in_itvoo (R : archiFieldType) (x y : R) :
x < y -> exists q, ratr q \in `]x, y[.
Proof.
move=> xy; move: (xy); rewrite -subr_gt0.
move=> /(archi_bound_divP 1); set n := bound_div _ _ => nyx.
have [m1 m1nx] : exists m1, m1.+1%:~R > x *+ n.
have := archi_bound_divP (x *+ n) ltr01; set p := bound_div _ _ => nxp.
have [x0|x0] := ltP 0 x.
exists p.-1; rewrite prednK // lt0n; apply: contraPN nxp => /eqP ->.
by apply/negP; rewrite -leNgt mulrn_wge0 // ltW.
by exists 0%N; rewrite (le_lt_trans _ ltr01) // mulrn_wle0.
have [m2 m2nx] : exists m2, m2.+1%:~R > - x *+ n.
have := archi_bound_divP (- x *+ n) ltr01; set p := bound_div _ _ => nxp.
have [x0|x0] := ltP 0 x.
by exists O; rewrite (le_lt_trans _ ltr01) // nmulrn_rle0// oppr_lt0.
exists p.-1; rewrite prednK // -(ltr_nat R) (le_lt_trans _ nxp) //.
by rewrite mulrn_wge0 // oppr_ge0.
have : exists m, -(m2.+1 : int) <= m <= m1.+1 /\ m%:~R - 1 <= x *+ n < m%:~R.
have m2m1 : - (m2.+1 : int) < m1.+1.
by rewrite -(ltr_int R) (lt_trans _ m1nx)// rmorphN /= ltrNl // -mulNrn.
pose B := [set m : int | m%:~R > x *+ n].
have m1B : B m1.+1 by [].
have m2B : lbound B (- m2.+1%:~R).
move=> i; rewrite /B /= -(opprK (x *+ n)) -ltrNl -mulNrn => nxi.
rewrite -(mulN1r m2.+1%:~R) mulN1r -lerNl.
by have := lt_trans nxi m2nx; rewrite intz -mulrNz ltr_int => /ltW.
have [m [Bm infB]] := int_lbound_has_minimum (ex_intro _ _ m1B) m2B.
have mN1B : ~ B (m - 1).
by move=> /infB; apply/negP; rewrite -ltNge ltrBlDr ltzD1.
exists m; split; [apply/andP; split|apply/andP; split] => //.
- by move: m2B; rewrite /lbound /= => /(_ _ Bm); rewrite intz.
- exact: infB.
- by rewrite leNgt; apply/negP; rewrite /B /= intrD in mN1B.
move=> [m [/andP[m2m mm1] /andP[mnx nxm]]].
have [/andP[a b] c] : x *+ n < m%:~R <= 1 + x *+ n /\ 1 + x *+ n < y *+ n.
split; [apply/andP; split|] => //; first by rewrite -lerBlDl.
by move: nyx; rewrite mulrnDl -ltrBrDr mulNrn.
have n_gt0 : n != 0%N by apply: contraTN nyx => /eqP ->; rewrite mulr0n ltr10.
exists (m%:Q / n%:Q); rewrite in_itv /=; apply/andP; split.
rewrite rmorphM/= (@rmorphV _ _ _ n%:~R); first by rewrite unitfE // intr_eq0.
rewrite ltr_pdivlMr /=; first by rewrite ltr0q ltr0z ltz_nat lt0n.
by rewrite mulrC // !ratr_int mulr_natl.
rewrite rmorphM /= (@rmorphV _ _ _ n%:~R); first by rewrite unitfE // intr_eq0.
rewrite ltr_pdivrMr /=; first by rewrite ltr0q ltr0z ltz_nat lt0n.
by rewrite 2!ratr_int mulr_natr (le_lt_trans _ c).
Qed.
move=> /(archi_bound_divP 1); set n := bound_div _ _ => nyx.
have [m1 m1nx] : exists m1, m1.+1%:~R > x *+ n.
have := archi_bound_divP (x *+ n) ltr01; set p := bound_div _ _ => nxp.
have [x0|x0] := ltP 0 x.
exists p.-1; rewrite prednK // lt0n; apply: contraPN nxp => /eqP ->.
by apply/negP; rewrite -leNgt mulrn_wge0 // ltW.
by exists 0%N; rewrite (le_lt_trans _ ltr01) // mulrn_wle0.
have [m2 m2nx] : exists m2, m2.+1%:~R > - x *+ n.
have := archi_bound_divP (- x *+ n) ltr01; set p := bound_div _ _ => nxp.
have [x0|x0] := ltP 0 x.
by exists O; rewrite (le_lt_trans _ ltr01) // nmulrn_rle0// oppr_lt0.
exists p.-1; rewrite prednK // -(ltr_nat R) (le_lt_trans _ nxp) //.
by rewrite mulrn_wge0 // oppr_ge0.
have : exists m, -(m2.+1 : int) <= m <= m1.+1 /\ m%:~R - 1 <= x *+ n < m%:~R.
have m2m1 : - (m2.+1 : int) < m1.+1.
by rewrite -(ltr_int R) (lt_trans _ m1nx)// rmorphN /= ltrNl // -mulNrn.
pose B := [set m : int | m%:~R > x *+ n].
have m1B : B m1.+1 by [].
have m2B : lbound B (- m2.+1%:~R).
move=> i; rewrite /B /= -(opprK (x *+ n)) -ltrNl -mulNrn => nxi.
rewrite -(mulN1r m2.+1%:~R) mulN1r -lerNl.
by have := lt_trans nxi m2nx; rewrite intz -mulrNz ltr_int => /ltW.
have [m [Bm infB]] := int_lbound_has_minimum (ex_intro _ _ m1B) m2B.
have mN1B : ~ B (m - 1).
by move=> /infB; apply/negP; rewrite -ltNge ltrBlDr ltzD1.
exists m; split; [apply/andP; split|apply/andP; split] => //.
- by move: m2B; rewrite /lbound /= => /(_ _ Bm); rewrite intz.
- exact: infB.
- by rewrite leNgt; apply/negP; rewrite /B /= intrD in mN1B.
move=> [m [/andP[m2m mm1] /andP[mnx nxm]]].
have [/andP[a b] c] : x *+ n < m%:~R <= 1 + x *+ n /\ 1 + x *+ n < y *+ n.
split; [apply/andP; split|] => //; first by rewrite -lerBlDl.
by move: nyx; rewrite mulrnDl -ltrBrDr mulNrn.
have n_gt0 : n != 0%N by apply: contraTN nyx => /eqP ->; rewrite mulr0n ltr10.
exists (m%:Q / n%:Q); rewrite in_itv /=; apply/andP; split.
rewrite rmorphM/= (@rmorphV _ _ _ n%:~R); first by rewrite unitfE // intr_eq0.
rewrite ltr_pdivlMr /=; first by rewrite ltr0q ltr0z ltz_nat lt0n.
by rewrite mulrC // !ratr_int mulr_natl.
rewrite rmorphM /= (@rmorphV _ _ _ n%:~R); first by rewrite unitfE // intr_eq0.
rewrite ltr_pdivrMr /=; first by rewrite ltr0q ltr0z ltz_nat lt0n.
by rewrite 2!ratr_int mulr_natr (le_lt_trans _ c).
Qed.
End rat_in_itvoo.