aboutsummaryrefslogtreecommitdiff
diff options
context:
space:
mode:
authorHugo Herbelin2019-09-25 19:24:35 +0200
committerHugo Herbelin2019-09-25 19:24:35 +0200
commit59079a232d2157c0c4bea4cb1a3cd68c9410e880 (patch)
treee5cbf8cdb687eb9d629621845bf5f7b978fc487c
parent227fcc4dff6fbb68ad6b91387b2f36705329a57e (diff)
parentbeeb3ed57c4749ff84524250a8b79dfa0b1e2f78 (diff)
Merge PR #10713: Define morphisms of real numbers and accelerate Cauchy reals
Reviewed-by: herbelin
-rw-r--r--doc/stdlib/index-list.html.template2
-rw-r--r--theories/Reals/ConstructiveCauchyReals.v1979
-rw-r--r--theories/Reals/ConstructiveCauchyRealsMult.v1415
-rw-r--r--theories/Reals/ConstructiveRIneq.v97
-rw-r--r--theories/Reals/ConstructiveRcomplete.v369
-rw-r--r--theories/Reals/ConstructiveReals.v746
-rw-r--r--theories/Reals/ConstructiveRealsLUB.v90
-rw-r--r--theories/Reals/ConstructiveRealsMorphisms.v1158
-rw-r--r--theories/Reals/Raxioms.v7
9 files changed, 3732 insertions, 2131 deletions
diff --git a/doc/stdlib/index-list.html.template b/doc/stdlib/index-list.html.template
index cc91776a4d..d1b98b94a3 100644
--- a/doc/stdlib/index-list.html.template
+++ b/doc/stdlib/index-list.html.template
@@ -515,7 +515,9 @@ through the <tt>Require Import</tt> command.</p>
<dd>
theories/Reals/Rdefinitions.v
theories/Reals/ConstructiveReals.v
+ theories/Reals/ConstructiveRealsMorphisms.v
theories/Reals/ConstructiveCauchyReals.v
+ theories/Reals/ConstructiveCauchyRealsMult.v
theories/Reals/Raxioms.v
theories/Reals/ConstructiveRIneq.v
theories/Reals/ConstructiveRealsLUB.v
diff --git a/theories/Reals/ConstructiveCauchyReals.v b/theories/Reals/ConstructiveCauchyReals.v
index 004854e751..965d31d403 100644
--- a/theories/Reals/ConstructiveCauchyReals.v
+++ b/theories/Reals/ConstructiveCauchyReals.v
@@ -15,34 +15,32 @@ Require Import Qround.
Require Import Logic.ConstructiveEpsilon.
Require CMorphisms.
-Open Scope Q.
+(** The constructive Cauchy real numbers, ie the Cauchy sequences
+ of rational numbers. This file is not supposed to be imported,
+ except in Rdefinitions.v, Raxioms.v, Rcomplete_constr.v
+ and ConstructiveRIneq.v.
-(* The constructive Cauchy real numbers, ie the Cauchy sequences
- of rational numbers. This file is not supposed to be imported,
- except in Rdefinitions.v, Raxioms.v, Rcomplete_constr.v
- and ConstructiveRIneq.v.
+ Constructive real numbers should be considered abstractly,
+ forgetting the fact that they are implemented as rational sequences.
+ All useful lemmas of this file are exposed in ConstructiveRIneq.v,
+ under more abstract names, like Rlt_asym instead of CRealLt_asym.
- Constructive real numbers should be considered abstractly,
- forgetting the fact that they are implemented as rational sequences.
- All useful lemmas of this file are exposed in ConstructiveRIneq.v,
- under more abstract names, like Rlt_asym instead of CRealLt_asym.
+ Cauchy reals are Cauchy sequences of rational numbers,
+ equipped with explicit moduli of convergence and
+ an equivalence relation (the difference converges to zero).
- Cauchy reals are Cauchy sequences of rational numbers,
- equipped with explicit moduli of convergence and
- an equivalence relation (the difference converges to zero).
+ Without convergence moduli, we would fail to prove that a Cauchy
+ sequence of constructive reals converges.
- Without convergence moduli, we would fail to prove that a Cauchy
- sequence of constructive reals converges.
+ Because of the Specker sequences (increasing, computable
+ and bounded sequences of rationals that do not converge
+ to a computable real number), constructive reals do not
+ follow the least upper bound principle.
- Because of the Specker sequences (increasing, computable
- and bounded sequences of rationals that do not converge
- to a computable real number), constructive reals do not
- follow the least upper bound principle.
-
- The double quantification on p q is needed to avoid
- forall un, QSeqEquiv un (fun _ => un O) (fun q => O)
- which says nothing about the limit of un.
+ The double quantification on p q is needed to avoid
+ forall un, QSeqEquiv un (fun _ => un O) (fun q => O)
+ which says nothing about the limit of un.
*)
Definition QSeqEquiv (un vn : nat -> Q) (cvmod : positive -> nat)
: Prop
@@ -213,7 +211,7 @@ Delimit Scope CReal_scope with CReal.
(* Automatically open scope R_scope for arguments of type R *)
Bind Scope CReal_scope with CReal.
-Open Scope CReal_scope.
+Local Open Scope CReal_scope.
(* So QSeqEquiv is the equivalence relation of this constructive pre-order *)
@@ -470,7 +468,8 @@ Qed.
Lemma CRealLt_above : forall (x y : CReal),
CRealLt x y
-> { k : positive | forall p:positive,
- Pos.le k p -> Qlt (2 # k) (proj1_sig y (Pos.to_nat p) - proj1_sig x (Pos.to_nat p)) }.
+ Pos.le k p -> Qlt (2 # k) (proj1_sig y (Pos.to_nat p)
+ - proj1_sig x (Pos.to_nat p)) }.
Proof.
intros x y [n maj].
pose proof (CRealLt_aboveSig x y n maj).
@@ -544,10 +543,9 @@ Proof.
Qed.
Lemma CRealLt_dec : forall x y z : CReal,
- CRealLt x y -> CRealLt x z + CRealLt z y.
+ x < y -> sum (x < z) (z < y).
Proof.
- intros [xn limx] [yn limy] [zn limz] clt.
- destruct clt as [n inf].
+ intros [xn limx] [yn limy] [zn limz] [n inf].
unfold proj1_sig in inf.
remember (yn (Pos.to_nat n) - xn (Pos.to_nat n) - (2 # n)) as eps.
assert (Qlt 0 eps) as epsPos.
@@ -629,33 +627,33 @@ Defined.
Definition linear_order_T x y z := CRealLt_dec x z y.
-Lemma CRealLe_Lt_trans : forall x y z : CReal,
+Lemma CReal_le_lt_trans : forall x y z : CReal,
x <= y -> y < z -> x < z.
Proof.
intros.
destruct (linear_order_T y x z H0). contradiction. apply c.
-Qed.
+Defined.
-Lemma CRealLt_Le_trans : forall x y z : CReal,
+Lemma CReal_lt_le_trans : forall x y z : CReal,
x < y -> y <= z -> x < z.
Proof.
intros.
destruct (linear_order_T x z y H). apply c. contradiction.
-Qed.
+Defined.
-Lemma CRealLe_trans : forall x y z : CReal,
+Lemma CReal_le_trans : forall x y z : CReal,
x <= y -> y <= z -> x <= z.
Proof.
intros. intro abs. apply H0.
- apply (CRealLt_Le_trans _ x); assumption.
+ apply (CReal_lt_le_trans _ x); assumption.
Qed.
-Lemma CRealLt_trans : forall x y z : CReal,
+Lemma CReal_lt_trans : forall x y z : CReal,
x < y -> y < z -> x < z.
Proof.
- intros. apply (CRealLt_Le_trans _ y _ H).
+ intros. apply (CReal_lt_le_trans _ y _ H).
apply CRealLt_asym. exact H0.
-Qed.
+Defined.
Lemma CRealEq_trans : forall x y z : CReal,
CRealEq x y -> CRealEq y z -> CRealEq x z.
@@ -776,6 +774,7 @@ Defined.
Notation "0" := (inject_Q 0) : CReal_scope.
Notation "1" := (inject_Q 1) : CReal_scope.
+Notation "2" := (inject_Q 2) : CReal_scope.
Lemma CRealLt_0_1 : CRealLt (inject_Q 0) (inject_Q 1).
Proof.
@@ -859,6 +858,56 @@ Proof.
apply Pos.le_max_r. apply le_refl.
Qed.
+Lemma inject_Q_compare : forall (x : CReal) (p : positive),
+ x <= inject_Q (proj1_sig x (Pos.to_nat p) + (1#p)).
+Proof.
+ intros. intros [n nmaj].
+ destruct x as [xn xcau]; simpl in nmaj.
+ apply (Qplus_lt_l _ _ (1#p)) in nmaj.
+ ring_simplify in nmaj.
+ destruct (Pos.max_dec p n).
+ - apply Pos.max_l_iff in e.
+ apply Pos2Nat.inj_le in e.
+ specialize (xcau n (Pos.to_nat n) (Pos.to_nat p) (le_refl _) e).
+ apply (Qlt_le_trans _ _ (Qabs (xn (Pos.to_nat n) + -1 * xn (Pos.to_nat p)))) in nmaj.
+ 2: apply Qle_Qabs.
+ apply (Qlt_trans _ _ _ nmaj) in xcau.
+ apply (Qplus_lt_l _ _ (-(1#n)-(1#p))) in xcau. ring_simplify in xcau.
+ setoid_replace ((2 # n) + -1 * (1 # n)) with (1#n)%Q in xcau.
+ discriminate xcau. setoid_replace (-1 * (1 # n)) with (-1#n)%Q. 2: reflexivity.
+ rewrite Qinv_plus_distr. reflexivity.
+ - apply Pos.max_r_iff, Pos2Nat.inj_le in e.
+ specialize (xcau p (Pos.to_nat n) (Pos.to_nat p) e (le_refl _)).
+ apply (Qlt_le_trans _ _ (Qabs (xn (Pos.to_nat n) + -1 * xn (Pos.to_nat p)))) in nmaj.
+ 2: apply Qle_Qabs.
+ apply (Qlt_trans _ _ _ nmaj) in xcau.
+ apply (Qplus_lt_l _ _ (-(1#p))) in xcau. ring_simplify in xcau. discriminate.
+Qed.
+
+
+Add Parametric Morphism : inject_Q
+ with signature Qeq ==> CRealEq
+ as inject_Q_morph.
+Proof.
+ split.
+ - intros [n abs]. simpl in abs. rewrite H in abs.
+ unfold Qminus in abs. rewrite Qplus_opp_r in abs. discriminate abs.
+ - intros [n abs]. simpl in abs. rewrite H in abs.
+ unfold Qminus in abs. rewrite Qplus_opp_r in abs. discriminate abs.
+Qed.
+
+Instance inject_Q_morph_T
+ : CMorphisms.Proper
+ (CMorphisms.respectful Qeq CRealEq) inject_Q.
+Proof.
+ split.
+ - intros [n abs]. simpl in abs. rewrite H in abs.
+ unfold Qminus in abs. rewrite Qplus_opp_r in abs. discriminate abs.
+ - intros [n abs]. simpl in abs. rewrite H in abs.
+ unfold Qminus in abs. rewrite Qplus_opp_r in abs. discriminate abs.
+Qed.
+
+
(* Algebraic operations *)
@@ -1029,9 +1078,7 @@ Proof.
Qed.
Lemma CReal_plus_lt_compat_l :
- forall x y z : CReal,
- CRealLt y z
- -> CRealLt (CReal_plus x y) (CReal_plus x z).
+ forall x y z : CReal, y < z -> x + y < x + z.
Proof.
intros.
apply CRealLt_above in H. destruct H as [n maj].
@@ -1047,6 +1094,13 @@ Proof.
simpl; ring.
Qed.
+Lemma CReal_plus_lt_compat_r :
+ forall x y z : CReal, y < z -> y + x < z + x.
+Proof.
+ intros. do 2 rewrite <- (CReal_plus_comm x).
+ apply CReal_plus_lt_compat_l. assumption.
+Qed.
+
Lemma CReal_plus_lt_reg_l :
forall x y z : CReal, x + y < x + z -> y < z.
Proof.
@@ -1068,6 +1122,20 @@ Proof.
apply CReal_plus_lt_reg_l in H. exact H.
Qed.
+Lemma CReal_plus_le_reg_l :
+ forall x y z : CReal, x + y <= x + z -> y <= z.
+Proof.
+ intros. intro abs. apply H. clear H.
+ apply CReal_plus_lt_compat_l. exact abs.
+Qed.
+
+Lemma CReal_plus_le_reg_r :
+ forall x y z : CReal, y + x <= z + x -> y <= z.
+Proof.
+ intros. intro abs. apply H. clear H.
+ apply CReal_plus_lt_compat_r. exact abs.
+Qed.
+
Lemma CReal_plus_le_compat_l : forall r r1 r2, r1 <= r2 -> r + r1 <= r + r2.
Proof.
intros. intro abs. apply CReal_plus_lt_reg_l in abs. contradiction.
@@ -1076,12 +1144,21 @@ Qed.
Lemma CReal_plus_le_lt_compat :
forall r1 r2 r3 r4 : CReal, r1 <= r2 -> r3 < r4 -> r1 + r3 < r2 + r4.
Proof.
- intros; apply CRealLe_Lt_trans with (r2 + r3).
+ intros; apply CReal_le_lt_trans with (r2 + r3).
intro abs. rewrite CReal_plus_comm, (CReal_plus_comm r1) in abs.
apply CReal_plus_lt_reg_l in abs. contradiction.
apply CReal_plus_lt_compat_l; exact H0.
Qed.
+Lemma CReal_plus_le_compat :
+ forall r1 r2 r3 r4 : CReal, r1 <= r2 -> r3 <= r4 -> r1 + r3 <= r2 + r4.
+Proof.
+ intros; apply CReal_le_trans with (r2 + r3).
+ intro abs. rewrite CReal_plus_comm, (CReal_plus_comm r1) in abs.
+ apply CReal_plus_lt_reg_l in abs. contradiction.
+ apply CReal_plus_le_compat_l; exact H0.
+Qed.
+
Lemma CReal_plus_opp_r : forall x : CReal,
x + - x == 0.
Proof.
@@ -1140,1812 +1217,110 @@ Proof.
Qed.
Lemma CReal_plus_eq_reg_l : forall (r r1 r2 : CReal),
- CRealEq (CReal_plus r r1) (CReal_plus r r2)
- -> CRealEq r1 r2.
+ r + r1 == r + r2 -> r1 == r2.
Proof.
intros. destruct H. split.
- intro abs. apply (CReal_plus_lt_compat_l r) in abs. contradiction.
- intro abs. apply (CReal_plus_lt_compat_l r) in abs. contradiction.
Qed.
-Fixpoint BoundFromZero (qn : nat -> Q) (k : nat) (A : positive) { struct k }
- : (forall n:nat, le k n -> Qlt (Qabs (qn n)) (Z.pos A # 1))
- -> { B : positive | forall n:nat, Qlt (Qabs (qn n)) (Z.pos B # 1) }.
-Proof.
- intro H. destruct k.
- - exists A. intros. apply H. apply le_0_n.
- - destruct (Qarchimedean (Qabs (qn k))) as [a maj].
- apply (BoundFromZero qn k (Pos.max A a)).
- intros n H0. destruct (Nat.le_gt_cases n k).
- + pose proof (Nat.le_antisymm n k H1 H0). subst k.
- apply (Qlt_le_trans _ (Z.pos a # 1)). apply maj.
- unfold Qle; simpl. rewrite Pos.mul_1_r. rewrite Pos.mul_1_r.
- apply Pos.le_max_r.
- + apply (Qlt_le_trans _ (Z.pos A # 1)). apply H.
- apply H1. unfold Qle; simpl. rewrite Pos.mul_1_r. rewrite Pos.mul_1_r.
- apply Pos.le_max_l.
-Qed.
-
-Lemma QCauchySeq_bounded (qn : nat -> Q) (cvmod : positive -> nat)
- : QCauchySeq qn cvmod
- -> { A : positive | forall n:nat, Qlt (Qabs (qn n)) (Z.pos A # 1) }.
-Proof.
- intros. remember (Zplus (Qnum (Qabs (qn (cvmod xH)))) 1) as z.
- assert (Z.lt 0 z) as zPos.
- { subst z. assert (Qle 0 (Qabs (qn (cvmod 1%positive)))).
- apply Qabs_nonneg. destruct (Qabs (qn (cvmod 1%positive))). simpl.
- unfold Qle in H0. simpl in H0. rewrite Zmult_1_r in H0.
- apply (Z.lt_le_trans 0 1). unfold Z.lt. auto.
- rewrite <- (Zplus_0_l 1). rewrite Zplus_assoc. apply Zplus_le_compat_r.
- rewrite Zplus_0_r. assumption. }
- assert { A : positive | forall n:nat,
- le (cvmod xH) n -> Qlt ((Qabs (qn n)) * (1#A)) 1 }.
- destruct z eqn:des.
- - exfalso. apply (Z.lt_irrefl 0). assumption.
- - exists p. intros. specialize (H xH (cvmod xH) n (le_refl _) H0).
- assert (Qlt (Qabs (qn n)) (Qabs (qn (cvmod 1%positive)) + 1)).
- { apply (Qplus_lt_l _ _ (-Qabs (qn (cvmod 1%positive)))).
- rewrite <- (Qplus_comm 1). rewrite <- Qplus_assoc. rewrite Qplus_opp_r.
- rewrite Qplus_0_r. apply (Qle_lt_trans _ (Qabs (qn n - qn (cvmod 1%positive)))).
- apply Qabs_triangle_reverse. rewrite Qabs_Qminus. assumption. }
- apply (Qlt_le_trans _ ((Qabs (qn (cvmod 1%positive)) + 1) * (1#p))).
- apply Qmult_lt_r. unfold Qlt. simpl. unfold Z.lt. auto. assumption.
- unfold Qle. simpl. rewrite Zmult_1_r. rewrite Zmult_1_r. rewrite Zmult_1_r.
- rewrite Pos.mul_1_r. rewrite Pos2Z.inj_mul. rewrite Heqz.
- destruct (Qabs (qn (cvmod 1%positive))) eqn:desAbs.
- rewrite Z.mul_add_distr_l. rewrite Zmult_1_r.
- apply Zplus_le_compat_r. rewrite <- (Zmult_1_l (QArith_base.Qnum (Qnum # Qden))).
- rewrite Zmult_assoc. apply Zmult_le_compat_r. rewrite Zmult_1_r.
- simpl. unfold Z.le. rewrite <- Pos2Z.inj_compare.
- unfold Pos.compare. destruct Qden; discriminate.
- simpl. assert (Qle 0 (Qnum # Qden)). rewrite <- desAbs.
- apply Qabs_nonneg. unfold Qle in H2. simpl in H2. rewrite Zmult_1_r in H2.
- assumption.
- - exfalso. inversion zPos.
- - destruct H0. apply (BoundFromZero _ (cvmod xH) x). intros n H0.
- specialize (q n H0). setoid_replace (Z.pos x # 1)%Q with (/(1#x))%Q.
- rewrite <- (Qmult_1_l (/(1#x))). apply Qlt_shift_div_l.
- reflexivity. apply q. reflexivity.
-Qed.
-
-Lemma CReal_mult_cauchy
- : forall (xn yn zn : nat -> Q) (Ay Az : positive) (cvmod : positive -> nat),
- QSeqEquiv xn yn cvmod
- -> QCauchySeq zn Pos.to_nat
- -> (forall n:nat, Qlt (Qabs (yn n)) (Z.pos Ay # 1))
- -> (forall n:nat, Qlt (Qabs (zn n)) (Z.pos Az # 1))
- -> QSeqEquiv (fun n:nat => xn n * zn n) (fun n:nat => yn n * zn n)
- (fun p => max (cvmod (2 * (Pos.max Ay Az) * p)%positive)
- (Pos.to_nat (2 * (Pos.max Ay Az) * p)%positive)).
-Proof.
- intros xn yn zn Ay Az cvmod limx limz majy majz.
- remember (Pos.mul 2 (Pos.max Ay Az)) as z.
- intros k p q H H0.
- assert (Pos.to_nat k <> O) as kPos.
- { intro absurd. pose proof (Pos2Nat.is_pos k).
- rewrite absurd in H1. inversion H1. }
- setoid_replace (xn p * zn p - yn q * zn q)%Q
- with ((xn p - yn q) * zn p + yn q * (zn p - zn q))%Q.
- 2: ring.
- apply (Qle_lt_trans _ (Qabs ((xn p - yn q) * zn p)
- + Qabs (yn q * (zn p - zn q)))).
- apply Qabs_triangle. rewrite Qabs_Qmult. rewrite Qabs_Qmult.
- setoid_replace (1#k)%Q with ((1#2*k) + (1#2*k))%Q.
- apply Qplus_lt_le_compat.
- - apply (Qle_lt_trans _ ((1#z * k) * Qabs (zn p)%nat)).
- + apply Qmult_le_compat_r. apply Qle_lteq. left. apply limx.
- apply (le_trans _ (Init.Nat.max (cvmod (z * k)%positive) (Pos.to_nat (z * k)))).
- apply Nat.le_max_l. assumption.
- apply (le_trans _ (Init.Nat.max (cvmod (z * k)%positive) (Pos.to_nat (z * k)))).
- apply Nat.le_max_l. assumption. apply Qabs_nonneg.
- + subst z. rewrite <- (Qmult_1_r (1 # 2 * k)).
- rewrite <- Pos.mul_assoc. rewrite <- (Pos.mul_comm k). rewrite Pos.mul_assoc.
- rewrite (factorDenom _ _ (2 * k)). rewrite <- Qmult_assoc.
- apply Qmult_lt_l. unfold Qlt. simpl. unfold Z.lt. auto.
- apply (Qle_lt_trans _ (Qabs (zn p)%nat * (1 # Az))).
- rewrite <- (Qmult_comm (1 # Az)). apply Qmult_le_compat_r.
- unfold Qle. simpl. rewrite Pos2Z.inj_max. apply Z.le_max_r.
- apply Qabs_nonneg. rewrite <- (Qmult_inv_r (1#Az)).
- rewrite Qmult_comm. apply Qmult_lt_l. reflexivity.
- setoid_replace (/(1#Az))%Q with (Z.pos Az # 1)%Q. apply majz.
- reflexivity. intro abs. inversion abs.
- - apply (Qle_trans _ ((1 # z * k) * Qabs (yn q)%nat)).
- + rewrite Qmult_comm. apply Qmult_le_compat_r. apply Qle_lteq.
- left. apply limz.
- apply (le_trans _ (max (cvmod (z * k)%positive)
- (Pos.to_nat (z * k)%positive))).
- apply Nat.le_max_r. assumption.
- apply (le_trans _ (max (cvmod (z * k)%positive)
- (Pos.to_nat (z * k)%positive))).
- apply Nat.le_max_r. assumption. apply Qabs_nonneg.
- + subst z. rewrite <- (Qmult_1_r (1 # 2 * k)).
- rewrite <- Pos.mul_assoc. rewrite <- (Pos.mul_comm k). rewrite Pos.mul_assoc.
- rewrite (factorDenom _ _ (2 * k)). rewrite <- Qmult_assoc.
- apply Qle_lteq. left.
- apply Qmult_lt_l. unfold Qlt. simpl. unfold Z.lt. auto.
- apply (Qle_lt_trans _ (Qabs (yn q)%nat * (1 # Ay))).
- rewrite <- (Qmult_comm (1 # Ay)). apply Qmult_le_compat_r.
- unfold Qle. simpl. rewrite Pos2Z.inj_max. apply Z.le_max_l.
- apply Qabs_nonneg. rewrite <- (Qmult_inv_r (1#Ay)).
- rewrite Qmult_comm. apply Qmult_lt_l. reflexivity.
- setoid_replace (/(1#Ay))%Q with (Z.pos Ay # 1)%Q. apply majy.
- reflexivity. intro abs. inversion abs.
- - rewrite Qinv_plus_distr. unfold Qeq. reflexivity.
-Qed.
-
-Lemma linear_max : forall (p Ax Ay : positive) (i : nat),
- le (Pos.to_nat p) i
- -> (Init.Nat.max (Pos.to_nat (2 * Pos.max Ax Ay * p))
- (Pos.to_nat (2 * Pos.max Ax Ay * p)) <= Pos.to_nat (2 * Pos.max Ax Ay) * i)%nat.
-Proof.
- intros. rewrite max_l. 2: apply le_refl.
- rewrite Pos2Nat.inj_mul. apply Nat.mul_le_mono_nonneg.
- apply le_0_n. apply le_refl. apply le_0_n. apply H.
-Qed.
-
-Definition CReal_mult (x y : CReal) : CReal.
-Proof.
- destruct x as [xn limx]. destruct y as [yn limy].
- destruct (QCauchySeq_bounded xn Pos.to_nat limx) as [Ax majx].
- destruct (QCauchySeq_bounded yn Pos.to_nat limy) as [Ay majy].
- pose proof (CReal_mult_cauchy xn xn yn Ax Ay Pos.to_nat limx limy majx majy).
- exists (fun n : nat => xn (Pos.to_nat (2 * Pos.max Ax Ay)* n)%nat
- * yn (Pos.to_nat (2 * Pos.max Ax Ay) * n)%nat).
- intros p n k H0 H1.
- apply H; apply linear_max; assumption.
-Defined.
-
-Infix "*" := CReal_mult : CReal_scope.
-
-Lemma CReal_mult_unfold : forall x y : CReal,
- QSeqEquivEx (proj1_sig (CReal_mult x y))
- (fun n : nat => proj1_sig x n * proj1_sig y n)%Q.
-Proof.
- intros [xn limx] [yn limy]. unfold CReal_mult ; simpl.
- destruct (QCauchySeq_bounded xn Pos.to_nat limx) as [Ax majx].
- destruct (QCauchySeq_bounded yn Pos.to_nat limy) as [Ay majy].
- simpl.
- pose proof (CReal_mult_cauchy xn xn yn Ax Ay Pos.to_nat limx limy majx majy).
- exists (fun p : positive =>
- Init.Nat.max (Pos.to_nat (2 * Pos.max Ax Ay * p))
- (Pos.to_nat (2 * Pos.max Ax Ay * p))).
- intros p n k H0 H1. rewrite max_l in H0, H1.
- 2: apply le_refl. 2: apply le_refl.
- apply H. apply linear_max.
- apply (le_trans _ (Pos.to_nat (2 * Pos.max Ax Ay * p))).
- rewrite <- (mult_1_l (Pos.to_nat p)). rewrite Pos2Nat.inj_mul.
- apply Nat.mul_le_mono_nonneg. auto. apply Pos2Nat.is_pos.
- apply le_0_n. apply le_refl. apply H0. rewrite max_l.
- apply H1. apply le_refl.
-Qed.
-
-Lemma CReal_mult_assoc_bounded_r : forall (xn yn zn : nat -> Q),
- QSeqEquivEx xn yn (* both are Cauchy with same limit *)
- -> QSeqEquiv zn zn Pos.to_nat
- -> QSeqEquivEx (fun n => xn n * zn n)%Q (fun n => yn n * zn n)%Q.
-Proof.
- intros. destruct H as [cvmod cveq].
- destruct (QCauchySeq_bounded yn (fun k => cvmod (2 * k)%positive)
- (QSeqEquiv_cau_r xn yn cvmod cveq))
- as [Ay majy].
- destruct (QCauchySeq_bounded zn Pos.to_nat H0) as [Az majz].
- exists (fun p => max (cvmod (2 * (Pos.max Ay Az) * p)%positive)
- (Pos.to_nat (2 * (Pos.max Ay Az) * p)%positive)).
- apply CReal_mult_cauchy; assumption.
-Qed.
-
-Lemma CReal_mult_assoc : forall x y z : CReal,
- CRealEq (CReal_mult (CReal_mult x y) z)
- (CReal_mult x (CReal_mult y z)).
-Proof.
- intros. apply CRealEq_diff. apply CRealEq_modindep.
- apply (QSeqEquivEx_trans _ (fun n => proj1_sig x n * proj1_sig y n * proj1_sig z n)%Q).
- - apply (QSeqEquivEx_trans _ (fun n => proj1_sig (CReal_mult x y) n * proj1_sig z n)%Q).
- apply CReal_mult_unfold.
- destruct x as [xn limx], y as [yn limy], z as [zn limz]; unfold CReal_mult; simpl.
- destruct (QCauchySeq_bounded xn Pos.to_nat limx) as [Ax majx].
- destruct (QCauchySeq_bounded yn Pos.to_nat limy) as [Ay majy].
- destruct (QCauchySeq_bounded zn Pos.to_nat limz) as [Az majz].
- apply CReal_mult_assoc_bounded_r. 2: apply limz.
- simpl.
- pose proof (CReal_mult_cauchy xn xn yn Ax Ay Pos.to_nat limx limy majx majy).
- exists (fun p : positive =>
- Init.Nat.max (Pos.to_nat (2 * Pos.max Ax Ay * p))
- (Pos.to_nat (2 * Pos.max Ax Ay * p))).
- intros p n k H0 H1. rewrite max_l in H0, H1.
- 2: apply le_refl. 2: apply le_refl.
- apply H. apply linear_max.
- apply (le_trans _ (Pos.to_nat (2 * Pos.max Ax Ay * p))).
- rewrite <- (mult_1_l (Pos.to_nat p)). rewrite Pos2Nat.inj_mul.
- apply Nat.mul_le_mono_nonneg. auto. apply Pos2Nat.is_pos.
- apply le_0_n. apply le_refl. apply H0. rewrite max_l.
- apply H1. apply le_refl.
- - apply (QSeqEquivEx_trans _ (fun n => proj1_sig x n * proj1_sig (CReal_mult y z) n)%Q).
- 2: apply QSeqEquivEx_sym; apply CReal_mult_unfold.
- destruct x as [xn limx], y as [yn limy], z as [zn limz]; unfold CReal_mult; simpl.
- destruct (QCauchySeq_bounded xn Pos.to_nat limx) as [Ax majx].
- destruct (QCauchySeq_bounded yn Pos.to_nat limy) as [Ay majy].
- destruct (QCauchySeq_bounded zn Pos.to_nat limz) as [Az majz].
- simpl.
- pose proof (CReal_mult_assoc_bounded_r (fun n0 : nat => yn n0 * zn n0)%Q (fun n : nat =>
- yn (Pos.to_nat (Pos.max Ay Az)~0 * n)%nat
- * zn (Pos.to_nat (Pos.max Ay Az)~0 * n)%nat)%Q xn)
- as [cvmod cveq].
-
- pose proof (CReal_mult_cauchy yn yn zn Ay Az Pos.to_nat limy limz majy majz).
- exists (fun p : positive =>
- Init.Nat.max (Pos.to_nat (2 * Pos.max Ay Az * p))
- (Pos.to_nat (2 * Pos.max Ay Az * p))).
- intros p n k H0 H1. rewrite max_l in H0, H1.
- 2: apply le_refl. 2: apply le_refl.
- apply H. rewrite max_l. apply H0. apply le_refl.
- apply linear_max.
- apply (le_trans _ (Pos.to_nat (2 * Pos.max Ay Az * p))).
- rewrite <- (mult_1_l (Pos.to_nat p)). rewrite Pos2Nat.inj_mul.
- apply Nat.mul_le_mono_nonneg. auto. apply Pos2Nat.is_pos.
- apply le_0_n. apply le_refl. apply H1.
- apply limx.
- exists cvmod. intros p k n H1 H2. specialize (cveq p k n H1 H2).
- setoid_replace (xn k * yn k * zn k -
- xn n *
- (yn (Pos.to_nat (Pos.max Ay Az)~0 * n)%nat *
- zn (Pos.to_nat (Pos.max Ay Az)~0 * n)%nat))%Q
- with ((fun n : nat => yn n * zn n * xn n) k -
- (fun n : nat =>
- yn (Pos.to_nat (Pos.max Ay Az)~0 * n)%nat *
- zn (Pos.to_nat (Pos.max Ay Az)~0 * n)%nat *
- xn n) n)%Q.
- apply cveq. ring.
-Qed.
-
-Lemma CReal_mult_comm : forall x y : CReal,
- CRealEq (CReal_mult x y) (CReal_mult y x).
-Proof.
- intros. apply CRealEq_diff. apply CRealEq_modindep.
- apply (QSeqEquivEx_trans _ (fun n => proj1_sig y n * proj1_sig x n)%Q).
- destruct x as [xn limx], y as [yn limy]; simpl.
- 2: apply QSeqEquivEx_sym; apply CReal_mult_unfold.
- destruct (QCauchySeq_bounded xn Pos.to_nat limx) as [Ax majx].
- destruct (QCauchySeq_bounded yn Pos.to_nat limy) as [Ay majy]; simpl.
- apply QSeqEquivEx_sym.
-
- pose proof (CReal_mult_cauchy yn yn xn Ay Ax Pos.to_nat limy limx majy majx).
- exists (fun p : positive =>
- Init.Nat.max (Pos.to_nat (2 * Pos.max Ay Ax * p))
- (Pos.to_nat (2 * Pos.max Ay Ax * p))).
- intros p n k H0 H1. rewrite max_l in H0, H1.
- 2: apply le_refl. 2: apply le_refl.
- rewrite (Qmult_comm (xn (Pos.to_nat (Pos.max Ax Ay)~0 * k)%nat)).
- apply (H p n). rewrite max_l. apply H0. apply le_refl.
- rewrite max_l. apply (le_trans _ k). apply H1.
- rewrite <- (mult_1_l k). rewrite mult_assoc.
- apply Nat.mul_le_mono_nonneg. auto. rewrite mult_1_r.
- apply Pos2Nat.is_pos. apply le_0_n. apply le_refl.
- apply le_refl.
-Qed.
-
-(* Axiom Rmult_eq_compat_l *)
-Lemma CReal_mult_proper_l : forall x y z : CReal,
- CRealEq y z -> CRealEq (CReal_mult x y) (CReal_mult x z).
-Proof.
- intros. apply CRealEq_diff. apply CRealEq_modindep.
- apply (QSeqEquivEx_trans _ (fun n => proj1_sig x n * proj1_sig y n)%Q).
- apply CReal_mult_unfold.
- rewrite CRealEq_diff in H. rewrite <- CRealEq_modindep in H.
- apply QSeqEquivEx_sym.
- apply (QSeqEquivEx_trans _ (fun n => proj1_sig x n * proj1_sig z n)%Q).
- apply CReal_mult_unfold.
- destruct x as [xn limx], y as [yn limy], z as [zn limz]; simpl.
- destruct H. simpl in H.
- destruct (QCauchySeq_bounded xn Pos.to_nat limx) as [Ax majx].
- destruct (QCauchySeq_bounded zn Pos.to_nat limz) as [Az majz].
- pose proof (CReal_mult_cauchy yn zn xn Az Ax x H limx majz majx).
- apply QSeqEquivEx_sym.
- exists (fun p : positive =>
- Init.Nat.max (x (2 * Pos.max Az Ax * p)%positive)
- (Pos.to_nat (2 * Pos.max Az Ax * p))).
- intros p n k H1 H2. specialize (H0 p n k H1 H2).
- setoid_replace (xn n * yn n - xn k * zn k)%Q
- with (yn n * xn n - zn k * xn k)%Q.
- apply H0. ring.
-Qed.
-
-Lemma CReal_mult_lt_0_compat : forall x y : CReal,
- CRealLt (inject_Q 0) x
- -> CRealLt (inject_Q 0) y
- -> CRealLt (inject_Q 0) (CReal_mult x y).
-Proof.
- intros. destruct H as [x0 H], H0 as [x1 H0].
- pose proof (CRealLt_aboveSig (inject_Q 0) x x0 H).
- pose proof (CRealLt_aboveSig (inject_Q 0) y x1 H0).
- destruct x as [xn limx], y as [yn limy].
- simpl in H, H1, H2. simpl.
- destruct (QCauchySeq_bounded xn Pos.to_nat limx) as [Ax majx].
- destruct (QCauchySeq_bounded yn Pos.to_nat limy) as [Ay majy].
- destruct (Qarchimedean (/ (xn (Pos.to_nat x0) - 0 - (2 # x0)))).
- destruct (Qarchimedean (/ (yn (Pos.to_nat x1) - 0 - (2 # x1)))).
- exists (Pos.max x0 x~0 * Pos.max x1 x2~0)%positive.
- simpl. unfold Qminus. rewrite Qplus_0_r.
- rewrite <- Pos2Nat.inj_mul.
- unfold Qminus in H1, H2.
- specialize (H1 ((Pos.max Ax Ay)~0 * (Pos.max x0 x~0 * Pos.max x1 x2~0))%positive).
- assert (Pos.max x1 x2~0 <= (Pos.max Ax Ay)~0 * (Pos.max x0 x~0 * Pos.max x1 x2~0))%positive.
- { apply Pos2Nat.inj_le.
- rewrite Pos.mul_assoc. rewrite Pos2Nat.inj_mul.
- rewrite <- (mult_1_l (Pos.to_nat (Pos.max x1 x2~0))).
- rewrite mult_assoc. apply Nat.mul_le_mono_nonneg. auto.
- rewrite mult_1_r. apply Pos2Nat.is_pos. apply le_0_n.
- apply le_refl. }
- specialize (H2 ((Pos.max Ax Ay)~0 * (Pos.max x0 x~0 * Pos.max x1 x2~0))%positive H3).
- rewrite Qplus_0_r in H1, H2.
- apply (Qlt_trans _ ((2 # Pos.max x0 x~0) * (2 # Pos.max x1 x2~0))).
- unfold Qlt; simpl. assert (forall p : positive, (Z.pos p < Z.pos p~0)%Z).
- intro p. rewrite <- (Z.mul_1_l (Z.pos p)).
- replace (Z.pos p~0) with (2 * Z.pos p)%Z. apply Z.mul_lt_mono_pos_r.
- apply Pos2Z.is_pos. reflexivity. reflexivity.
- apply H4.
- apply (Qlt_trans _ ((2 # Pos.max x0 x~0) * (yn (Pos.to_nat ((Pos.max Ax Ay)~0 * (Pos.max x0 x~0 * Pos.max x1 x2~0)))))).
- apply Qmult_lt_l. reflexivity. apply H2. apply Qmult_lt_r.
- apply (Qlt_trans 0 (2 # Pos.max x1 x2~0)). reflexivity. apply H2.
- apply H1. rewrite Pos.mul_comm. apply Pos2Nat.inj_le.
- rewrite <- Pos.mul_assoc. rewrite Pos2Nat.inj_mul.
- rewrite <- (mult_1_r (Pos.to_nat (Pos.max x0 x~0))).
- rewrite <- mult_assoc. apply Nat.mul_le_mono_nonneg.
- apply le_0_n. apply le_refl. auto.
- rewrite mult_1_l. apply Pos2Nat.is_pos.
-Qed.
-
-Lemma CReal_mult_plus_distr_l : forall r1 r2 r3 : CReal,
- r1 * (r2 + r3) == (r1 * r2) + (r1 * r3).
-Proof.
- intros x y z. apply CRealEq_diff. apply CRealEq_modindep.
- apply (QSeqEquivEx_trans _ (fun n => proj1_sig x n
- * (proj1_sig (CReal_plus y z) n))%Q).
- apply CReal_mult_unfold.
- apply (QSeqEquivEx_trans _ (fun n => proj1_sig (CReal_mult x y) n
- + proj1_sig (CReal_mult x z) n))%Q.
- 2: apply QSeqEquivEx_sym; exists (fun p => Pos.to_nat (2 * p))
- ; apply CReal_plus_unfold.
- apply (QSeqEquivEx_trans _ (fun n => proj1_sig x n
- * (proj1_sig y n + proj1_sig z n))%Q).
- - pose proof (CReal_plus_unfold y z).
- destruct x as [xn limx], y as [yn limy], z as [zn limz]; simpl; simpl in H.
- destruct (QCauchySeq_bounded xn Pos.to_nat limx) as [Ax majx].
- destruct (QCauchySeq_bounded yn Pos.to_nat limy) as [Ay majy].
- destruct (QCauchySeq_bounded zn Pos.to_nat limz) as [Az majz].
- pose proof (CReal_mult_cauchy (fun n => yn (n + (n + 0))%nat + zn (n + (n + 0))%nat)%Q
- (fun n => yn n + zn n)%Q
- xn (Ay + Az) Ax
- (fun p => Pos.to_nat (2 * p)) H limx).
- exists (fun p : positive => (Pos.to_nat (2 * (2 * Pos.max (Ay + Az) Ax * p)))).
- intros p n k H1 H2.
- setoid_replace (xn n * (yn (n + (n + 0))%nat + zn (n + (n + 0))%nat) - xn k * (yn k + zn k))%Q
- with ((yn (n + (n + 0))%nat + zn (n + (n + 0))%nat) * xn n - (yn k + zn k) * xn k)%Q.
- 2: ring.
- assert (Pos.to_nat (2 * Pos.max (Ay + Az) Ax * p) <=
- Pos.to_nat 2 * Pos.to_nat (2 * Pos.max (Ay + Az) Ax * p))%nat.
- { rewrite (Pos2Nat.inj_mul 2).
- rewrite <- (mult_1_l (Pos.to_nat (2 * Pos.max (Ay + Az) Ax * p))).
- rewrite mult_assoc. apply Nat.mul_le_mono_nonneg. auto.
- simpl. auto. apply le_0_n. apply le_refl. }
- apply H0. intro n0. apply (Qle_lt_trans _ (Qabs (yn n0) + Qabs (zn n0))).
- apply Qabs_triangle. rewrite Pos2Z.inj_add.
- rewrite <- Qinv_plus_distr. apply Qplus_lt_le_compat.
- apply majy. apply Qlt_le_weak. apply majz.
- apply majx. rewrite max_l.
- apply H1. rewrite (Pos2Nat.inj_mul 2). apply H3.
- rewrite max_l. apply H2. rewrite (Pos2Nat.inj_mul 2).
- apply H3.
- - destruct x as [xn limx], y as [yn limy], z as [zn limz]; simpl.
- destruct (QCauchySeq_bounded xn Pos.to_nat limx) as [Ax majx].
- destruct (QCauchySeq_bounded yn Pos.to_nat limy) as [Ay majy].
- destruct (QCauchySeq_bounded zn Pos.to_nat limz) as [Az majz].
- simpl.
- exists (fun p : positive => (Pos.to_nat (2 * (Pos.max (Pos.max Ax Ay) Az) * (2 * p)))).
- intros p n k H H0.
- setoid_replace (xn n * (yn n + zn n) -
- (xn (Pos.to_nat (Pos.max Ax Ay)~0 * k)%nat *
- yn (Pos.to_nat (Pos.max Ax Ay)~0 * k)%nat +
- xn (Pos.to_nat (Pos.max Ax Az)~0 * k)%nat *
- zn (Pos.to_nat (Pos.max Ax Az)~0 * k)%nat))%Q
- with (xn n * yn n - (xn (Pos.to_nat (Pos.max Ax Ay)~0 * k)%nat *
- yn (Pos.to_nat (Pos.max Ax Ay)~0 * k)%nat)
- + (xn n * zn n - xn (Pos.to_nat (Pos.max Ax Az)~0 * k)%nat *
- zn (Pos.to_nat (Pos.max Ax Az)~0 * k)%nat))%Q.
- 2: ring.
- apply (Qle_lt_trans _ (Qabs (xn n * yn n - (xn (Pos.to_nat (Pos.max Ax Ay)~0 * k)%nat *
- yn (Pos.to_nat (Pos.max Ax Ay)~0 * k)%nat))
- + Qabs (xn n * zn n - xn (Pos.to_nat (Pos.max Ax Az)~0 * k)%nat *
- zn (Pos.to_nat (Pos.max Ax Az)~0 * k)%nat))).
- apply Qabs_triangle.
- setoid_replace (1#p)%Q with ((1#2*p) + (1#2*p))%Q.
- apply Qplus_lt_le_compat.
- + pose proof (CReal_mult_cauchy xn xn yn Ax Ay Pos.to_nat limx limy).
- apply H1. apply majx. apply majy. rewrite max_l.
- apply (le_trans _ (Pos.to_nat (2 * Pos.max (Pos.max Ax Ay) Az * (2 * p)))).
- rewrite (Pos.mul_comm 2). rewrite <- Pos.mul_assoc.
- rewrite <- (Pos.mul_comm (Pos.max (Pos.max Ax Ay) Az)).
- rewrite <- Pos.mul_assoc.
- rewrite Pos2Nat.inj_mul.
- rewrite (Pos2Nat.inj_mul (Pos.max (Pos.max Ax Ay) Az)).
- apply Nat.mul_le_mono_nonneg. apply le_0_n.
- apply Pos2Nat.inj_le. apply Pos.le_max_l.
- apply le_0_n. apply le_refl. apply H. apply le_refl.
- rewrite max_l. apply (le_trans _ k).
- apply (le_trans _ (Pos.to_nat (2 * Pos.max (Pos.max Ax Ay) Az * (2 * p)))).
- rewrite (Pos.mul_comm 2). rewrite <- Pos.mul_assoc.
- rewrite <- (Pos.mul_comm (Pos.max (Pos.max Ax Ay) Az)).
- rewrite <- Pos.mul_assoc.
- rewrite Pos2Nat.inj_mul.
- rewrite (Pos2Nat.inj_mul (Pos.max (Pos.max Ax Ay) Az)).
- apply Nat.mul_le_mono_nonneg. apply le_0_n.
- apply Pos2Nat.inj_le. apply Pos.le_max_l.
- apply le_0_n. apply le_refl. apply H0.
- rewrite <- (mult_1_l k). rewrite mult_assoc.
- apply Nat.mul_le_mono_nonneg. auto.
- rewrite mult_1_r. apply Pos2Nat.is_pos. apply le_0_n.
- apply le_refl. apply le_refl.
- + apply Qlt_le_weak.
- pose proof (CReal_mult_cauchy xn xn zn Ax Az Pos.to_nat limx limz).
- apply H1. apply majx. apply majz. rewrite max_l. 2: apply le_refl.
- apply (le_trans _ (Pos.to_nat (2 * Pos.max (Pos.max Ax Ay) Az * (2 * p)))).
- rewrite (Pos.mul_comm 2). rewrite <- Pos.mul_assoc.
- rewrite <- (Pos.mul_comm (Pos.max (Pos.max Ax Ay) Az)).
- rewrite <- Pos.mul_assoc.
- rewrite Pos2Nat.inj_mul.
- rewrite (Pos2Nat.inj_mul (Pos.max (Pos.max Ax Ay) Az)).
- apply Nat.mul_le_mono_nonneg. apply le_0_n.
- rewrite <- Pos.max_assoc. rewrite (Pos.max_comm Ay Az).
- rewrite Pos.max_assoc. apply Pos2Nat.inj_le. apply Pos.le_max_l.
- apply le_0_n. apply le_refl. apply H.
- rewrite max_l. apply (le_trans _ k).
- apply (le_trans _ (Pos.to_nat (2 * Pos.max (Pos.max Ax Ay) Az * (2 * p)))).
- rewrite (Pos.mul_comm 2). rewrite <- Pos.mul_assoc.
- rewrite <- (Pos.mul_comm (Pos.max (Pos.max Ax Ay) Az)).
- rewrite <- Pos.mul_assoc.
- rewrite Pos2Nat.inj_mul.
- rewrite (Pos2Nat.inj_mul (Pos.max (Pos.max Ax Ay) Az)).
- apply Nat.mul_le_mono_nonneg. apply le_0_n.
- rewrite <- Pos.max_assoc. rewrite (Pos.max_comm Ay Az).
- rewrite Pos.max_assoc. apply Pos2Nat.inj_le. apply Pos.le_max_l.
- apply le_0_n. apply le_refl. apply H0.
- rewrite <- (mult_1_l k). rewrite mult_assoc.
- apply Nat.mul_le_mono_nonneg. auto.
- rewrite mult_1_r. apply Pos2Nat.is_pos. apply le_0_n.
- apply le_refl. apply le_refl.
- + rewrite Qinv_plus_distr. unfold Qeq. reflexivity.
-Qed.
-
-Lemma CReal_mult_plus_distr_r : forall r1 r2 r3 : CReal,
- (r2 + r3) * r1 == (r2 * r1) + (r3 * r1).
-Proof.
- intros.
- rewrite CReal_mult_comm, CReal_mult_plus_distr_l,
- <- (CReal_mult_comm r1), <- (CReal_mult_comm r1).
- reflexivity.
-Qed.
-
-Lemma CReal_mult_1_l : forall r: CReal, 1 * r == r.
-Proof.
- intros [rn limr]. split.
- - intros [m maj]. simpl in maj.
- destruct (QCauchySeq_bounded (fun _ : nat => 1%Q) Pos.to_nat (ConstCauchy 1)).
- destruct (QCauchySeq_bounded rn Pos.to_nat limr).
- simpl in maj. rewrite Qmult_1_l in maj.
- specialize (limr m).
- apply (Qlt_not_le (2 # m) (1 # m)).
- apply (Qlt_trans _ (rn (Pos.to_nat m) - rn (Pos.to_nat (Pos.max x x0)~0 * Pos.to_nat m)%nat)).
- apply maj.
- apply (Qle_lt_trans _ (Qabs (rn (Pos.to_nat m) - rn (Pos.to_nat (Pos.max x x0)~0 * Pos.to_nat m)%nat))).
- apply Qle_Qabs. apply limr. apply le_refl.
- rewrite <- (mult_1_l (Pos.to_nat m)). rewrite mult_assoc.
- apply Nat.mul_le_mono_nonneg. auto. rewrite mult_1_r.
- apply Pos2Nat.is_pos. apply le_0_n. apply le_refl.
- apply Z.mul_le_mono_nonneg. discriminate. discriminate.
- discriminate. apply Z.le_refl.
- - intros [m maj]. simpl in maj.
- destruct (QCauchySeq_bounded (fun _ : nat => 1%Q) Pos.to_nat (ConstCauchy 1)).
- destruct (QCauchySeq_bounded rn Pos.to_nat limr).
- simpl in maj. rewrite Qmult_1_l in maj.
- specialize (limr m).
- apply (Qlt_not_le (2 # m) (1 # m)).
- apply (Qlt_trans _ (rn (Pos.to_nat (Pos.max x x0)~0 * Pos.to_nat m)%nat - rn (Pos.to_nat m))).
- apply maj.
- apply (Qle_lt_trans _ (Qabs (rn (Pos.to_nat (Pos.max x x0)~0 * Pos.to_nat m)%nat - rn (Pos.to_nat m)))).
- apply Qle_Qabs. apply limr.
- rewrite <- (mult_1_l (Pos.to_nat m)). rewrite mult_assoc.
- apply Nat.mul_le_mono_nonneg. auto. rewrite mult_1_r.
- apply Pos2Nat.is_pos. apply le_0_n. apply le_refl.
- apply le_refl. apply Z.mul_le_mono_nonneg. discriminate. discriminate.
- discriminate. apply Z.le_refl.
-Qed.
-
-Lemma CReal_isRingExt : ring_eq_ext CReal_plus CReal_mult CReal_opp CRealEq.
-Proof.
- split.
- - intros x y H z t H0. apply CReal_plus_morph; assumption.
- - intros x y H z t H0. apply (CRealEq_trans _ (CReal_mult x t)).
- apply CReal_mult_proper_l. apply H0.
- apply (CRealEq_trans _ (CReal_mult t x)). apply CReal_mult_comm.
- apply (CRealEq_trans _ (CReal_mult t y)).
- apply CReal_mult_proper_l. apply H. apply CReal_mult_comm.
- - intros x y H. apply (CReal_plus_eq_reg_l x).
- apply (CRealEq_trans _ (inject_Q 0)). apply CReal_plus_opp_r.
- apply (CRealEq_trans _ (CReal_plus y (CReal_opp y))).
- apply CRealEq_sym. apply CReal_plus_opp_r.
- apply CReal_plus_proper_r. apply CRealEq_sym. apply H.
-Qed.
-
-Lemma CReal_isRing : ring_theory (inject_Q 0) (inject_Q 1)
- CReal_plus CReal_mult
- CReal_minus CReal_opp
- CRealEq.
-Proof.
- intros. split.
- - apply CReal_plus_0_l.
- - apply CReal_plus_comm.
- - intros x y z. symmetry. apply CReal_plus_assoc.
- - apply CReal_mult_1_l.
- - apply CReal_mult_comm.
- - intros x y z. symmetry. apply CReal_mult_assoc.
- - intros x y z. rewrite <- (CReal_mult_comm z).
- rewrite CReal_mult_plus_distr_l.
- apply (CRealEq_trans _ (CReal_plus (CReal_mult x z) (CReal_mult z y))).
- apply CReal_plus_proper_r. apply CReal_mult_comm.
- apply CReal_plus_proper_l. apply CReal_mult_comm.
- - intros x y. apply CRealEq_refl.
- - apply CReal_plus_opp_r.
-Qed.
-
-Add Parametric Morphism : CReal_mult
- with signature CRealEq ==> CRealEq ==> CRealEq
- as CReal_mult_morph.
-Proof.
- apply CReal_isRingExt.
-Qed.
-
-Instance CReal_mult_morph_T
- : CMorphisms.Proper
- (CMorphisms.respectful CRealEq (CMorphisms.respectful CRealEq CRealEq)) CReal_mult.
-Proof.
- apply CReal_isRingExt.
-Qed.
-
-Add Parametric Morphism : CReal_opp
- with signature CRealEq ==> CRealEq
- as CReal_opp_morph.
-Proof.
- apply (Ropp_ext CReal_isRingExt).
-Qed.
-
-Instance CReal_opp_morph_T
- : CMorphisms.Proper
- (CMorphisms.respectful CRealEq CRealEq) CReal_opp.
-Proof.
- apply CReal_isRingExt.
-Qed.
-
-Add Parametric Morphism : CReal_minus
- with signature CRealEq ==> CRealEq ==> CRealEq
- as CReal_minus_morph.
-Proof.
- intros. unfold CReal_minus. rewrite H,H0. reflexivity.
-Qed.
-
-Instance CReal_minus_morph_T
- : CMorphisms.Proper
- (CMorphisms.respectful CRealEq (CMorphisms.respectful CRealEq CRealEq)) CReal_minus.
-Proof.
- intros x y exy z t ezt. unfold CReal_minus. rewrite exy,ezt. reflexivity.
-Qed.
-
-Add Ring CRealRing : CReal_isRing.
-
Lemma CReal_opp_0 : -0 == 0.
Proof.
- ring.
+ apply (CReal_plus_eq_reg_l 0).
+ rewrite CReal_plus_0_r, CReal_plus_opp_r. reflexivity.
Qed.
Lemma CReal_opp_plus_distr : forall r1 r2, - (r1 + r2) == - r1 + - r2.
Proof.
- intros; ring.
+ intros. apply (CReal_plus_eq_reg_l (r1+r2)).
+ rewrite CReal_plus_opp_r, (CReal_plus_comm (-r1)), CReal_plus_assoc.
+ rewrite <- (CReal_plus_assoc r2), CReal_plus_opp_r, CReal_plus_0_l.
+ rewrite CReal_plus_opp_r. reflexivity.
Qed.
Lemma CReal_opp_involutive : forall x:CReal, --x == x.
Proof.
- intro x. ring.
+ intros. apply (CReal_plus_eq_reg_l (-x)).
+ rewrite CReal_plus_opp_l, CReal_plus_opp_r. reflexivity.
Qed.
Lemma CReal_opp_gt_lt_contravar : forall r1 r2, r1 > r2 -> - r1 < - r2.
Proof.
unfold CRealGt; intros.
apply (CReal_plus_lt_reg_l (r2 + r1)).
- setoid_replace (r2 + r1 + - r1) with r2 by ring.
- setoid_replace (r2 + r1 + - r2) with r1 by ring.
- exact H.
-Qed.
-
-(**********)
-Lemma CReal_mult_0_l : forall r, 0 * r == 0.
-Proof.
- intro; ring.
-Qed.
-
-Lemma CReal_mult_0_r : forall r, r * 0 == 0.
-Proof.
- intro; ring.
-Qed.
-
-(**********)
-Lemma CReal_mult_1_r : forall r, r * 1 == r.
-Proof.
- intro; ring.
-Qed.
-
-Lemma CReal_opp_mult_distr_l
- : forall r1 r2 : CReal, CRealEq (CReal_opp (CReal_mult r1 r2))
- (CReal_mult (CReal_opp r1) r2).
-Proof.
- intros. ring.
+ rewrite CReal_plus_assoc, CReal_plus_opp_r, CReal_plus_0_r.
+ rewrite CReal_plus_comm, <- CReal_plus_assoc, CReal_plus_opp_l.
+ rewrite CReal_plus_0_l. exact H.
Qed.
-Lemma CReal_mult_lt_compat_l : forall x y z : CReal,
- 0 < x -> y < z -> x*y < x*z.
+Lemma CReal_opp_ge_le_contravar : forall r1 r2, r1 >= r2 -> - r1 <= - r2.
Proof.
- intros. apply (CReal_plus_lt_reg_l
- (CReal_opp (CReal_mult x y))).
- rewrite CReal_plus_comm. pose proof CReal_plus_opp_r.
- unfold CReal_minus in H1. rewrite H1.
- rewrite CReal_mult_comm, CReal_opp_mult_distr_l, CReal_mult_comm.
- rewrite <- CReal_mult_plus_distr_l.
- apply CReal_mult_lt_0_compat. exact H.
- apply (CReal_plus_lt_reg_l y).
- rewrite CReal_plus_comm, CReal_plus_0_l.
- rewrite <- CReal_plus_assoc, H1, CReal_plus_0_l. exact H0.
+ intros. intro abs. apply H. clear H.
+ apply (CReal_plus_lt_reg_r (-r1-r2)).
+ unfold CReal_minus. rewrite <- CReal_plus_assoc, CReal_plus_opp_r, CReal_plus_0_l.
+ rewrite (CReal_plus_comm (-r1)), <- CReal_plus_assoc, CReal_plus_opp_r, CReal_plus_0_l.
+ exact abs.
Qed.
-Lemma CReal_mult_lt_compat_r : forall x y z : CReal,
- 0 < x -> y < z -> y*x < z*x.
+Lemma inject_Q_plus : forall q r : Q,
+ inject_Q (q + r) == inject_Q q + inject_Q r.
Proof.
- intros. rewrite <- (CReal_mult_comm x), <- (CReal_mult_comm x).
- apply (CReal_mult_lt_compat_l x); assumption.
-Qed.
-
-Lemma CReal_mult_eq_reg_l : forall (r r1 r2 : CReal),
- r # 0
- -> CRealEq (CReal_mult r r1) (CReal_mult r r2)
- -> CRealEq r1 r2.
-Proof.
- intros. destruct H; split.
- - intro abs. apply (CReal_mult_lt_compat_l (-r)) in abs.
- rewrite <- CReal_opp_mult_distr_l, <- CReal_opp_mult_distr_l, H0 in abs.
- exact (CRealLt_irrefl _ abs). apply (CReal_plus_lt_reg_l r).
- rewrite CReal_plus_opp_r, CReal_plus_comm, CReal_plus_0_l. exact c.
- - intro abs. apply (CReal_mult_lt_compat_l (-r)) in abs.
- rewrite <- CReal_opp_mult_distr_l, <- CReal_opp_mult_distr_l, H0 in abs.
- exact (CRealLt_irrefl _ abs). apply (CReal_plus_lt_reg_l r).
- rewrite CReal_plus_opp_r, CReal_plus_comm, CReal_plus_0_l. exact c.
- - intro abs. apply (CReal_mult_lt_compat_l r) in abs. rewrite H0 in abs.
- exact (CRealLt_irrefl _ abs). exact c.
- - intro abs. apply (CReal_mult_lt_compat_l r) in abs. rewrite H0 in abs.
- exact (CRealLt_irrefl _ abs). exact c.
-Qed.
-
-
-
-(*********************************************************)
-(** * Field *)
-(*********************************************************)
-
-(**********)
-Fixpoint INR (n:nat) : CReal :=
- match n with
- | O => 0
- | S O => 1
- | S n => INR n + 1
- end.
-Arguments INR n%nat.
-
-(* compact representation for 2*p *)
-Fixpoint IPR_2 (p:positive) : CReal :=
- match p with
- | xH => 1 + 1
- | xO p => IPR_2 p + IPR_2 p
- | xI p => (1 + IPR_2 p) + (1 + IPR_2 p)
- end.
-
-Definition IPR (p:positive) : CReal :=
- match p with
- | xH => 1
- | xO p => IPR_2 p
- | xI p => 1 + IPR_2 p
- end.
-Arguments IPR p%positive : simpl never.
-
-(**********)
-Definition IZR (z:Z) : CReal :=
- match z with
- | Z0 => 0
- | Zpos n => IPR n
- | Zneg n => - IPR n
- end.
-Arguments IZR z%Z : simpl never.
-
-Notation "2" := (IZR 2) : CReal_scope.
-
-(**********)
-Lemma S_INR : forall n:nat, INR (S n) == INR n + 1.
-Proof.
- intro; destruct n. rewrite CReal_plus_0_l. reflexivity. reflexivity.
-Qed.
-
-Lemma le_succ_r_T : forall n m : nat, (n <= S m)%nat -> {(n <= m)%nat} + {n = S m}.
-Proof.
- intros. destruct (le_lt_dec n m). left. exact l.
- right. apply Nat.le_succ_r in H. destruct H.
- exfalso. apply (le_not_lt n m); assumption. exact H.
-Qed.
-
-Lemma lt_INR : forall n m:nat, (n < m)%nat -> INR n < INR m.
-Proof.
- induction m.
- - intros. exfalso. inversion H.
- - intros. unfold lt in H. apply le_S_n in H. destruct m.
- assert (n = 0)%nat.
- { inversion H. reflexivity. }
- subst n. apply CRealLt_0_1. apply le_succ_r_T in H. destruct H.
- rewrite S_INR. apply (CRealLt_trans _ (INR (S m) + 0)).
- rewrite CReal_plus_comm, CReal_plus_0_l. apply IHm.
- apply le_n_S. exact l.
- apply CReal_plus_lt_compat_l. exact CRealLt_0_1.
- subst n. rewrite (S_INR (S m)). rewrite <- (CReal_plus_0_l).
- rewrite (CReal_plus_comm 0), CReal_plus_assoc.
- apply CReal_plus_lt_compat_l. rewrite CReal_plus_0_l.
- exact CRealLt_0_1.
-Qed.
-
-(**********)
-Lemma S_O_plus_INR : forall n:nat, INR (1 + n) == INR 1 + INR n.
-Proof.
- intros; destruct n.
- - rewrite CReal_plus_comm, CReal_plus_0_l. reflexivity.
- - rewrite CReal_plus_comm. reflexivity.
-Qed.
-
-(**********)
-Lemma plus_INR : forall n m:nat, INR (n + m) == INR n + INR m.
-Proof.
- intros n m; induction n as [| n Hrecn].
- - rewrite CReal_plus_0_l. reflexivity.
- - replace (S n + m)%nat with (S (n + m)); auto with arith.
- repeat rewrite S_INR.
- rewrite Hrecn; ring.
-Qed.
-
-(**********)
-Lemma minus_INR : forall n m:nat, (m <= n)%nat -> INR (n - m) == INR n - INR m.
-Proof.
- intros n m le; pattern m, n; apply le_elim_rel.
- intros. rewrite <- minus_n_O. unfold CReal_minus.
- unfold INR. ring.
- intros; repeat rewrite S_INR; simpl.
- unfold CReal_minus. rewrite H0. ring. exact le.
-Qed.
-
-(*********)
-Lemma mult_INR : forall n m:nat, INR (n * m) == INR n * INR m.
-Proof.
- intros n m; induction n as [| n Hrecn].
- - rewrite CReal_mult_0_l. reflexivity.
- - intros; repeat rewrite S_INR; simpl.
- rewrite plus_INR. rewrite Hrecn; ring.
-Qed.
-
-(**********)
-Lemma IZN : forall n:Z, (0 <= n)%Z -> { m : nat | n = Z.of_nat m }.
-Proof.
- intros. exists (Z.to_nat n). rewrite Z2Nat.id. reflexivity. assumption.
-Qed.
-
-Lemma INR_IPR : forall p, INR (Pos.to_nat p) == IPR p.
-Proof.
- assert (H: forall p, INR (Pos.to_nat p) + INR (Pos.to_nat p) == IPR_2 p).
- { induction p as [p|p|].
- - unfold IPR_2; rewrite Pos2Nat.inj_xI, S_INR, mult_INR, <- IHp.
- setoid_replace (INR 2) with (1 + 1). 2: reflexivity. ring.
- - unfold IPR_2; rewrite Pos2Nat.inj_xO, mult_INR, <- IHp.
- setoid_replace (INR 2) with (1 + 1). 2: reflexivity. ring.
- - reflexivity. }
- intros [p|p|] ; unfold IPR.
- rewrite Pos2Nat.inj_xI, S_INR, mult_INR, <- H.
- setoid_replace (INR 2) with (1 + 1). 2: reflexivity. ring.
- rewrite Pos2Nat.inj_xO, mult_INR, <- H.
- setoid_replace (INR 2) with (1 + 1). 2: reflexivity. ring.
- easy.
-Qed.
-
-(* This is stronger than Req to injectQ, because it
- concerns all the rational sequence, not only its limit. *)
-Lemma FinjectP2_CReal : forall (p:positive) (k:nat),
- (proj1_sig (IPR_2 p) k == Z.pos p~0 # 1)%Q.
-Proof.
- induction p.
- - intros. replace (IPR_2 p~1) with (1 + IPR_2 p + (1+ IPR_2 p)).
- 2: reflexivity. do 2 rewrite CReal_plus_nth. rewrite IHp.
- simpl. rewrite Pos2Z.inj_xO, (Pos2Z.inj_xO (p~1)), Pos2Z.inj_xI.
- generalize (2*Z.pos p)%Z. intro z.
- do 2 rewrite Qinv_plus_distr. apply f_equal2.
- 2: reflexivity. unfold Qnum. ring.
- - intros. replace (IPR_2 p~0) with (IPR_2 p + IPR_2 p).
- 2: reflexivity. rewrite CReal_plus_nth, IHp.
- rewrite Qinv_plus_distr. apply f_equal2. 2: reflexivity.
- unfold Qnum. rewrite (Pos2Z.inj_xO (p~0)). ring.
- - intros. reflexivity.
-Qed.
-
-Lemma FinjectP_CReal : forall (p:positive) (k:nat),
- (proj1_sig (IPR p) k == Z.pos p # 1)%Q.
-Proof.
- destruct p.
- - intros. unfold IPR.
- rewrite CReal_plus_nth, FinjectP2_CReal. unfold Qeq; simpl.
- rewrite Pos.mul_1_r. reflexivity.
- - intros. unfold IPR. rewrite FinjectP2_CReal. reflexivity.
- - intros. reflexivity.
-Qed.
-
-(* Inside this Cauchy real implementation, we can give
- an instantaneous witness of this inequality, because
- we know a priori that it will work. *)
-Lemma IPR_pos : forall p:positive, 0 < IPR p.
-Proof.
- intro p. exists 3%positive. simpl.
- rewrite FinjectP_CReal. apply (Qlt_le_trans _ 1). reflexivity.
- unfold Qle; simpl.
- rewrite <- (Zpos_max_1 (p*1*1)). apply Z.le_max_l.
-Defined.
-
-Lemma IPR_double : forall p:positive, IPR (2*p) == 2 * IPR p.
-Proof.
- intro p.
- destruct p; rewrite (CReal_mult_plus_distr_r _ 1 1), CReal_mult_1_l; reflexivity.
-Qed.
-
-(**********)
-Lemma INR_IZR_INZ : forall n:nat, INR n == IZR (Z.of_nat n).
-Proof.
- intros [|n].
- easy.
- simpl Z.of_nat. unfold IZR.
- now rewrite <- INR_IPR, SuccNat2Pos.id_succ.
-Qed.
-
-Lemma plus_IZR_NEG_POS :
- forall p q:positive, IZR (Zpos p + Zneg q) == IZR (Zpos p) + IZR (Zneg q).
-Proof.
- intros p q; simpl. rewrite Z.pos_sub_spec.
- case Pos.compare_spec; intros H; unfold IZR.
- subst. ring.
- rewrite <- 3!INR_IPR, Pos2Nat.inj_sub.
- rewrite minus_INR.
- 2: (now apply lt_le_weak, Pos2Nat.inj_lt).
- ring.
- trivial.
- rewrite <- 3!INR_IPR, Pos2Nat.inj_sub.
- rewrite minus_INR.
- 2: (now apply lt_le_weak, Pos2Nat.inj_lt).
- ring. trivial.
-Qed.
-
-Lemma plus_IPR : forall n m:positive, IPR (n + m) == IPR n + IPR m.
-Proof.
- intros. repeat rewrite <- INR_IPR.
- rewrite Pos2Nat.inj_add. apply plus_INR.
-Qed.
-
-(**********)
-Lemma plus_IZR : forall n m:Z, IZR (n + m) == IZR n + IZR m.
-Proof.
- intro z; destruct z; intro t; destruct t; intros.
- - rewrite CReal_plus_0_l. reflexivity.
- - rewrite CReal_plus_0_l. rewrite Z.add_0_l. reflexivity.
- - rewrite CReal_plus_0_l. reflexivity.
- - rewrite CReal_plus_comm,CReal_plus_0_l. reflexivity.
- - rewrite <- Pos2Z.inj_add. unfold IZR. apply plus_IPR.
- - apply plus_IZR_NEG_POS.
- - rewrite CReal_plus_comm,CReal_plus_0_l, Z.add_0_r. reflexivity.
- - rewrite Z.add_comm; rewrite CReal_plus_comm; apply plus_IZR_NEG_POS.
- - simpl. unfold IZR. rewrite <- 3!INR_IPR, Pos2Nat.inj_add, plus_INR.
- ring.
-Qed.
-
-Lemma mult_IPR : forall n m:positive, IPR (n * m) == IPR n * IPR m.
-Proof.
- intros. repeat rewrite <- INR_IPR.
- rewrite Pos2Nat.inj_mul. apply mult_INR.
-Qed.
-
-Lemma mult_IZR : forall n m:Z, IZR (n * m) == IZR n * IZR m.
-Proof.
- intros n m. destruct n.
- - rewrite CReal_mult_0_l. rewrite Z.mul_0_l. reflexivity.
- - destruct m. rewrite Z.mul_0_r, CReal_mult_0_r. reflexivity.
- simpl; unfold IZR. apply mult_IPR.
- simpl. unfold IZR. rewrite mult_IPR. ring.
- - destruct m. rewrite Z.mul_0_r, CReal_mult_0_r. reflexivity.
- simpl. unfold IZR. rewrite mult_IPR. ring.
- simpl. unfold IZR. rewrite mult_IPR. ring.
-Qed.
-
-Lemma opp_IZR : forall n:Z, IZR (- n) == - IZR n.
-Proof.
- intros [|z|z]; unfold IZR. rewrite CReal_opp_0. reflexivity.
- reflexivity. rewrite CReal_opp_involutive. reflexivity.
-Qed.
-
-Lemma minus_IZR : forall n m:Z, IZR (n - m) == IZR n - IZR m.
-Proof.
- intros; unfold Z.sub, CReal_minus.
- rewrite <- opp_IZR.
- apply plus_IZR.
-Qed.
-
-Lemma IZR_lt : forall n m:Z, (n < m)%Z -> IZR n < IZR m.
-Proof.
- assert (forall n:Z, Z.lt 0 n -> 0 < IZR n) as posCase.
- { intros. destruct (IZN n). apply Z.lt_le_incl. apply H.
- subst n. rewrite <- INR_IZR_INZ. apply (lt_INR 0).
- apply Nat2Z.inj_lt. apply H. }
- intros. apply (CReal_plus_lt_reg_r (-(IZR n))).
- pose proof minus_IZR. unfold CReal_minus in H0.
- repeat rewrite <- H0. unfold Zminus.
- rewrite Z.add_opp_diag_r. apply posCase.
- rewrite (Z.add_lt_mono_l _ _ n). ring_simplify. apply H.
-Qed.
-
-Lemma Z_R_minus : forall n m:Z, IZR n - IZR m == IZR (n - m).
-Proof.
- intros z1 z2; unfold CReal_minus; unfold Z.sub.
- rewrite plus_IZR, opp_IZR. reflexivity.
-Qed.
-
-Lemma lt_0_IZR : forall n:Z, 0 < IZR n -> (0 < n)%Z.
-Proof.
- intro z; case z; simpl; intros.
- elim (CRealLt_irrefl _ H).
- easy. exfalso.
- apply (CRealLt_asym 0 (IZR (Z.neg p))). exact H.
- apply (IZR_lt (Z.neg p) 0). reflexivity.
-Qed.
-
-Lemma lt_IZR : forall n m:Z, IZR n < IZR m -> (n < m)%Z.
-Proof.
- intros z1 z2 H; apply Z.lt_0_sub.
- apply lt_0_IZR.
- rewrite <- Z_R_minus. apply (CReal_plus_lt_reg_l (IZR z1)).
- ring_simplify. exact H.
-Qed.
-
-Lemma IZR_le : forall n m:Z, (n <= m)%Z -> IZR n <= IZR m.
-Proof.
- intros m n H. intro abs. apply (lt_IZR n m) in abs. omega.
-Qed.
-
-Lemma CReal_iterate_one : forall (n : nat),
- IZR (Z.of_nat n) == inject_Q (Z.of_nat n # 1).
-Proof.
- induction n.
- - apply CRealEq_refl.
- - replace (Z.of_nat (S n)) with (1 + Z.of_nat n)%Z.
- rewrite plus_IZR.
- rewrite IHn. clear IHn. apply CRealEq_diff. intro k. simpl.
- rewrite Z.mul_1_r. rewrite Z.mul_1_r. rewrite Z.mul_1_r.
- rewrite Z.add_opp_diag_r. discriminate.
- replace (S n) with (1 + n)%nat. 2: reflexivity.
- rewrite (Nat2Z.inj_add 1 n). reflexivity.
-Qed.
-
-(* The constant sequences of rationals are CRealEq to
- the rational operations on the unity. *)
-Lemma FinjectZ_CReal : forall z : Z,
- IZR z == inject_Q (z # 1).
-Proof.
- intros. destruct z.
- - apply CRealEq_refl.
- - simpl. pose proof (CReal_iterate_one (Pos.to_nat p)).
- rewrite positive_nat_Z in H. apply H.
- - simpl. apply (CReal_plus_eq_reg_l (IZR (Z.pos p))).
- pose proof CReal_plus_opp_r. rewrite H.
- pose proof (CReal_iterate_one (Pos.to_nat p)).
- rewrite positive_nat_Z in H0. rewrite H0.
- apply CRealEq_diff. intro n. simpl. rewrite Z.pos_sub_diag.
- discriminate.
-Qed.
-
-
-(* Axiom Rarchimed_constr *)
-Lemma Rarchimedean
- : forall x:CReal,
- { n:Z & x < IZR n < x+2 }.
-Proof.
- (* Locate x within 1/4 and pick the first integer above this interval. *)
- intros [xn limx].
- pose proof (Qlt_floor (xn 4%nat + (1#4))). unfold inject_Z in H.
- pose proof (Qfloor_le (xn 4%nat + (1#4))). unfold inject_Z in H0.
- remember (Qfloor (xn 4%nat + (1#4)))%Z as n.
- exists (n+1)%Z. split.
- - rewrite FinjectZ_CReal.
- assert (Qlt 0 ((n + 1 # 1) - (xn 4%nat + (1 # 4)))) as epsPos.
- { unfold Qminus. rewrite <- Qlt_minus_iff. exact H. }
- destruct (Qarchimedean (/((1#2)*((n + 1 # 1) - (xn 4%nat + (1 # 4)))))) as [k kmaj].
- exists (Pos.max 4 k). simpl.
- apply (Qlt_trans _ ((n + 1 # 1) - (xn 4%nat + (1 # 4)))).
- + setoid_replace (Z.pos k # 1)%Q with (/(1#k))%Q in kmaj. 2: reflexivity.
- rewrite <- Qinv_lt_contravar in kmaj. 2: reflexivity.
- apply (Qle_lt_trans _ (2#k)).
- rewrite <- (Qmult_le_l _ _ (1#2)).
- setoid_replace ((1 # 2) * (2 # k))%Q with (1#k)%Q. 2: reflexivity.
- setoid_replace ((1 # 2) * (2 # Pos.max 4 k))%Q with (1#Pos.max 4 k)%Q. 2: reflexivity.
- unfold Qle; simpl. apply Pos2Z.pos_le_pos. apply Pos.le_max_r.
- reflexivity.
- rewrite <- (Qmult_lt_l _ _ (1#2)).
- setoid_replace ((1 # 2) * (2 # k))%Q with (1#k)%Q. exact kmaj.
- reflexivity. reflexivity. rewrite <- (Qmult_0_r (1#2)).
- rewrite Qmult_lt_l. exact epsPos. reflexivity.
- + rewrite <- (Qplus_lt_r _ _ (xn (Pos.to_nat (Pos.max 4 k)) - (n + 1 # 1) + (1#4))).
- ring_simplify.
- apply (Qle_lt_trans _ (Qabs (xn (Pos.to_nat (Pos.max 4 k)) - xn 4%nat))).
- apply Qle_Qabs. apply limx.
- rewrite Pos2Nat.inj_max. apply Nat.le_max_l. apply le_refl.
- - apply (CReal_plus_lt_reg_l (-IZR 2)). ring_simplify.
- do 2 rewrite FinjectZ_CReal.
- exists 4%positive. simpl.
- rewrite <- Qinv_plus_distr.
- rewrite <- (Qplus_lt_r _ _ ((n#1) - (1#2))). ring_simplify.
- apply (Qle_lt_trans _ (xn 4%nat + (1 # 4)) _ H0).
- unfold Pos.to_nat; simpl.
- rewrite <- (Qplus_lt_r _ _ (-xn 4%nat)). ring_simplify.
- reflexivity.
-Qed.
-
-Lemma CRealLtDisjunctEpsilon : forall a b c d : CReal,
- (CRealLtProp a b \/ CRealLtProp c d) -> CRealLt a b + CRealLt c d.
-Proof.
- intros.
- assert (exists n : nat, n <> O /\
- (Qlt (2 # Pos.of_nat n) (proj1_sig b n - proj1_sig a n)
- \/ Qlt (2 # Pos.of_nat n) (proj1_sig d n - proj1_sig c n))).
- { destruct H. destruct H as [n maj]. exists (Pos.to_nat n). split.
- intro abs. destruct (Pos2Nat.is_succ n). rewrite H in abs.
- inversion abs. left. rewrite Pos2Nat.id. apply maj.
- destruct H as [n maj]. exists (Pos.to_nat n). split.
- intro abs. destruct (Pos2Nat.is_succ n). rewrite H in abs.
- inversion abs. right. rewrite Pos2Nat.id. apply maj. }
- apply constructive_indefinite_ground_description_nat in H0.
- - destruct H0 as [n [nPos maj]].
- destruct (Qlt_le_dec (2 # Pos.of_nat n)
- (proj1_sig b n - proj1_sig a n)).
- left. exists (Pos.of_nat n). rewrite Nat2Pos.id. apply q. apply nPos.
- assert (2 # Pos.of_nat n < proj1_sig d n - proj1_sig c n)%Q.
- destruct maj. exfalso.
- apply (Qlt_not_le (2 # Pos.of_nat n) (proj1_sig b n - proj1_sig a n)); assumption.
- assumption. clear maj. right. exists (Pos.of_nat n). rewrite Nat2Pos.id.
- apply H0. apply nPos.
- - clear H0. clear H. intro n. destruct n. right.
- intros [abs _]. exact (abs (eq_refl O)).
- destruct (Qlt_le_dec (2 # Pos.of_nat (S n)) (proj1_sig b (S n) - proj1_sig a (S n))).
- left. split. discriminate. left. apply q.
- destruct (Qlt_le_dec (2 # Pos.of_nat (S n)) (proj1_sig d (S n) - proj1_sig c (S n))).
- left. split. discriminate. right. apply q0.
- right. intros [_ [abs|abs]].
- apply (Qlt_not_le (2 # Pos.of_nat (S n))
- (proj1_sig b (S n) - proj1_sig a (S n))); assumption.
- apply (Qlt_not_le (2 # Pos.of_nat (S n))
- (proj1_sig d (S n) - proj1_sig c (S n))); assumption.
-Qed.
-
-Lemma CRealShiftReal : forall (x : CReal) (k : nat),
- QCauchySeq (fun n => proj1_sig x (plus n k)) Pos.to_nat.
-Proof.
- intros x k n p q H H0.
- destruct x as [xn cau]; unfold proj1_sig.
- destruct k. rewrite plus_0_r. rewrite plus_0_r. apply cau; assumption.
- specialize (cau (n + Pos.of_nat (S k))%positive (p + S k)%nat (q + S k)%nat).
- apply (Qlt_trans _ (1 # n + Pos.of_nat (S k))).
- apply cau. rewrite Pos2Nat.inj_add. rewrite Nat2Pos.id.
- apply Nat.add_le_mono_r. apply H. discriminate.
- rewrite Pos2Nat.inj_add. rewrite Nat2Pos.id.
- apply Nat.add_le_mono_r. apply H0. discriminate.
- apply Pos2Nat.inj_lt; simpl. rewrite Pos2Nat.inj_add.
- rewrite <- (plus_0_r (Pos.to_nat n)). rewrite <- plus_assoc.
- apply Nat.add_lt_mono_l. apply Pos2Nat.is_pos.
-Qed.
-
-Lemma CRealShiftEqual : forall (x : CReal) (k : nat),
- CRealEq x (exist _ (fun n => proj1_sig x (plus n k)) (CRealShiftReal x k)).
-Proof.
- intros. split.
- - intros [n maj]. destruct x as [xn cau]; simpl in maj.
- specialize (cau n (Pos.to_nat n + k)%nat (Pos.to_nat n)).
- apply Qlt_not_le in maj. apply maj. clear maj.
- apply (Qle_trans _ (Qabs (xn (Pos.to_nat n + k)%nat - xn (Pos.to_nat n)))).
- apply Qle_Qabs. apply (Qle_trans _ (1#n)). apply Zlt_le_weak.
- apply cau. rewrite <- (plus_0_r (Pos.to_nat n)).
- rewrite <- plus_assoc. apply Nat.add_le_mono_l. apply le_0_n.
- apply le_refl. apply Z.mul_le_mono_pos_r. apply Pos2Z.is_pos.
- discriminate.
- - intros [n maj]. destruct x as [xn cau]; simpl in maj.
- specialize (cau n (Pos.to_nat n) (Pos.to_nat n + k)%nat).
- apply Qlt_not_le in maj. apply maj. clear maj.
- apply (Qle_trans _ (Qabs (xn (Pos.to_nat n) - xn (Pos.to_nat n + k)%nat))).
- apply Qle_Qabs. apply (Qle_trans _ (1#n)). apply Zlt_le_weak.
- apply cau. apply le_refl. rewrite <- (plus_0_r (Pos.to_nat n)).
- rewrite <- plus_assoc. apply Nat.add_le_mono_l. apply le_0_n.
- apply Z.mul_le_mono_pos_r. apply Pos2Z.is_pos. discriminate.
-Qed.
-
-(* Find an equal negative real number, which rational sequence
- stays below 0, so that it can be inversed. *)
-Definition CRealNegShift (x : CReal)
- : CRealLt x (inject_Q 0)
- -> { y : prod positive CReal | CRealEq x (snd y)
- /\ forall n:nat, Qlt (proj1_sig (snd y) n) (-1 # fst y) }.
-Proof.
- intro xNeg.
- pose proof (CRealLt_aboveSig x (inject_Q 0)).
- pose proof (CRealShiftReal x).
- pose proof (CRealShiftEqual x).
- destruct xNeg as [n maj], x as [xn cau]; simpl in maj.
- specialize (H n maj); simpl in H.
- destruct (Qarchimedean (/ (0 - xn (Pos.to_nat n) - (2 # n)))) as [a _].
- remember (Pos.max n a~0) as k.
- clear Heqk. clear maj. clear n.
- exists (pair k
- (exist _ (fun n => xn (plus n (Pos.to_nat k))) (H0 (Pos.to_nat k)))).
- split. apply H1. intro n. simpl. apply Qlt_minus_iff.
- destruct n.
- - specialize (H k).
- unfold Qminus in H. rewrite Qplus_0_l in H. apply Qlt_minus_iff in H.
- unfold Qminus. rewrite Qplus_comm.
- apply (Qlt_trans _ (- xn (Pos.to_nat k)%nat - (2 #k))). apply H.
- unfold Qminus. simpl. apply Qplus_lt_r.
- apply Z.mul_lt_mono_pos_r. simpl. apply Pos2Z.is_pos.
- reflexivity. apply Pos.le_refl.
- - apply (Qlt_trans _ (-(2 # k) - xn (S n + Pos.to_nat k)%nat)).
- rewrite <- (Nat2Pos.id (S n)). rewrite <- Pos2Nat.inj_add.
- specialize (H (Pos.of_nat (S n) + k)%positive).
- unfold Qminus in H. rewrite Qplus_0_l in H. apply Qlt_minus_iff in H.
- unfold Qminus. rewrite Qplus_comm. apply H. apply Pos2Nat.inj_le.
- rewrite <- (plus_0_l (Pos.to_nat k)). rewrite Pos2Nat.inj_add.
- apply Nat.add_le_mono_r. apply le_0_n. discriminate.
- apply Qplus_lt_l.
- apply Z.mul_lt_mono_pos_r. simpl. apply Pos2Z.is_pos.
- reflexivity.
-Qed.
-
-Definition CRealPosShift (x : CReal)
- : CRealLt (inject_Q 0) x
- -> { y : prod positive CReal | CRealEq x (snd y)
- /\ forall n:nat, Qlt (1 # fst y) (proj1_sig (snd y) n) }.
-Proof.
- intro xPos.
- pose proof (CRealLt_aboveSig (inject_Q 0) x).
- pose proof (CRealShiftReal x).
- pose proof (CRealShiftEqual x).
- destruct xPos as [n maj], x as [xn cau]; simpl in maj.
- simpl in H. specialize (H n).
- destruct (Qarchimedean (/ (xn (Pos.to_nat n) - 0 - (2 # n)))) as [a _].
- specialize (H maj); simpl in H.
- remember (Pos.max n a~0) as k.
- clear Heqk. clear maj. clear n.
- exists (pair k
- (exist _ (fun n => xn (plus n (Pos.to_nat k))) (H0 (Pos.to_nat k)))).
- split. apply H1. intro n. simpl. apply Qlt_minus_iff.
- destruct n.
- - specialize (H k).
- unfold Qminus in H. rewrite Qplus_0_r in H.
- simpl. rewrite <- Qlt_minus_iff.
- apply (Qlt_trans _ (2 #k)).
- apply Z.mul_lt_mono_pos_r. simpl. apply Pos2Z.is_pos.
- reflexivity. apply H. apply Pos.le_refl.
- - rewrite <- Qlt_minus_iff. apply (Qlt_trans _ (2 # k)).
- apply Z.mul_lt_mono_pos_r. simpl. apply Pos2Z.is_pos.
- reflexivity. specialize (H (Pos.of_nat (S n) + k)%positive).
- unfold Qminus in H. rewrite Qplus_0_r in H.
- rewrite Pos2Nat.inj_add in H. rewrite Nat2Pos.id in H.
- apply H. apply Pos2Nat.inj_le.
- rewrite <- (plus_0_l (Pos.to_nat k)). rewrite Pos2Nat.inj_add.
- apply Nat.add_le_mono_r. apply le_0_n. discriminate.
-Qed.
-
-Lemma CReal_inv_neg : forall (yn : nat -> Q) (k : positive),
- (QCauchySeq yn Pos.to_nat)
- -> (forall n : nat, yn n < -1 # k)%Q
- -> QCauchySeq (fun n : nat => / yn (Pos.to_nat k ^ 2 * n)%nat) Pos.to_nat.
-Proof.
- (* Prove the inverse sequence is Cauchy *)
- intros yn k cau maj n p q H0 H1.
- setoid_replace (/ yn (Pos.to_nat k ^ 2 * p)%nat -
- / yn (Pos.to_nat k ^ 2 * q)%nat)%Q
- with ((yn (Pos.to_nat k ^ 2 * q)%nat -
- yn (Pos.to_nat k ^ 2 * p)%nat)
- / (yn (Pos.to_nat k ^ 2 * q)%nat *
- yn (Pos.to_nat k ^ 2 * p)%nat)).
- + apply (Qle_lt_trans _ (Qabs (yn (Pos.to_nat k ^ 2 * q)%nat
- - yn (Pos.to_nat k ^ 2 * p)%nat)
- / (1 # (k^2)))).
- assert (1 # k ^ 2
- < Qabs (yn (Pos.to_nat k ^ 2 * q)%nat * yn (Pos.to_nat k ^ 2 * p)%nat))%Q.
- { rewrite Qabs_Qmult. unfold "^"%positive; simpl.
- rewrite factorDenom. rewrite Pos.mul_1_r.
- apply (Qlt_trans _ ((1#k) * Qabs (yn (Pos.to_nat k * (Pos.to_nat k * 1) * p)%nat))).
- apply Qmult_lt_l. reflexivity. rewrite Qabs_neg.
- specialize (maj (Pos.to_nat k * (Pos.to_nat k * 1) * p)%nat).
- apply Qlt_minus_iff in maj. apply Qlt_minus_iff.
- rewrite Qplus_comm. setoid_replace (-(1#k))%Q with (-1 # k)%Q. apply maj.
- reflexivity. apply (Qle_trans _ (-1 # k)). apply Zlt_le_weak.
- apply maj. discriminate.
- apply Qmult_lt_r. apply (Qlt_trans 0 (1#k)). reflexivity.
- rewrite Qabs_neg.
- specialize (maj (Pos.to_nat k * (Pos.to_nat k * 1) * p)%nat).
- apply Qlt_minus_iff in maj. apply Qlt_minus_iff.
- rewrite Qplus_comm. setoid_replace (-(1#k))%Q with (-1 # k)%Q. apply maj.
- reflexivity. apply (Qle_trans _ (-1 # k)). apply Zlt_le_weak.
- apply maj. discriminate.
- rewrite Qabs_neg.
- specialize (maj (Pos.to_nat k * (Pos.to_nat k * 1) * q)%nat).
- apply Qlt_minus_iff in maj. apply Qlt_minus_iff.
- rewrite Qplus_comm. setoid_replace (-(1#k))%Q with (-1 # k)%Q. apply maj.
- reflexivity. apply (Qle_trans _ (-1 # k)). apply Zlt_le_weak.
- apply maj. discriminate. }
- unfold Qdiv. rewrite Qabs_Qmult. rewrite Qabs_Qinv.
- rewrite Qmult_comm. rewrite <- (Qmult_comm (/ (1 # k ^ 2))).
- apply Qmult_le_compat_r. apply Qlt_le_weak.
- rewrite <- Qmult_1_l. apply Qlt_shift_div_r.
- apply (Qlt_trans 0 (1 # k ^ 2)). reflexivity. apply H.
- rewrite Qmult_comm. apply Qlt_shift_div_l.
- reflexivity. rewrite Qmult_1_l. apply H.
- apply Qabs_nonneg. simpl in maj.
- specialize (cau (n * (k^2))%positive
- (Pos.to_nat k ^ 2 * q)%nat
- (Pos.to_nat k ^ 2 * p)%nat).
- apply Qlt_shift_div_r. reflexivity.
- apply (Qlt_le_trans _ (1 # n * k ^ 2)). apply cau.
- rewrite Pos2Nat.inj_mul. rewrite mult_comm.
- unfold "^"%positive. simpl. rewrite Pos2Nat.inj_mul.
- rewrite <- mult_assoc. rewrite <- mult_assoc.
- apply Nat.mul_le_mono_nonneg_l. apply le_0_n.
- rewrite (mult_1_r). rewrite Pos.mul_1_r.
- apply Nat.mul_le_mono_nonneg_l. apply le_0_n.
- apply (le_trans _ (q+0)). rewrite plus_0_r. assumption.
- rewrite plus_0_r. apply le_refl.
- rewrite Pos2Nat.inj_mul. rewrite mult_comm.
- unfold "^"%positive; simpl. rewrite Pos2Nat.inj_mul.
- rewrite <- mult_assoc. rewrite <- mult_assoc.
- apply Nat.mul_le_mono_nonneg_l. apply le_0_n.
- rewrite (mult_1_r). rewrite Pos.mul_1_r.
- apply Nat.mul_le_mono_nonneg_l. apply le_0_n.
- apply (le_trans _ (p+0)). rewrite plus_0_r. assumption.
- rewrite plus_0_r. apply le_refl.
- rewrite factorDenom. apply Qle_refl.
- + field. split. intro abs.
- specialize (maj (Pos.to_nat k ^ 2 * p)%nat).
- rewrite abs in maj. inversion maj.
- intro abs.
- specialize (maj (Pos.to_nat k ^ 2 * q)%nat).
- rewrite abs in maj. inversion maj.
-Qed.
-
-Lemma CReal_inv_pos : forall (yn : nat -> Q) (k : positive),
- (QCauchySeq yn Pos.to_nat)
- -> (forall n : nat, 1 # k < yn n)%Q
- -> QCauchySeq (fun n : nat => / yn (Pos.to_nat k ^ 2 * n)%nat) Pos.to_nat.
-Proof.
- intros yn k cau maj n p q H0 H1.
- setoid_replace (/ yn (Pos.to_nat k ^ 2 * p)%nat -
- / yn (Pos.to_nat k ^ 2 * q)%nat)%Q
- with ((yn (Pos.to_nat k ^ 2 * q)%nat -
- yn (Pos.to_nat k ^ 2 * p)%nat)
- / (yn (Pos.to_nat k ^ 2 * q)%nat *
- yn (Pos.to_nat k ^ 2 * p)%nat)).
- + apply (Qle_lt_trans _ (Qabs (yn (Pos.to_nat k ^ 2 * q)%nat
- - yn (Pos.to_nat k ^ 2 * p)%nat)
- / (1 # (k^2)))).
- assert (1 # k ^ 2
- < Qabs (yn (Pos.to_nat k ^ 2 * q)%nat * yn (Pos.to_nat k ^ 2 * p)%nat))%Q.
- { rewrite Qabs_Qmult. unfold "^"%positive; simpl.
- rewrite factorDenom. rewrite Pos.mul_1_r.
- apply (Qlt_trans _ ((1#k) * Qabs (yn (Pos.to_nat k * (Pos.to_nat k * 1) * p)%nat))).
- apply Qmult_lt_l. reflexivity. rewrite Qabs_pos.
- specialize (maj (Pos.to_nat k * (Pos.to_nat k * 1) * p)%nat).
- apply maj. apply (Qle_trans _ (1 # k)).
- discriminate. apply Zlt_le_weak. apply maj.
- apply Qmult_lt_r. apply (Qlt_trans 0 (1#k)). reflexivity.
- rewrite Qabs_pos.
- specialize (maj (Pos.to_nat k * (Pos.to_nat k * 1) * p)%nat).
- apply maj. apply (Qle_trans _ (1 # k)). discriminate.
- apply Zlt_le_weak. apply maj.
- rewrite Qabs_pos.
- specialize (maj (Pos.to_nat k * (Pos.to_nat k * 1) * q)%nat).
- apply maj. apply (Qle_trans _ (1 # k)). discriminate.
- apply Zlt_le_weak. apply maj. }
- unfold Qdiv. rewrite Qabs_Qmult. rewrite Qabs_Qinv.
- rewrite Qmult_comm. rewrite <- (Qmult_comm (/ (1 # k ^ 2))).
- apply Qmult_le_compat_r. apply Qlt_le_weak.
- rewrite <- Qmult_1_l. apply Qlt_shift_div_r.
- apply (Qlt_trans 0 (1 # k ^ 2)). reflexivity. apply H.
- rewrite Qmult_comm. apply Qlt_shift_div_l.
- reflexivity. rewrite Qmult_1_l. apply H.
- apply Qabs_nonneg. simpl in maj.
- specialize (cau (n * (k^2))%positive
- (Pos.to_nat k ^ 2 * q)%nat
- (Pos.to_nat k ^ 2 * p)%nat).
- apply Qlt_shift_div_r. reflexivity.
- apply (Qlt_le_trans _ (1 # n * k ^ 2)). apply cau.
- rewrite Pos2Nat.inj_mul. rewrite mult_comm.
- unfold "^"%positive. simpl. rewrite Pos2Nat.inj_mul.
- rewrite <- mult_assoc. rewrite <- mult_assoc.
- apply Nat.mul_le_mono_nonneg_l. apply le_0_n.
- rewrite (mult_1_r). rewrite Pos.mul_1_r.
- apply Nat.mul_le_mono_nonneg_l. apply le_0_n.
- apply (le_trans _ (q+0)). rewrite plus_0_r. assumption.
- rewrite plus_0_r. apply le_refl.
- rewrite Pos2Nat.inj_mul. rewrite mult_comm.
- unfold "^"%positive; simpl. rewrite Pos2Nat.inj_mul.
- rewrite <- mult_assoc. rewrite <- mult_assoc.
- apply Nat.mul_le_mono_nonneg_l. apply le_0_n.
- rewrite (mult_1_r). rewrite Pos.mul_1_r.
- apply Nat.mul_le_mono_nonneg_l. apply le_0_n.
- apply (le_trans _ (p+0)). rewrite plus_0_r. assumption.
- rewrite plus_0_r. apply le_refl.
- rewrite factorDenom. apply Qle_refl.
- + field. split. intro abs.
- specialize (maj (Pos.to_nat k ^ 2 * p)%nat).
- rewrite abs in maj. inversion maj.
- intro abs.
- specialize (maj (Pos.to_nat k ^ 2 * q)%nat).
- rewrite abs in maj. inversion maj.
-Qed.
-
-Definition CReal_inv (x : CReal) (xnz : x # 0) : CReal.
-Proof.
- destruct xnz as [xNeg | xPos].
- - destruct (CRealNegShift x xNeg) as [[k y] [_ maj]].
- destruct y as [yn cau]; unfold proj1_sig, snd, fst in maj.
- exists (fun n => Qinv (yn (mult (Pos.to_nat k^2) n))).
- apply (CReal_inv_neg yn). apply cau. apply maj.
- - destruct (CRealPosShift x xPos) as [[k y] [_ maj]].
- destruct y as [yn cau]; unfold proj1_sig, snd, fst in maj.
- exists (fun n => Qinv (yn (mult (Pos.to_nat k^2) n))).
- apply (CReal_inv_pos yn). apply cau. apply maj.
-Defined.
-
-Notation "/ x" := (CReal_inv x) (at level 35, right associativity) : CReal_scope.
-
-Lemma CReal_inv_0_lt_compat
- : forall (r : CReal) (rnz : r # 0),
- 0 < r -> 0 < ((/ r) rnz).
-Proof.
- intros. unfold CReal_inv. simpl.
- destruct rnz.
- - exfalso. apply CRealLt_asym in H. contradiction.
- - destruct (CRealPosShift r c) as [[k rpos] [req maj]].
- clear req. destruct rpos as [rn cau]; simpl in maj.
- unfold CRealLt; simpl.
- destruct (Qarchimedean (rn 1%nat)) as [A majA].
- exists (2 * (A + 1))%positive. unfold Qminus. rewrite Qplus_0_r.
- rewrite <- (Qmult_1_l (Qinv (rn (Pos.to_nat k * (Pos.to_nat k * 1) * Pos.to_nat (2 * (A + 1)))%nat))).
- apply Qlt_shift_div_l. apply (Qlt_trans 0 (1#k)). reflexivity.
- apply maj. rewrite <- (Qmult_inv_r (Z.pos A + 1 # 1)).
- setoid_replace (2 # 2 * (A + 1))%Q with (Qinv (Z.pos A + 1 # 1)).
- 2: reflexivity.
- rewrite Qmult_comm. apply Qmult_lt_r. reflexivity.
- rewrite mult_1_r. rewrite <- Pos2Nat.inj_mul. rewrite <- Pos2Nat.inj_mul.
- rewrite <- (Qplus_lt_l _ _ (- rn 1%nat)).
- apply (Qle_lt_trans _ (Qabs (rn (Pos.to_nat (k * k * (2 * (A + 1)))) + - rn 1%nat))).
- apply Qle_Qabs. apply (Qlt_le_trans _ 1). apply cau.
- apply Pos2Nat.is_pos. apply le_refl.
- rewrite <- Qinv_plus_distr. rewrite <- (Qplus_comm 1).
- rewrite <- Qplus_0_r. rewrite <- Qplus_assoc. rewrite <- Qplus_assoc.
- rewrite Qplus_le_r. rewrite Qplus_0_l. apply Qlt_le_weak.
- apply Qlt_minus_iff in majA. apply majA.
- intro abs. inversion abs.
-Qed.
-
-Lemma CReal_linear_shift : forall (x : CReal) (k : nat),
- le 1 k -> QCauchySeq (fun n => proj1_sig x (k * n)%nat) Pos.to_nat.
-Proof.
- intros [xn limx] k lek p n m H H0. unfold proj1_sig.
- apply limx. apply (le_trans _ n). apply H.
- rewrite <- (mult_1_l n). rewrite mult_assoc.
- apply Nat.mul_le_mono_nonneg_r. apply le_0_n.
- rewrite mult_1_r. apply lek. apply (le_trans _ m). apply H0.
- rewrite <- (mult_1_l m). rewrite mult_assoc.
- apply Nat.mul_le_mono_nonneg_r. apply le_0_n.
- rewrite mult_1_r. apply lek.
-Qed.
-
-Lemma CReal_linear_shift_eq : forall (x : CReal) (k : nat) (kPos : le 1 k),
- CRealEq x
- (exist (fun n : nat -> Q => QCauchySeq n Pos.to_nat)
- (fun n : nat => proj1_sig x (k * n)%nat) (CReal_linear_shift x k kPos)).
-Proof.
- intros. apply CRealEq_diff. intro n.
- destruct x as [xn limx]; unfold proj1_sig.
- specialize (limx n (Pos.to_nat n) (k * Pos.to_nat n)%nat).
- apply (Qle_trans _ (1 # n)). apply Qlt_le_weak. apply limx.
- apply le_refl. rewrite <- (mult_1_l (Pos.to_nat n)).
- rewrite mult_assoc. apply Nat.mul_le_mono_nonneg_r. apply le_0_n.
- rewrite mult_1_r. apply kPos. apply Z.mul_le_mono_nonneg_r.
- discriminate. discriminate.
-Qed.
-
-Lemma CReal_inv_l : forall (r:CReal) (rnz : r # 0),
- ((/ r) rnz) * r == 1.
-Proof.
- intros. unfold CReal_inv; simpl.
- destruct rnz.
- - (* r < 0 *) destruct (CRealNegShift r c) as [[k rneg] [req maj]].
- simpl in req. apply CRealEq_diff. apply CRealEq_modindep.
- apply (QSeqEquivEx_trans _
- (proj1_sig (CReal_mult ((let
- (yn, cau) as s
- return ((forall n : nat, proj1_sig s n < -1 # k) -> CReal) := rneg in
- fun maj0 : forall n : nat, yn n < -1 # k =>
- exist (fun x : nat -> Q => QCauchySeq x Pos.to_nat)
- (fun n : nat => Qinv (yn (Pos.to_nat k * (Pos.to_nat k * 1) * n))%nat)
- (CReal_inv_neg yn k cau maj0)) maj) rneg)))%Q.
- + apply CRealEq_modindep. apply CRealEq_diff.
- apply CReal_mult_proper_l. apply req.
- + assert (le 1 (Pos.to_nat k * (Pos.to_nat k * 1))%nat). rewrite mult_1_r.
- rewrite <- Pos2Nat.inj_mul. apply Pos2Nat.is_pos.
- apply (QSeqEquivEx_trans _
- (proj1_sig (CReal_mult ((let
- (yn, cau) as s
- return ((forall n : nat, proj1_sig s n < -1 # k) -> CReal) := rneg in
- fun maj0 : forall n : nat, yn n < -1 # k =>
- exist (fun x : nat -> Q => QCauchySeq x Pos.to_nat)
- (fun n : nat => Qinv (yn (Pos.to_nat k * (Pos.to_nat k * 1) * n)%nat))
- (CReal_inv_neg yn k cau maj0)) maj)
- (exist _ (fun n => proj1_sig rneg (Pos.to_nat k * (Pos.to_nat k * 1) * n)%nat) (CReal_linear_shift rneg _ H)))))%Q.
- apply CRealEq_modindep. apply CRealEq_diff.
- apply CReal_mult_proper_l. apply CReal_linear_shift_eq.
- destruct r as [rn limr], rneg as [rnn limneg]; simpl.
- destruct (QCauchySeq_bounded
- (fun n : nat => Qinv (rnn (Pos.to_nat k * (Pos.to_nat k * 1) * n)%nat))
- Pos.to_nat (CReal_inv_neg rnn k limneg maj)).
- destruct (QCauchySeq_bounded
- (fun n : nat => rnn (Pos.to_nat k * (Pos.to_nat k * 1) * n)%nat)
- Pos.to_nat
- (CReal_linear_shift
- (exist (fun x0 : nat -> Q => QCauchySeq x0 Pos.to_nat) rnn limneg)
- (Pos.to_nat k * (Pos.to_nat k * 1)) H)) ; simpl.
- exists (fun n => 1%nat). intros p n m H0 H1. rewrite Qmult_comm.
- rewrite Qmult_inv_r. unfold Qminus. rewrite Qplus_opp_r.
- reflexivity. intro abs.
- specialize (maj (Pos.to_nat k * (Pos.to_nat k * 1)
- * (Pos.to_nat (Pos.max x x0)~0 * n))%nat).
- simpl in maj. rewrite abs in maj. inversion maj.
- - (* r > 0 *) destruct (CRealPosShift r c) as [[k rneg] [req maj]].
- simpl in req. apply CRealEq_diff. apply CRealEq_modindep.
- apply (QSeqEquivEx_trans _
- (proj1_sig (CReal_mult ((let
- (yn, cau) as s
- return ((forall n : nat, 1 # k < proj1_sig s n) -> CReal) := rneg in
- fun maj0 : forall n : nat, 1 # k < yn n =>
- exist (fun x : nat -> Q => QCauchySeq x Pos.to_nat)
- (fun n : nat => Qinv (yn (Pos.to_nat k * (Pos.to_nat k * 1) * n)%nat))
- (CReal_inv_pos yn k cau maj0)) maj) rneg)))%Q.
- + apply CRealEq_modindep. apply CRealEq_diff.
- apply CReal_mult_proper_l. apply req.
- + assert (le 1 (Pos.to_nat k * (Pos.to_nat k * 1))%nat). rewrite mult_1_r.
- rewrite <- Pos2Nat.inj_mul. apply Pos2Nat.is_pos.
- apply (QSeqEquivEx_trans _
- (proj1_sig (CReal_mult ((let
- (yn, cau) as s
- return ((forall n : nat, 1 # k < proj1_sig s n) -> CReal) := rneg in
- fun maj0 : forall n : nat, 1 # k < yn n =>
- exist (fun x : nat -> Q => QCauchySeq x Pos.to_nat)
- (fun n : nat => Qinv (yn (Pos.to_nat k * (Pos.to_nat k * 1) * n)%nat))
- (CReal_inv_pos yn k cau maj0)) maj)
- (exist _ (fun n => proj1_sig rneg (Pos.to_nat k * (Pos.to_nat k * 1) * n)%nat) (CReal_linear_shift rneg _ H)))))%Q.
- apply CRealEq_modindep. apply CRealEq_diff.
- apply CReal_mult_proper_l. apply CReal_linear_shift_eq.
- destruct r as [rn limr], rneg as [rnn limneg]; simpl.
- destruct (QCauchySeq_bounded
- (fun n : nat => Qinv (rnn (Pos.to_nat k * (Pos.to_nat k * 1) * n)%nat))
- Pos.to_nat (CReal_inv_pos rnn k limneg maj)).
- destruct (QCauchySeq_bounded
- (fun n : nat => rnn (Pos.to_nat k * (Pos.to_nat k * 1) * n)%nat)
- Pos.to_nat
- (CReal_linear_shift
- (exist (fun x0 : nat -> Q => QCauchySeq x0 Pos.to_nat) rnn limneg)
- (Pos.to_nat k * (Pos.to_nat k * 1)) H)) ; simpl.
- exists (fun n => 1%nat). intros p n m H0 H1. rewrite Qmult_comm.
- rewrite Qmult_inv_r. unfold Qminus. rewrite Qplus_opp_r.
- reflexivity. intro abs.
- specialize (maj (Pos.to_nat k * (Pos.to_nat k * 1)
- * (Pos.to_nat (Pos.max x x0)~0 * n))%nat).
- simpl in maj. rewrite abs in maj. inversion maj.
-Qed.
-
-Lemma CReal_inv_r : forall (r:CReal) (rnz : r # 0),
- r * ((/ r) rnz) == 1.
-Proof.
- intros. rewrite CReal_mult_comm, CReal_inv_l.
- reflexivity.
-Qed.
-
-Lemma CReal_inv_1 : forall nz : 1 # 0, (/ 1) nz == 1.
-Proof.
- intros. rewrite <- (CReal_mult_1_l ((/1) nz)). rewrite CReal_inv_r.
- reflexivity.
-Qed.
-
-Lemma CReal_inv_mult_distr :
- forall r1 r2 (r1nz : r1 # 0) (r2nz : r2 # 0) (rmnz : (r1*r2) # 0),
- (/ (r1 * r2)) rmnz == (/ r1) r1nz * (/ r2) r2nz.
-Proof.
- intros. apply (CReal_mult_eq_reg_l r1). exact r1nz.
- rewrite <- CReal_mult_assoc. rewrite CReal_inv_r. rewrite CReal_mult_1_l.
- apply (CReal_mult_eq_reg_l r2). exact r2nz.
- rewrite CReal_inv_r. rewrite <- CReal_mult_assoc.
- rewrite (CReal_mult_comm r2 r1). rewrite CReal_inv_r.
- reflexivity.
-Qed.
-
-Lemma Rinv_eq_compat : forall x y (rxnz : x # 0) (rynz : y # 0),
- x == y
- -> (/ x) rxnz == (/ y) rynz.
-Proof.
- intros. apply (CReal_mult_eq_reg_l x). exact rxnz.
- rewrite CReal_inv_r, H, CReal_inv_r. reflexivity.
-Qed.
-
-Lemma CReal_mult_lt_reg_l : forall r r1 r2, 0 < r -> r * r1 < r * r2 -> r1 < r2.
-Proof.
- intros z x y H H0.
- apply (CReal_mult_lt_compat_l ((/z) (inr H))) in H0.
- repeat rewrite <- CReal_mult_assoc in H0. rewrite CReal_inv_l in H0.
- repeat rewrite CReal_mult_1_l in H0. apply H0.
- apply CReal_inv_0_lt_compat. exact H.
-Qed.
-
-Lemma CReal_mult_lt_reg_r : forall r r1 r2, 0 < r -> r1 * r < r2 * r -> r1 < r2.
-Proof.
- intros.
- apply CReal_mult_lt_reg_l with r.
- exact H.
- now rewrite 2!(CReal_mult_comm r).
-Qed.
-
-Lemma CReal_mult_eq_reg_r : forall r r1 r2, r1 * r == r2 * r -> r # 0 -> r1 == r2.
-Proof.
- intros. apply (CReal_mult_eq_reg_l r). exact H0.
- now rewrite 2!(CReal_mult_comm r).
-Qed.
-
-Lemma CReal_mult_eq_compat_l : forall r r1 r2, r1 == r2 -> r * r1 == r * r2.
-Proof.
- intros. rewrite H. reflexivity.
-Qed.
-
-Lemma CReal_mult_eq_compat_r : forall r r1 r2, r1 == r2 -> r1 * r == r2 * r.
-Proof.
- intros. rewrite H. reflexivity.
-Qed.
-
-Fixpoint pow (r:CReal) (n:nat) : CReal :=
- match n with
- | O => 1
- | S n => r * (pow r n)
- end.
-
-
-(**********)
-Definition IQR (q:Q) : CReal :=
- match q with
- | Qmake a b => IZR a * (CReal_inv (IPR b)) (inr (IPR_pos b))
- end.
-Arguments IQR q%Q : simpl never.
-
-Lemma mult_IPR_IZR : forall (n:positive) (m:Z), IZR (Z.pos n * m) == IPR n * IZR m.
-Proof.
- intros. rewrite mult_IZR. apply CReal_mult_eq_compat_r. reflexivity.
-Qed.
-
-Lemma plus_IQR : forall n m:Q, IQR (n + m) == IQR n + IQR m.
-Proof.
- intros. destruct n,m; unfold Qplus,IQR; simpl.
- rewrite plus_IZR. repeat rewrite mult_IZR.
- setoid_replace ((/ IPR (Qden * Qden0)) (inr (IPR_pos (Qden * Qden0))))
- with ((/ IPR Qden) (inr (IPR_pos Qden))
- * (/ IPR Qden0) (inr (IPR_pos Qden0))).
- rewrite CReal_mult_plus_distr_r.
- repeat rewrite CReal_mult_assoc. rewrite <- (CReal_mult_assoc (IZR (Z.pos Qden))).
- rewrite CReal_inv_r, CReal_mult_1_l.
- rewrite (CReal_mult_comm ((/IPR Qden) (inr (IPR_pos Qden)))).
- rewrite <- (CReal_mult_assoc (IZR (Z.pos Qden0))).
- rewrite CReal_inv_r, CReal_mult_1_l. reflexivity. unfold IZR.
- rewrite <- (CReal_inv_mult_distr
- _ _ _ _ (inr (CReal_mult_lt_0_compat _ _ (IPR_pos _) (IPR_pos _)))).
- apply Rinv_eq_compat. apply mult_IPR.
-Qed.
-
-Lemma IQR_pos : forall q:Q, Qlt 0 q -> 0 < IQR q.
-Proof.
- intros. destruct q; unfold IQR.
- apply CReal_mult_lt_0_compat. apply (IZR_lt 0).
- unfold Qlt in H; simpl in H.
- rewrite Z.mul_1_r in H. apply H.
- apply CReal_inv_0_lt_compat. apply IPR_pos.
+ split.
+ - intros [n nmaj]. simpl in nmaj.
+ ring_simplify in nmaj. discriminate.
+ - intros [n nmaj]. simpl in nmaj.
+ ring_simplify in nmaj. discriminate.
Qed.
-Lemma opp_IQR : forall q:Q, IQR (- q) == - IQR q.
+Lemma inject_Q_one : inject_Q 1 == 1.
Proof.
- intros [a b]; unfold IQR; simpl.
- rewrite CReal_opp_mult_distr_l.
- rewrite opp_IZR. reflexivity.
+ split.
+ - intros [n nmaj]. simpl in nmaj.
+ ring_simplify in nmaj. discriminate.
+ - intros [n nmaj]. simpl in nmaj.
+ ring_simplify in nmaj. discriminate.
Qed.
-Lemma lt_IQR : forall n m:Q, IQR n < IQR m -> (n < m)%Q.
+Lemma inject_Q_lt : forall q r : Q,
+ Qlt q r -> inject_Q q < inject_Q r.
Proof.
- intros. destruct n,m; unfold IQR in H.
- unfold Qlt; simpl. apply (CReal_mult_lt_compat_r (IPR Qden)) in H.
- rewrite CReal_mult_assoc in H. rewrite CReal_inv_l in H.
- rewrite CReal_mult_1_r in H. rewrite (CReal_mult_comm (IZR Qnum0)) in H.
- apply (CReal_mult_lt_compat_l (IPR Qden0)) in H.
- do 2 rewrite <- CReal_mult_assoc in H. rewrite CReal_inv_r in H.
- rewrite CReal_mult_1_l in H.
- rewrite (CReal_mult_comm (IZR Qnum0)) in H.
- do 2 rewrite <- mult_IPR_IZR in H. apply lt_IZR in H.
- rewrite Z.mul_comm. rewrite (Z.mul_comm Qnum0).
- apply H. apply IPR_pos. apply IPR_pos.
+ intros. destruct (Qarchimedean (/(r-q))).
+ exists (2*x)%positive; simpl.
+ setoid_replace (2 # x~0)%Q with (/(Z.pos x#1))%Q. 2: reflexivity.
+ apply Qlt_shift_inv_r. reflexivity.
+ apply (Qmult_lt_l _ _ (r-q)) in q0. rewrite Qmult_inv_r in q0.
+ exact q0. intro abs. rewrite Qlt_minus_iff in H.
+ unfold Qminus in abs. rewrite abs in H. discriminate H.
+ unfold Qminus. rewrite <- Qlt_minus_iff. exact H.
Qed.
-Lemma CReal_mult_le_compat_l_half : forall r r1 r2,
- 0 < r -> r1 <= r2 -> r * r1 <= r * r2.
+Lemma opp_inject_Q : forall q : Q,
+ inject_Q (-q) == - inject_Q q.
Proof.
- intros. intro abs. apply (CReal_mult_lt_reg_l) in abs.
- contradiction. apply H.
+ split.
+ - intros [n maj]. simpl in maj. ring_simplify in maj. discriminate.
+ - intros [n maj]. simpl in maj. ring_simplify in maj. discriminate.
Qed.
-Lemma IQR_lt : forall n m:Q, Qlt n m -> IQR n < IQR m.
+Lemma lt_inject_Q : forall q r : Q,
+ inject_Q q < inject_Q r -> Qlt q r.
Proof.
- intros. apply (CReal_plus_lt_reg_r (-IQR n)).
- rewrite CReal_plus_opp_r. rewrite <- opp_IQR. rewrite <- plus_IQR.
- apply IQR_pos. apply (Qplus_lt_l _ _ n).
- ring_simplify. apply H.
+ intros. destruct H. simpl in q0.
+ apply Qlt_minus_iff, (Qlt_trans _ (2#x)).
+ reflexivity. exact q0.
Qed.
-Lemma IQR_nonneg : forall q:Q, Qle 0 q -> 0 <= (IQR q).
+Lemma le_inject_Q : forall q r : Q,
+ inject_Q q <= inject_Q r -> Qle q r.
Proof.
- intros [a b] H. unfold IQR.
- apply (CRealLe_trans _ ((/ IPR b) (inr (IPR_pos b)) * 0)).
- rewrite CReal_mult_0_r. apply CRealLe_refl.
- rewrite (CReal_mult_comm (IZR a)). apply CReal_mult_le_compat_l_half.
- apply CReal_inv_0_lt_compat. apply IPR_pos.
- apply (IZR_le 0 a). unfold Qle in H; simpl in H.
- rewrite Z.mul_1_r in H. apply H.
+ intros. destruct (Qlt_le_dec r q). 2: exact q0.
+ exfalso. apply H. clear H. apply inject_Q_lt. exact q0.
Qed.
-Lemma IQR_le : forall n m:Q, Qle n m -> IQR n <= IQR m.
+Lemma inject_Q_le : forall q r : Q,
+ Qle q r -> inject_Q q <= inject_Q r.
Proof.
- intros. intro abs. apply (CReal_plus_lt_compat_l (-IQR n)) in abs.
- rewrite CReal_plus_opp_l, <- opp_IQR, <- plus_IQR in abs.
- apply IQR_nonneg in abs. contradiction. apply (Qplus_le_l _ _ n).
- ring_simplify. apply H.
+ intros. intros [n maj]. simpl in maj.
+ apply (Qlt_not_le _ _ maj). apply (Qle_trans _ 0).
+ apply (Qplus_le_l _ _ r). ring_simplify. exact H. discriminate.
Qed.
-
-Add Parametric Morphism : IQR
- with signature Qeq ==> CRealEq
- as IQR_morph.
-Proof.
- intros. destruct x,y; unfold IQR.
- unfold Qeq in H; simpl in H.
- apply (CReal_mult_eq_reg_r (IZR (Z.pos Qden))).
- 2: right; apply IPR_pos.
- rewrite CReal_mult_assoc. rewrite CReal_inv_l. rewrite CReal_mult_1_r.
- rewrite (CReal_mult_comm (IZR Qnum0)).
- apply (CReal_mult_eq_reg_l (IZR (Z.pos Qden0))).
- right; apply IPR_pos.
- rewrite <- CReal_mult_assoc, <- CReal_mult_assoc, CReal_inv_r.
- rewrite CReal_mult_1_l.
- repeat rewrite <- mult_IZR.
- rewrite <- H. rewrite Zmult_comm. reflexivity.
-Qed.
-
-Instance IQR_morph_T
- : CMorphisms.Proper
- (CMorphisms.respectful Qeq CRealEq) IQR.
-Proof.
- intros x y H. destruct x,y; unfold IQR.
- unfold Qeq in H; simpl in H.
- apply (CReal_mult_eq_reg_r (IZR (Z.pos Qden))).
- 2: right; apply IPR_pos.
- rewrite CReal_mult_assoc. rewrite CReal_inv_l. rewrite CReal_mult_1_r.
- rewrite (CReal_mult_comm (IZR Qnum0)).
- apply (CReal_mult_eq_reg_l (IZR (Z.pos Qden0))).
- right; apply IPR_pos.
- rewrite <- CReal_mult_assoc, <- CReal_mult_assoc, CReal_inv_r.
- rewrite CReal_mult_1_l.
- repeat rewrite <- mult_IZR.
- rewrite <- H. rewrite Zmult_comm. reflexivity.
-Qed.
-
-Lemma CReal_invQ : forall (b : positive) (pos : Qlt 0 (Z.pos b # 1)),
- CRealEq (CReal_inv (inject_Q (Z.pos b # 1)) (inr (CReal_injectQPos (Z.pos b # 1) pos)))
- (inject_Q (1 # b)).
-Proof.
- intros.
- apply (CReal_mult_eq_reg_l (inject_Q (Z.pos b # 1))).
- - right. apply CReal_injectQPos. exact pos.
- - rewrite CReal_mult_comm, CReal_inv_l.
- apply CRealEq_diff. intro n. simpl;
- destruct (QCauchySeq_bounded (fun _ : nat => 1 # b)%Q Pos.to_nat (ConstCauchy (1 # b))),
- (QCauchySeq_bounded (fun _ : nat => Z.pos b # 1)%Q Pos.to_nat (ConstCauchy (Z.pos b # 1))); simpl.
- do 2 rewrite Pos.mul_1_r. rewrite Z.pos_sub_diag. discriminate.
-Qed.
-
-(* The constant sequences of rationals are CRealEq to
- the rational operations on the unity. *)
-Lemma FinjectQ_CReal : forall q : Q,
- IQR q == inject_Q q.
-Proof.
- intros [a b]. unfold IQR.
- pose proof (CReal_iterate_one (Pos.to_nat b)).
- rewrite positive_nat_Z in H. simpl in H.
- assert (0 < Z.pos b # 1)%Q as pos. reflexivity.
- apply (CRealEq_trans _ (CReal_mult (IZR a)
- (CReal_inv (inject_Q (Z.pos b # 1)) (inr (CReal_injectQPos (Z.pos b # 1) pos))))).
- - apply CReal_mult_proper_l.
- apply (CReal_mult_eq_reg_l (IPR b)).
- right. apply IPR_pos.
- rewrite CReal_mult_comm, CReal_inv_l, H, CReal_mult_comm, CReal_inv_l. reflexivity.
- - rewrite FinjectZ_CReal. rewrite CReal_invQ. apply CRealEq_diff. intro n.
- simpl;
- destruct (QCauchySeq_bounded (fun _ : nat => a # 1)%Q Pos.to_nat (ConstCauchy (a # 1))),
- (QCauchySeq_bounded (fun _ : nat => 1 # b)%Q Pos.to_nat (ConstCauchy (1 # b))); simpl.
- rewrite Z.mul_1_r. rewrite <- Z.mul_add_distr_r.
- rewrite Z.add_opp_diag_r. rewrite Z.mul_0_l. simpl.
- discriminate.
-Qed.
-
-Lemma CReal_gen_inject : forall (n : nat),
- gen_phiZ (inject_Q 0) (inject_Q 1) CReal_plus CReal_mult CReal_opp
- (Z.of_nat n)
- == inject_Q (Z.of_nat n # 1).
-Proof.
- induction n.
- - apply CRealEq_refl.
- - replace (Z.of_nat (S n)) with (1 + Z.of_nat n)%Z.
- rewrite (gen_phiZ_add CRealEq_rel CReal_isRingExt CReal_isRing).
- rewrite IHn. clear IHn. apply CRealEq_diff. intro k. simpl.
- rewrite Z.mul_1_r. rewrite Z.mul_1_r. rewrite Z.mul_1_r.
- rewrite Z.add_opp_diag_r. discriminate.
- replace (S n) with (1 + n)%nat. 2: reflexivity.
- rewrite (Nat2Z.inj_add 1 n). reflexivity.
-Qed.
-
-Lemma CRealArchimedean
- : forall x:CReal, { n:Z & CRealLt x (gen_phiZ (inject_Q 0) (inject_Q 1) CReal_plus
- CReal_mult CReal_opp n) }.
-Proof.
- intros [xn limx]. destruct (Qarchimedean (xn 1%nat)) as [k kmaj].
- exists (Z.pos (2 + k)). rewrite <- (positive_nat_Z (2 + k)).
- rewrite CReal_gen_inject. rewrite (positive_nat_Z (2 + k)).
- exists xH.
- setoid_replace (2 # 1)%Q with
- ((Z.pos (2 + k) # 1) - (Z.pos k # 1))%Q.
- - apply Qplus_lt_r. apply Qlt_minus_iff. rewrite Qopp_involutive.
- apply Qlt_minus_iff in kmaj. rewrite Qplus_comm. apply kmaj.
- - unfold Qminus. setoid_replace (- (Z.pos k # 1))%Q with (-Z.pos k # 1)%Q.
- 2: reflexivity. rewrite Qinv_plus_distr.
- rewrite Pos2Z.inj_add. rewrite <- Zplus_assoc.
- rewrite Zplus_opp_r. reflexivity.
-Qed.
-
-
-Close Scope CReal_scope.
-
-Close Scope Q.
diff --git a/theories/Reals/ConstructiveCauchyRealsMult.v b/theories/Reals/ConstructiveCauchyRealsMult.v
new file mode 100644
index 0000000000..d6d4e84560
--- /dev/null
+++ b/theories/Reals/ConstructiveCauchyRealsMult.v
@@ -0,0 +1,1415 @@
+(************************************************************************)
+(* * The Coq Proof Assistant / The Coq Development Team *)
+(* v * INRIA, CNRS and contributors - Copyright 1999-2019 *)
+(* <O___,, * (see CREDITS file for the list of authors) *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(* * (see LICENSE file for the text of the license) *)
+(************************************************************************)
+(************************************************************************)
+
+(* The multiplication and division of Cauchy reals. *)
+
+Require Import QArith.
+Require Import Qabs.
+Require Import Qround.
+Require Import Logic.ConstructiveEpsilon.
+Require Export Reals.ConstructiveCauchyReals.
+Require CMorphisms.
+
+Local Open Scope CReal_scope.
+
+Fixpoint BoundFromZero (qn : nat -> Q) (k : nat) (A : positive) { struct k }
+ : (forall n:nat, le k n -> Qlt (Qabs (qn n)) (Z.pos A # 1))
+ -> { B : positive | forall n:nat, Qlt (Qabs (qn n)) (Z.pos B # 1) }.
+Proof.
+ intro H. destruct k.
+ - exists A. intros. apply H. apply le_0_n.
+ - destruct (Qarchimedean (Qabs (qn k))) as [a maj].
+ apply (BoundFromZero qn k (Pos.max A a)).
+ intros n H0. destruct (Nat.le_gt_cases n k).
+ + pose proof (Nat.le_antisymm n k H1 H0). subst k.
+ apply (Qlt_le_trans _ (Z.pos a # 1)). apply maj.
+ unfold Qle; simpl. rewrite Pos.mul_1_r. rewrite Pos.mul_1_r.
+ apply Pos.le_max_r.
+ + apply (Qlt_le_trans _ (Z.pos A # 1)). apply H.
+ apply H1. unfold Qle; simpl. rewrite Pos.mul_1_r. rewrite Pos.mul_1_r.
+ apply Pos.le_max_l.
+Qed.
+
+Lemma QCauchySeq_bounded (qn : nat -> Q) (cvmod : positive -> nat)
+ : QCauchySeq qn cvmod
+ -> { A : positive | forall n:nat, Qlt (Qabs (qn n)) (Z.pos A # 1) }.
+Proof.
+ intros. remember (Zplus (Qnum (Qabs (qn (cvmod xH)))) 1) as z.
+ assert (Z.lt 0 z) as zPos.
+ { subst z. assert (Qle 0 (Qabs (qn (cvmod 1%positive)))).
+ apply Qabs_nonneg. destruct (Qabs (qn (cvmod 1%positive))). simpl.
+ unfold Qle in H0. simpl in H0. rewrite Zmult_1_r in H0.
+ apply (Z.lt_le_trans 0 1). unfold Z.lt. auto.
+ rewrite <- (Zplus_0_l 1). rewrite Zplus_assoc. apply Zplus_le_compat_r.
+ rewrite Zplus_0_r. assumption. }
+ assert { A : positive | forall n:nat,
+ le (cvmod xH) n -> Qlt ((Qabs (qn n)) * (1#A)) 1 }.
+ destruct z eqn:des.
+ - exfalso. apply (Z.lt_irrefl 0). assumption.
+ - exists p. intros. specialize (H xH (cvmod xH) n (le_refl _) H0).
+ assert (Qlt (Qabs (qn n)) (Qabs (qn (cvmod 1%positive)) + 1)).
+ { apply (Qplus_lt_l _ _ (-Qabs (qn (cvmod 1%positive)))).
+ rewrite <- (Qplus_comm 1). rewrite <- Qplus_assoc. rewrite Qplus_opp_r.
+ rewrite Qplus_0_r. apply (Qle_lt_trans _ (Qabs (qn n - qn (cvmod 1%positive)))).
+ apply Qabs_triangle_reverse. rewrite Qabs_Qminus. assumption. }
+ apply (Qlt_le_trans _ ((Qabs (qn (cvmod 1%positive)) + 1) * (1#p))).
+ apply Qmult_lt_r. unfold Qlt. simpl. unfold Z.lt. auto. assumption.
+ unfold Qle. simpl. rewrite Zmult_1_r. rewrite Zmult_1_r. rewrite Zmult_1_r.
+ rewrite Pos.mul_1_r. rewrite Pos2Z.inj_mul. rewrite Heqz.
+ destruct (Qabs (qn (cvmod 1%positive))) eqn:desAbs.
+ rewrite Z.mul_add_distr_l. rewrite Zmult_1_r.
+ apply Zplus_le_compat_r. rewrite <- (Zmult_1_l (QArith_base.Qnum (Qnum # Qden))).
+ rewrite Zmult_assoc. apply Zmult_le_compat_r. rewrite Zmult_1_r.
+ simpl. unfold Z.le. rewrite <- Pos2Z.inj_compare.
+ unfold Pos.compare. destruct Qden; discriminate.
+ simpl. assert (Qle 0 (Qnum # Qden)). rewrite <- desAbs.
+ apply Qabs_nonneg. unfold Qle in H2. simpl in H2. rewrite Zmult_1_r in H2.
+ assumption.
+ - exfalso. inversion zPos.
+ - destruct H0. apply (BoundFromZero _ (cvmod xH) x). intros n H0.
+ specialize (q n H0). setoid_replace (Z.pos x # 1)%Q with (/(1#x))%Q.
+ rewrite <- (Qmult_1_l (/(1#x))). apply Qlt_shift_div_l.
+ reflexivity. apply q. reflexivity.
+Qed.
+
+Lemma CReal_mult_cauchy
+ : forall (xn yn zn : nat -> Q) (Ay Az : positive) (cvmod : positive -> nat),
+ QSeqEquiv xn yn cvmod
+ -> QCauchySeq zn Pos.to_nat
+ -> (forall n:nat, Qlt (Qabs (yn n)) (Z.pos Ay # 1))
+ -> (forall n:nat, Qlt (Qabs (zn n)) (Z.pos Az # 1))
+ -> QSeqEquiv (fun n:nat => xn n * zn n) (fun n:nat => yn n * zn n)
+ (fun p => max (cvmod (2 * (Pos.max Ay Az) * p)%positive)
+ (Pos.to_nat (2 * (Pos.max Ay Az) * p)%positive)).
+Proof.
+ intros xn yn zn Ay Az cvmod limx limz majy majz.
+ remember (Pos.mul 2 (Pos.max Ay Az)) as z.
+ intros k p q H H0.
+ assert (Pos.to_nat k <> O) as kPos.
+ { intro absurd. pose proof (Pos2Nat.is_pos k).
+ rewrite absurd in H1. inversion H1. }
+ setoid_replace (xn p * zn p - yn q * zn q)%Q
+ with ((xn p - yn q) * zn p + yn q * (zn p - zn q))%Q.
+ 2: ring.
+ apply (Qle_lt_trans _ (Qabs ((xn p - yn q) * zn p)
+ + Qabs (yn q * (zn p - zn q)))).
+ apply Qabs_triangle. rewrite Qabs_Qmult. rewrite Qabs_Qmult.
+ setoid_replace (1#k)%Q with ((1#2*k) + (1#2*k))%Q.
+ apply Qplus_lt_le_compat.
+ - apply (Qle_lt_trans _ ((1#z * k) * Qabs (zn p)%nat)).
+ + apply Qmult_le_compat_r. apply Qle_lteq. left. apply limx.
+ apply (le_trans _ (Init.Nat.max (cvmod (z * k)%positive) (Pos.to_nat (z * k)))).
+ apply Nat.le_max_l. assumption.
+ apply (le_trans _ (Init.Nat.max (cvmod (z * k)%positive) (Pos.to_nat (z * k)))).
+ apply Nat.le_max_l. assumption. apply Qabs_nonneg.
+ + subst z. rewrite <- (Qmult_1_r (1 # 2 * k)).
+ rewrite <- Pos.mul_assoc. rewrite <- (Pos.mul_comm k). rewrite Pos.mul_assoc.
+ rewrite (factorDenom _ _ (2 * k)). rewrite <- Qmult_assoc.
+ apply Qmult_lt_l. unfold Qlt. simpl. unfold Z.lt. auto.
+ apply (Qle_lt_trans _ (Qabs (zn p)%nat * (1 # Az))).
+ rewrite <- (Qmult_comm (1 # Az)). apply Qmult_le_compat_r.
+ unfold Qle. simpl. rewrite Pos2Z.inj_max. apply Z.le_max_r.
+ apply Qabs_nonneg. rewrite <- (Qmult_inv_r (1#Az)).
+ rewrite Qmult_comm. apply Qmult_lt_l. reflexivity.
+ setoid_replace (/(1#Az))%Q with (Z.pos Az # 1)%Q. apply majz.
+ reflexivity. intro abs. inversion abs.
+ - apply (Qle_trans _ ((1 # z * k) * Qabs (yn q)%nat)).
+ + rewrite Qmult_comm. apply Qmult_le_compat_r. apply Qle_lteq.
+ left. apply limz.
+ apply (le_trans _ (max (cvmod (z * k)%positive)
+ (Pos.to_nat (z * k)%positive))).
+ apply Nat.le_max_r. assumption.
+ apply (le_trans _ (max (cvmod (z * k)%positive)
+ (Pos.to_nat (z * k)%positive))).
+ apply Nat.le_max_r. assumption. apply Qabs_nonneg.
+ + subst z. rewrite <- (Qmult_1_r (1 # 2 * k)).
+ rewrite <- Pos.mul_assoc. rewrite <- (Pos.mul_comm k). rewrite Pos.mul_assoc.
+ rewrite (factorDenom _ _ (2 * k)). rewrite <- Qmult_assoc.
+ apply Qle_lteq. left.
+ apply Qmult_lt_l. unfold Qlt. simpl. unfold Z.lt. auto.
+ apply (Qle_lt_trans _ (Qabs (yn q)%nat * (1 # Ay))).
+ rewrite <- (Qmult_comm (1 # Ay)). apply Qmult_le_compat_r.
+ unfold Qle. simpl. rewrite Pos2Z.inj_max. apply Z.le_max_l.
+ apply Qabs_nonneg. rewrite <- (Qmult_inv_r (1#Ay)).
+ rewrite Qmult_comm. apply Qmult_lt_l. reflexivity.
+ setoid_replace (/(1#Ay))%Q with (Z.pos Ay # 1)%Q. apply majy.
+ reflexivity. intro abs. inversion abs.
+ - rewrite Qinv_plus_distr. unfold Qeq. reflexivity.
+Qed.
+
+Lemma linear_max : forall (p Ax Ay : positive) (i : nat),
+ le (Pos.to_nat p) i
+ -> (Init.Nat.max (Pos.to_nat (2 * Pos.max Ax Ay * p))
+ (Pos.to_nat (2 * Pos.max Ax Ay * p)) <= Pos.to_nat (2 * Pos.max Ax Ay) * i)%nat.
+Proof.
+ intros. rewrite max_l. 2: apply le_refl.
+ rewrite Pos2Nat.inj_mul. apply Nat.mul_le_mono_nonneg.
+ apply le_0_n. apply le_refl. apply le_0_n. apply H.
+Qed.
+
+Definition CReal_mult (x y : CReal) : CReal.
+Proof.
+ destruct x as [xn limx]. destruct y as [yn limy].
+ destruct (QCauchySeq_bounded xn Pos.to_nat limx) as [Ax majx].
+ destruct (QCauchySeq_bounded yn Pos.to_nat limy) as [Ay majy].
+ pose proof (CReal_mult_cauchy xn xn yn Ax Ay Pos.to_nat limx limy majx majy).
+ exists (fun n : nat => xn (Pos.to_nat (2 * Pos.max Ax Ay)* n)%nat
+ * yn (Pos.to_nat (2 * Pos.max Ax Ay) * n)%nat).
+ intros p n k H0 H1.
+ apply H; apply linear_max; assumption.
+Defined.
+
+Infix "*" := CReal_mult : CReal_scope.
+
+Lemma CReal_mult_unfold : forall x y : CReal,
+ QSeqEquivEx (proj1_sig (CReal_mult x y))
+ (fun n : nat => proj1_sig x n * proj1_sig y n)%Q.
+Proof.
+ intros [xn limx] [yn limy]. unfold CReal_mult ; simpl.
+ destruct (QCauchySeq_bounded xn Pos.to_nat limx) as [Ax majx].
+ destruct (QCauchySeq_bounded yn Pos.to_nat limy) as [Ay majy].
+ simpl.
+ pose proof (CReal_mult_cauchy xn xn yn Ax Ay Pos.to_nat limx limy majx majy).
+ exists (fun p : positive =>
+ Init.Nat.max (Pos.to_nat (2 * Pos.max Ax Ay * p))
+ (Pos.to_nat (2 * Pos.max Ax Ay * p))).
+ intros p n k H0 H1. rewrite max_l in H0, H1.
+ 2: apply le_refl. 2: apply le_refl.
+ apply H. apply linear_max.
+ apply (le_trans _ (Pos.to_nat (2 * Pos.max Ax Ay * p))).
+ rewrite <- (mult_1_l (Pos.to_nat p)). rewrite Pos2Nat.inj_mul.
+ apply Nat.mul_le_mono_nonneg. auto. apply Pos2Nat.is_pos.
+ apply le_0_n. apply le_refl. apply H0. rewrite max_l.
+ apply H1. apply le_refl.
+Qed.
+
+Lemma CReal_mult_assoc_bounded_r : forall (xn yn zn : nat -> Q),
+ QSeqEquivEx xn yn (* both are Cauchy with same limit *)
+ -> QSeqEquiv zn zn Pos.to_nat
+ -> QSeqEquivEx (fun n => xn n * zn n)%Q (fun n => yn n * zn n)%Q.
+Proof.
+ intros. destruct H as [cvmod cveq].
+ destruct (QCauchySeq_bounded yn (fun k => cvmod (2 * k)%positive)
+ (QSeqEquiv_cau_r xn yn cvmod cveq))
+ as [Ay majy].
+ destruct (QCauchySeq_bounded zn Pos.to_nat H0) as [Az majz].
+ exists (fun p => max (cvmod (2 * (Pos.max Ay Az) * p)%positive)
+ (Pos.to_nat (2 * (Pos.max Ay Az) * p)%positive)).
+ apply CReal_mult_cauchy; assumption.
+Qed.
+
+Lemma CReal_mult_assoc : forall x y z : CReal,
+ CRealEq (CReal_mult (CReal_mult x y) z)
+ (CReal_mult x (CReal_mult y z)).
+Proof.
+ intros. apply CRealEq_diff. apply CRealEq_modindep.
+ apply (QSeqEquivEx_trans _ (fun n => proj1_sig x n * proj1_sig y n * proj1_sig z n)%Q).
+ - apply (QSeqEquivEx_trans _ (fun n => proj1_sig (CReal_mult x y) n * proj1_sig z n)%Q).
+ apply CReal_mult_unfold.
+ destruct x as [xn limx], y as [yn limy], z as [zn limz]; unfold CReal_mult; simpl.
+ destruct (QCauchySeq_bounded xn Pos.to_nat limx) as [Ax majx].
+ destruct (QCauchySeq_bounded yn Pos.to_nat limy) as [Ay majy].
+ destruct (QCauchySeq_bounded zn Pos.to_nat limz) as [Az majz].
+ apply CReal_mult_assoc_bounded_r. 2: apply limz.
+ simpl.
+ pose proof (CReal_mult_cauchy xn xn yn Ax Ay Pos.to_nat limx limy majx majy).
+ exists (fun p : positive =>
+ Init.Nat.max (Pos.to_nat (2 * Pos.max Ax Ay * p))
+ (Pos.to_nat (2 * Pos.max Ax Ay * p))).
+ intros p n k H0 H1. rewrite max_l in H0, H1.
+ 2: apply le_refl. 2: apply le_refl.
+ apply H. apply linear_max.
+ apply (le_trans _ (Pos.to_nat (2 * Pos.max Ax Ay * p))).
+ rewrite <- (mult_1_l (Pos.to_nat p)). rewrite Pos2Nat.inj_mul.
+ apply Nat.mul_le_mono_nonneg. auto. apply Pos2Nat.is_pos.
+ apply le_0_n. apply le_refl. apply H0. rewrite max_l.
+ apply H1. apply le_refl.
+ - apply (QSeqEquivEx_trans _ (fun n => proj1_sig x n * proj1_sig (CReal_mult y z) n)%Q).
+ 2: apply QSeqEquivEx_sym; apply CReal_mult_unfold.
+ destruct x as [xn limx], y as [yn limy], z as [zn limz]; unfold CReal_mult; simpl.
+ destruct (QCauchySeq_bounded xn Pos.to_nat limx) as [Ax majx].
+ destruct (QCauchySeq_bounded yn Pos.to_nat limy) as [Ay majy].
+ destruct (QCauchySeq_bounded zn Pos.to_nat limz) as [Az majz].
+ simpl.
+ pose proof (CReal_mult_assoc_bounded_r (fun n0 : nat => yn n0 * zn n0)%Q (fun n : nat =>
+ yn (Pos.to_nat (Pos.max Ay Az)~0 * n)%nat
+ * zn (Pos.to_nat (Pos.max Ay Az)~0 * n)%nat)%Q xn)
+ as [cvmod cveq].
+
+ pose proof (CReal_mult_cauchy yn yn zn Ay Az Pos.to_nat limy limz majy majz).
+ exists (fun p : positive =>
+ Init.Nat.max (Pos.to_nat (2 * Pos.max Ay Az * p))
+ (Pos.to_nat (2 * Pos.max Ay Az * p))).
+ intros p n k H0 H1. rewrite max_l in H0, H1.
+ 2: apply le_refl. 2: apply le_refl.
+ apply H. rewrite max_l. apply H0. apply le_refl.
+ apply linear_max.
+ apply (le_trans _ (Pos.to_nat (2 * Pos.max Ay Az * p))).
+ rewrite <- (mult_1_l (Pos.to_nat p)). rewrite Pos2Nat.inj_mul.
+ apply Nat.mul_le_mono_nonneg. auto. apply Pos2Nat.is_pos.
+ apply le_0_n. apply le_refl. apply H1.
+ apply limx.
+ exists cvmod. intros p k n H1 H2. specialize (cveq p k n H1 H2).
+ setoid_replace (xn k * yn k * zn k -
+ xn n *
+ (yn (Pos.to_nat (Pos.max Ay Az)~0 * n)%nat *
+ zn (Pos.to_nat (Pos.max Ay Az)~0 * n)%nat))%Q
+ with ((fun n : nat => yn n * zn n * xn n) k -
+ (fun n : nat =>
+ yn (Pos.to_nat (Pos.max Ay Az)~0 * n)%nat *
+ zn (Pos.to_nat (Pos.max Ay Az)~0 * n)%nat *
+ xn n) n)%Q.
+ apply cveq. ring.
+Qed.
+
+Lemma CReal_mult_comm : forall x y : CReal,
+ CRealEq (CReal_mult x y) (CReal_mult y x).
+Proof.
+ intros. apply CRealEq_diff. apply CRealEq_modindep.
+ apply (QSeqEquivEx_trans _ (fun n => proj1_sig y n * proj1_sig x n)%Q).
+ destruct x as [xn limx], y as [yn limy]; simpl.
+ 2: apply QSeqEquivEx_sym; apply CReal_mult_unfold.
+ destruct (QCauchySeq_bounded xn Pos.to_nat limx) as [Ax majx].
+ destruct (QCauchySeq_bounded yn Pos.to_nat limy) as [Ay majy]; simpl.
+ apply QSeqEquivEx_sym.
+
+ pose proof (CReal_mult_cauchy yn yn xn Ay Ax Pos.to_nat limy limx majy majx).
+ exists (fun p : positive =>
+ Init.Nat.max (Pos.to_nat (2 * Pos.max Ay Ax * p))
+ (Pos.to_nat (2 * Pos.max Ay Ax * p))).
+ intros p n k H0 H1. rewrite max_l in H0, H1.
+ 2: apply le_refl. 2: apply le_refl.
+ rewrite (Qmult_comm (xn (Pos.to_nat (Pos.max Ax Ay)~0 * k)%nat)).
+ apply (H p n). rewrite max_l. apply H0. apply le_refl.
+ rewrite max_l. apply (le_trans _ k). apply H1.
+ rewrite <- (mult_1_l k). rewrite mult_assoc.
+ apply Nat.mul_le_mono_nonneg. auto. rewrite mult_1_r.
+ apply Pos2Nat.is_pos. apply le_0_n. apply le_refl.
+ apply le_refl.
+Qed.
+
+Lemma CReal_mult_proper_l : forall x y z : CReal,
+ CRealEq y z -> CRealEq (CReal_mult x y) (CReal_mult x z).
+Proof.
+ intros. apply CRealEq_diff. apply CRealEq_modindep.
+ apply (QSeqEquivEx_trans _ (fun n => proj1_sig x n * proj1_sig y n)%Q).
+ apply CReal_mult_unfold.
+ rewrite CRealEq_diff in H. rewrite <- CRealEq_modindep in H.
+ apply QSeqEquivEx_sym.
+ apply (QSeqEquivEx_trans _ (fun n => proj1_sig x n * proj1_sig z n)%Q).
+ apply CReal_mult_unfold.
+ destruct x as [xn limx], y as [yn limy], z as [zn limz]; simpl.
+ destruct H. simpl in H.
+ destruct (QCauchySeq_bounded xn Pos.to_nat limx) as [Ax majx].
+ destruct (QCauchySeq_bounded zn Pos.to_nat limz) as [Az majz].
+ pose proof (CReal_mult_cauchy yn zn xn Az Ax x H limx majz majx).
+ apply QSeqEquivEx_sym.
+ exists (fun p : positive =>
+ Init.Nat.max (x (2 * Pos.max Az Ax * p)%positive)
+ (Pos.to_nat (2 * Pos.max Az Ax * p))).
+ intros p n k H1 H2. specialize (H0 p n k H1 H2).
+ setoid_replace (xn n * yn n - xn k * zn k)%Q
+ with (yn n * xn n - zn k * xn k)%Q.
+ apply H0. ring.
+Qed.
+
+Lemma CReal_mult_lt_0_compat : forall x y : CReal,
+ CRealLt (inject_Q 0) x
+ -> CRealLt (inject_Q 0) y
+ -> CRealLt (inject_Q 0) (CReal_mult x y).
+Proof.
+ intros. destruct H as [x0 H], H0 as [x1 H0].
+ pose proof (CRealLt_aboveSig (inject_Q 0) x x0 H).
+ pose proof (CRealLt_aboveSig (inject_Q 0) y x1 H0).
+ destruct x as [xn limx], y as [yn limy].
+ simpl in H, H1, H2. simpl.
+ destruct (QCauchySeq_bounded xn Pos.to_nat limx) as [Ax majx].
+ destruct (QCauchySeq_bounded yn Pos.to_nat limy) as [Ay majy].
+ destruct (Qarchimedean (/ (xn (Pos.to_nat x0) - 0 - (2 # x0)))).
+ destruct (Qarchimedean (/ (yn (Pos.to_nat x1) - 0 - (2 # x1)))).
+ exists (Pos.max x0 x~0 * Pos.max x1 x2~0)%positive.
+ simpl. unfold Qminus. rewrite Qplus_0_r.
+ rewrite <- Pos2Nat.inj_mul.
+ unfold Qminus in H1, H2.
+ specialize (H1 ((Pos.max Ax Ay)~0 * (Pos.max x0 x~0 * Pos.max x1 x2~0))%positive).
+ assert (Pos.max x1 x2~0 <= (Pos.max Ax Ay)~0 * (Pos.max x0 x~0 * Pos.max x1 x2~0))%positive.
+ { apply Pos2Nat.inj_le.
+ rewrite Pos.mul_assoc. rewrite Pos2Nat.inj_mul.
+ rewrite <- (mult_1_l (Pos.to_nat (Pos.max x1 x2~0))).
+ rewrite mult_assoc. apply Nat.mul_le_mono_nonneg. auto.
+ rewrite mult_1_r. apply Pos2Nat.is_pos. apply le_0_n.
+ apply le_refl. }
+ specialize (H2 ((Pos.max Ax Ay)~0 * (Pos.max x0 x~0 * Pos.max x1 x2~0))%positive H3).
+ rewrite Qplus_0_r in H1, H2.
+ apply (Qlt_trans _ ((2 # Pos.max x0 x~0) * (2 # Pos.max x1 x2~0))).
+ unfold Qlt; simpl. assert (forall p : positive, (Z.pos p < Z.pos p~0)%Z).
+ intro p. rewrite <- (Z.mul_1_l (Z.pos p)).
+ replace (Z.pos p~0) with (2 * Z.pos p)%Z. apply Z.mul_lt_mono_pos_r.
+ apply Pos2Z.is_pos. reflexivity. reflexivity.
+ apply H4.
+ apply (Qlt_trans _ ((2 # Pos.max x0 x~0) * (yn (Pos.to_nat ((Pos.max Ax Ay)~0 * (Pos.max x0 x~0 * Pos.max x1 x2~0)))))).
+ apply Qmult_lt_l. reflexivity. apply H2. apply Qmult_lt_r.
+ apply (Qlt_trans 0 (2 # Pos.max x1 x2~0)). reflexivity. apply H2.
+ apply H1. rewrite Pos.mul_comm. apply Pos2Nat.inj_le.
+ rewrite <- Pos.mul_assoc. rewrite Pos2Nat.inj_mul.
+ rewrite <- (mult_1_r (Pos.to_nat (Pos.max x0 x~0))).
+ rewrite <- mult_assoc. apply Nat.mul_le_mono_nonneg.
+ apply le_0_n. apply le_refl. auto.
+ rewrite mult_1_l. apply Pos2Nat.is_pos.
+Qed.
+
+Lemma CReal_mult_plus_distr_l : forall r1 r2 r3 : CReal,
+ r1 * (r2 + r3) == (r1 * r2) + (r1 * r3).
+Proof.
+ intros x y z. apply CRealEq_diff. apply CRealEq_modindep.
+ apply (QSeqEquivEx_trans _ (fun n => proj1_sig x n
+ * (proj1_sig (CReal_plus y z) n))%Q).
+ apply CReal_mult_unfold.
+ apply (QSeqEquivEx_trans _ (fun n => proj1_sig (CReal_mult x y) n
+ + proj1_sig (CReal_mult x z) n))%Q.
+ 2: apply QSeqEquivEx_sym; exists (fun p => Pos.to_nat (2 * p))
+ ; apply CReal_plus_unfold.
+ apply (QSeqEquivEx_trans _ (fun n => proj1_sig x n
+ * (proj1_sig y n + proj1_sig z n))%Q).
+ - pose proof (CReal_plus_unfold y z).
+ destruct x as [xn limx], y as [yn limy], z as [zn limz]; simpl; simpl in H.
+ destruct (QCauchySeq_bounded xn Pos.to_nat limx) as [Ax majx].
+ destruct (QCauchySeq_bounded yn Pos.to_nat limy) as [Ay majy].
+ destruct (QCauchySeq_bounded zn Pos.to_nat limz) as [Az majz].
+ pose proof (CReal_mult_cauchy (fun n => yn (n + (n + 0))%nat + zn (n + (n + 0))%nat)%Q
+ (fun n => yn n + zn n)%Q
+ xn (Ay + Az) Ax
+ (fun p => Pos.to_nat (2 * p)) H limx).
+ exists (fun p : positive => (Pos.to_nat (2 * (2 * Pos.max (Ay + Az) Ax * p)))).
+ intros p n k H1 H2.
+ setoid_replace (xn n * (yn (n + (n + 0))%nat + zn (n + (n + 0))%nat) - xn k * (yn k + zn k))%Q
+ with ((yn (n + (n + 0))%nat + zn (n + (n + 0))%nat) * xn n - (yn k + zn k) * xn k)%Q.
+ 2: ring.
+ assert (Pos.to_nat (2 * Pos.max (Ay + Az) Ax * p) <=
+ Pos.to_nat 2 * Pos.to_nat (2 * Pos.max (Ay + Az) Ax * p))%nat.
+ { rewrite (Pos2Nat.inj_mul 2).
+ rewrite <- (mult_1_l (Pos.to_nat (2 * Pos.max (Ay + Az) Ax * p))).
+ rewrite mult_assoc. apply Nat.mul_le_mono_nonneg. auto.
+ simpl. auto. apply le_0_n. apply le_refl. }
+ apply H0. intro n0. apply (Qle_lt_trans _ (Qabs (yn n0) + Qabs (zn n0))).
+ apply Qabs_triangle. rewrite Pos2Z.inj_add.
+ rewrite <- Qinv_plus_distr. apply Qplus_lt_le_compat.
+ apply majy. apply Qlt_le_weak. apply majz.
+ apply majx. rewrite max_l.
+ apply H1. rewrite (Pos2Nat.inj_mul 2). apply H3.
+ rewrite max_l. apply H2. rewrite (Pos2Nat.inj_mul 2).
+ apply H3.
+ - destruct x as [xn limx], y as [yn limy], z as [zn limz]; simpl.
+ destruct (QCauchySeq_bounded xn Pos.to_nat limx) as [Ax majx].
+ destruct (QCauchySeq_bounded yn Pos.to_nat limy) as [Ay majy].
+ destruct (QCauchySeq_bounded zn Pos.to_nat limz) as [Az majz].
+ simpl.
+ exists (fun p : positive => (Pos.to_nat (2 * (Pos.max (Pos.max Ax Ay) Az) * (2 * p)))).
+ intros p n k H H0.
+ setoid_replace (xn n * (yn n + zn n) -
+ (xn (Pos.to_nat (Pos.max Ax Ay)~0 * k)%nat *
+ yn (Pos.to_nat (Pos.max Ax Ay)~0 * k)%nat +
+ xn (Pos.to_nat (Pos.max Ax Az)~0 * k)%nat *
+ zn (Pos.to_nat (Pos.max Ax Az)~0 * k)%nat))%Q
+ with (xn n * yn n - (xn (Pos.to_nat (Pos.max Ax Ay)~0 * k)%nat *
+ yn (Pos.to_nat (Pos.max Ax Ay)~0 * k)%nat)
+ + (xn n * zn n - xn (Pos.to_nat (Pos.max Ax Az)~0 * k)%nat *
+ zn (Pos.to_nat (Pos.max Ax Az)~0 * k)%nat))%Q.
+ 2: ring.
+ apply (Qle_lt_trans _ (Qabs (xn n * yn n - (xn (Pos.to_nat (Pos.max Ax Ay)~0 * k)%nat *
+ yn (Pos.to_nat (Pos.max Ax Ay)~0 * k)%nat))
+ + Qabs (xn n * zn n - xn (Pos.to_nat (Pos.max Ax Az)~0 * k)%nat *
+ zn (Pos.to_nat (Pos.max Ax Az)~0 * k)%nat))).
+ apply Qabs_triangle.
+ setoid_replace (1#p)%Q with ((1#2*p) + (1#2*p))%Q.
+ apply Qplus_lt_le_compat.
+ + pose proof (CReal_mult_cauchy xn xn yn Ax Ay Pos.to_nat limx limy).
+ apply H1. apply majx. apply majy. rewrite max_l.
+ apply (le_trans _ (Pos.to_nat (2 * Pos.max (Pos.max Ax Ay) Az * (2 * p)))).
+ rewrite (Pos.mul_comm 2). rewrite <- Pos.mul_assoc.
+ rewrite <- (Pos.mul_comm (Pos.max (Pos.max Ax Ay) Az)).
+ rewrite <- Pos.mul_assoc.
+ rewrite Pos2Nat.inj_mul.
+ rewrite (Pos2Nat.inj_mul (Pos.max (Pos.max Ax Ay) Az)).
+ apply Nat.mul_le_mono_nonneg. apply le_0_n.
+ apply Pos2Nat.inj_le. apply Pos.le_max_l.
+ apply le_0_n. apply le_refl. apply H. apply le_refl.
+ rewrite max_l. apply (le_trans _ k).
+ apply (le_trans _ (Pos.to_nat (2 * Pos.max (Pos.max Ax Ay) Az * (2 * p)))).
+ rewrite (Pos.mul_comm 2). rewrite <- Pos.mul_assoc.
+ rewrite <- (Pos.mul_comm (Pos.max (Pos.max Ax Ay) Az)).
+ rewrite <- Pos.mul_assoc.
+ rewrite Pos2Nat.inj_mul.
+ rewrite (Pos2Nat.inj_mul (Pos.max (Pos.max Ax Ay) Az)).
+ apply Nat.mul_le_mono_nonneg. apply le_0_n.
+ apply Pos2Nat.inj_le. apply Pos.le_max_l.
+ apply le_0_n. apply le_refl. apply H0.
+ rewrite <- (mult_1_l k). rewrite mult_assoc.
+ apply Nat.mul_le_mono_nonneg. auto.
+ rewrite mult_1_r. apply Pos2Nat.is_pos. apply le_0_n.
+ apply le_refl. apply le_refl.
+ + apply Qlt_le_weak.
+ pose proof (CReal_mult_cauchy xn xn zn Ax Az Pos.to_nat limx limz).
+ apply H1. apply majx. apply majz. rewrite max_l. 2: apply le_refl.
+ apply (le_trans _ (Pos.to_nat (2 * Pos.max (Pos.max Ax Ay) Az * (2 * p)))).
+ rewrite (Pos.mul_comm 2). rewrite <- Pos.mul_assoc.
+ rewrite <- (Pos.mul_comm (Pos.max (Pos.max Ax Ay) Az)).
+ rewrite <- Pos.mul_assoc.
+ rewrite Pos2Nat.inj_mul.
+ rewrite (Pos2Nat.inj_mul (Pos.max (Pos.max Ax Ay) Az)).
+ apply Nat.mul_le_mono_nonneg. apply le_0_n.
+ rewrite <- Pos.max_assoc. rewrite (Pos.max_comm Ay Az).
+ rewrite Pos.max_assoc. apply Pos2Nat.inj_le. apply Pos.le_max_l.
+ apply le_0_n. apply le_refl. apply H.
+ rewrite max_l. apply (le_trans _ k).
+ apply (le_trans _ (Pos.to_nat (2 * Pos.max (Pos.max Ax Ay) Az * (2 * p)))).
+ rewrite (Pos.mul_comm 2). rewrite <- Pos.mul_assoc.
+ rewrite <- (Pos.mul_comm (Pos.max (Pos.max Ax Ay) Az)).
+ rewrite <- Pos.mul_assoc.
+ rewrite Pos2Nat.inj_mul.
+ rewrite (Pos2Nat.inj_mul (Pos.max (Pos.max Ax Ay) Az)).
+ apply Nat.mul_le_mono_nonneg. apply le_0_n.
+ rewrite <- Pos.max_assoc. rewrite (Pos.max_comm Ay Az).
+ rewrite Pos.max_assoc. apply Pos2Nat.inj_le. apply Pos.le_max_l.
+ apply le_0_n. apply le_refl. apply H0.
+ rewrite <- (mult_1_l k). rewrite mult_assoc.
+ apply Nat.mul_le_mono_nonneg. auto.
+ rewrite mult_1_r. apply Pos2Nat.is_pos. apply le_0_n.
+ apply le_refl. apply le_refl.
+ + rewrite Qinv_plus_distr. unfold Qeq. reflexivity.
+Qed.
+
+Lemma CReal_mult_plus_distr_r : forall r1 r2 r3 : CReal,
+ (r2 + r3) * r1 == (r2 * r1) + (r3 * r1).
+Proof.
+ intros.
+ rewrite CReal_mult_comm, CReal_mult_plus_distr_l,
+ <- (CReal_mult_comm r1), <- (CReal_mult_comm r1).
+ reflexivity.
+Qed.
+
+Lemma CReal_mult_1_l : forall r: CReal, 1 * r == r.
+Proof.
+ intros [rn limr]. split.
+ - intros [m maj]. simpl in maj.
+ destruct (QCauchySeq_bounded (fun _ : nat => 1%Q) Pos.to_nat (ConstCauchy 1)).
+ destruct (QCauchySeq_bounded rn Pos.to_nat limr).
+ simpl in maj. rewrite Qmult_1_l in maj.
+ specialize (limr m).
+ apply (Qlt_not_le (2 # m) (1 # m)).
+ apply (Qlt_trans _ (rn (Pos.to_nat m) - rn (Pos.to_nat (Pos.max x x0)~0 * Pos.to_nat m)%nat)).
+ apply maj.
+ apply (Qle_lt_trans _ (Qabs (rn (Pos.to_nat m) - rn (Pos.to_nat (Pos.max x x0)~0 * Pos.to_nat m)%nat))).
+ apply Qle_Qabs. apply limr. apply le_refl.
+ rewrite <- (mult_1_l (Pos.to_nat m)). rewrite mult_assoc.
+ apply Nat.mul_le_mono_nonneg. auto. rewrite mult_1_r.
+ apply Pos2Nat.is_pos. apply le_0_n. apply le_refl.
+ apply Z.mul_le_mono_nonneg. discriminate. discriminate.
+ discriminate. apply Z.le_refl.
+ - intros [m maj]. simpl in maj.
+ destruct (QCauchySeq_bounded (fun _ : nat => 1%Q) Pos.to_nat (ConstCauchy 1)).
+ destruct (QCauchySeq_bounded rn Pos.to_nat limr).
+ simpl in maj. rewrite Qmult_1_l in maj.
+ specialize (limr m).
+ apply (Qlt_not_le (2 # m) (1 # m)).
+ apply (Qlt_trans _ (rn (Pos.to_nat (Pos.max x x0)~0 * Pos.to_nat m)%nat - rn (Pos.to_nat m))).
+ apply maj.
+ apply (Qle_lt_trans _ (Qabs (rn (Pos.to_nat (Pos.max x x0)~0 * Pos.to_nat m)%nat - rn (Pos.to_nat m)))).
+ apply Qle_Qabs. apply limr.
+ rewrite <- (mult_1_l (Pos.to_nat m)). rewrite mult_assoc.
+ apply Nat.mul_le_mono_nonneg. auto. rewrite mult_1_r.
+ apply Pos2Nat.is_pos. apply le_0_n. apply le_refl.
+ apply le_refl. apply Z.mul_le_mono_nonneg. discriminate. discriminate.
+ discriminate. apply Z.le_refl.
+Qed.
+
+Lemma CReal_isRingExt : ring_eq_ext CReal_plus CReal_mult CReal_opp CRealEq.
+Proof.
+ split.
+ - intros x y H z t H0. apply CReal_plus_morph; assumption.
+ - intros x y H z t H0. apply (CRealEq_trans _ (CReal_mult x t)).
+ apply CReal_mult_proper_l. apply H0.
+ apply (CRealEq_trans _ (CReal_mult t x)). apply CReal_mult_comm.
+ apply (CRealEq_trans _ (CReal_mult t y)).
+ apply CReal_mult_proper_l. apply H. apply CReal_mult_comm.
+ - intros x y H. apply (CReal_plus_eq_reg_l x).
+ apply (CRealEq_trans _ (inject_Q 0)). apply CReal_plus_opp_r.
+ apply (CRealEq_trans _ (CReal_plus y (CReal_opp y))).
+ apply CRealEq_sym. apply CReal_plus_opp_r.
+ apply CReal_plus_proper_r. apply CRealEq_sym. apply H.
+Qed.
+
+Lemma CReal_isRing : ring_theory (inject_Q 0) (inject_Q 1)
+ CReal_plus CReal_mult
+ CReal_minus CReal_opp
+ CRealEq.
+Proof.
+ intros. split.
+ - apply CReal_plus_0_l.
+ - apply CReal_plus_comm.
+ - intros x y z. symmetry. apply CReal_plus_assoc.
+ - apply CReal_mult_1_l.
+ - apply CReal_mult_comm.
+ - intros x y z. symmetry. apply CReal_mult_assoc.
+ - intros x y z. rewrite <- (CReal_mult_comm z).
+ rewrite CReal_mult_plus_distr_l.
+ apply (CRealEq_trans _ (CReal_plus (CReal_mult x z) (CReal_mult z y))).
+ apply CReal_plus_proper_r. apply CReal_mult_comm.
+ apply CReal_plus_proper_l. apply CReal_mult_comm.
+ - intros x y. apply CRealEq_refl.
+ - apply CReal_plus_opp_r.
+Qed.
+
+Add Parametric Morphism : CReal_mult
+ with signature CRealEq ==> CRealEq ==> CRealEq
+ as CReal_mult_morph.
+Proof.
+ apply CReal_isRingExt.
+Qed.
+
+Instance CReal_mult_morph_T
+ : CMorphisms.Proper
+ (CMorphisms.respectful CRealEq (CMorphisms.respectful CRealEq CRealEq)) CReal_mult.
+Proof.
+ apply CReal_isRingExt.
+Qed.
+
+Add Parametric Morphism : CReal_opp
+ with signature CRealEq ==> CRealEq
+ as CReal_opp_morph.
+Proof.
+ apply (Ropp_ext CReal_isRingExt).
+Qed.
+
+Instance CReal_opp_morph_T
+ : CMorphisms.Proper
+ (CMorphisms.respectful CRealEq CRealEq) CReal_opp.
+Proof.
+ apply CReal_isRingExt.
+Qed.
+
+Add Parametric Morphism : CReal_minus
+ with signature CRealEq ==> CRealEq ==> CRealEq
+ as CReal_minus_morph.
+Proof.
+ intros. unfold CReal_minus. rewrite H,H0. reflexivity.
+Qed.
+
+Instance CReal_minus_morph_T
+ : CMorphisms.Proper
+ (CMorphisms.respectful CRealEq (CMorphisms.respectful CRealEq CRealEq)) CReal_minus.
+Proof.
+ intros x y exy z t ezt. unfold CReal_minus. rewrite exy,ezt. reflexivity.
+Qed.
+
+Add Ring CRealRing : CReal_isRing.
+
+(**********)
+Lemma CReal_mult_0_l : forall r, 0 * r == 0.
+Proof.
+ intro; ring.
+Qed.
+
+Lemma CReal_mult_0_r : forall r, r * 0 == 0.
+Proof.
+ intro; ring.
+Qed.
+
+(**********)
+Lemma CReal_mult_1_r : forall r, r * 1 == r.
+Proof.
+ intro; ring.
+Qed.
+
+Lemma CReal_opp_mult_distr_l
+ : forall r1 r2 : CReal, - (r1 * r2) == (- r1) * r2.
+Proof.
+ intros. ring.
+Qed.
+
+Lemma CReal_opp_mult_distr_r
+ : forall r1 r2 : CReal, - (r1 * r2) == r1 * (- r2).
+Proof.
+ intros. ring.
+Qed.
+
+Lemma CReal_mult_lt_compat_l : forall x y z : CReal,
+ 0 < x -> y < z -> x*y < x*z.
+Proof.
+ intros. apply (CReal_plus_lt_reg_l
+ (CReal_opp (CReal_mult x y))).
+ rewrite CReal_plus_comm. pose proof CReal_plus_opp_r.
+ unfold CReal_minus in H1. rewrite H1.
+ rewrite CReal_mult_comm, CReal_opp_mult_distr_l, CReal_mult_comm.
+ rewrite <- CReal_mult_plus_distr_l.
+ apply CReal_mult_lt_0_compat. exact H.
+ apply (CReal_plus_lt_reg_l y).
+ rewrite CReal_plus_comm, CReal_plus_0_l.
+ rewrite <- CReal_plus_assoc, H1, CReal_plus_0_l. exact H0.
+Qed.
+
+Lemma CReal_mult_lt_compat_r : forall x y z : CReal,
+ 0 < x -> y < z -> y*x < z*x.
+Proof.
+ intros. rewrite <- (CReal_mult_comm x), <- (CReal_mult_comm x).
+ apply (CReal_mult_lt_compat_l x); assumption.
+Qed.
+
+Lemma CReal_mult_eq_reg_l : forall (r r1 r2 : CReal),
+ r # 0
+ -> CRealEq (CReal_mult r r1) (CReal_mult r r2)
+ -> CRealEq r1 r2.
+Proof.
+ intros. destruct H; split.
+ - intro abs. apply (CReal_mult_lt_compat_l (-r)) in abs.
+ rewrite <- CReal_opp_mult_distr_l, <- CReal_opp_mult_distr_l, H0 in abs.
+ exact (CRealLt_irrefl _ abs). apply (CReal_plus_lt_reg_l r).
+ rewrite CReal_plus_opp_r, CReal_plus_comm, CReal_plus_0_l. exact c.
+ - intro abs. apply (CReal_mult_lt_compat_l (-r)) in abs.
+ rewrite <- CReal_opp_mult_distr_l, <- CReal_opp_mult_distr_l, H0 in abs.
+ exact (CRealLt_irrefl _ abs). apply (CReal_plus_lt_reg_l r).
+ rewrite CReal_plus_opp_r, CReal_plus_comm, CReal_plus_0_l. exact c.
+ - intro abs. apply (CReal_mult_lt_compat_l r) in abs. rewrite H0 in abs.
+ exact (CRealLt_irrefl _ abs). exact c.
+ - intro abs. apply (CReal_mult_lt_compat_l r) in abs. rewrite H0 in abs.
+ exact (CRealLt_irrefl _ abs). exact c.
+Qed.
+
+Lemma CReal_abs_appart_zero : forall (x : CReal) (n : positive),
+ Qlt (2#n) (Qabs (proj1_sig x (Pos.to_nat n)))
+ -> 0 # x.
+Proof.
+ intros. destruct x as [xn xcau]. simpl in H.
+ destruct (Qlt_le_dec 0 (xn (Pos.to_nat n))).
+ - left. exists n; simpl. rewrite Qabs_pos in H.
+ ring_simplify. exact H. apply Qlt_le_weak. exact q.
+ - right. exists n; simpl. rewrite Qabs_neg in H.
+ unfold Qminus. rewrite Qplus_0_l. exact H. exact q.
+Qed.
+
+
+(*********************************************************)
+(** * Field *)
+(*********************************************************)
+
+Lemma CRealArchimedean
+ : forall x:CReal, { n:Z & x < inject_Q (n#1) < x+2 }.
+Proof.
+ (* Locate x within 1/4 and pick the first integer above this interval. *)
+ intros [xn limx].
+ pose proof (Qlt_floor (xn 4%nat + (1#4))). unfold inject_Z in H.
+ pose proof (Qfloor_le (xn 4%nat + (1#4))). unfold inject_Z in H0.
+ remember (Qfloor (xn 4%nat + (1#4)))%Z as n.
+ exists (n+1)%Z. split.
+ - assert (Qlt 0 ((n + 1 # 1) - (xn 4%nat + (1 # 4)))) as epsPos.
+ { unfold Qminus. rewrite <- Qlt_minus_iff. exact H. }
+ destruct (Qarchimedean (/((1#2)*((n + 1 # 1) - (xn 4%nat + (1 # 4)))))) as [k kmaj].
+ exists (Pos.max 4 k). simpl.
+ apply (Qlt_trans _ ((n + 1 # 1) - (xn 4%nat + (1 # 4)))).
+ + setoid_replace (Z.pos k # 1)%Q with (/(1#k))%Q in kmaj. 2: reflexivity.
+ rewrite <- Qinv_lt_contravar in kmaj. 2: reflexivity.
+ apply (Qle_lt_trans _ (2#k)).
+ rewrite <- (Qmult_le_l _ _ (1#2)).
+ setoid_replace ((1 # 2) * (2 # k))%Q with (1#k)%Q. 2: reflexivity.
+ setoid_replace ((1 # 2) * (2 # Pos.max 4 k))%Q with (1#Pos.max 4 k)%Q. 2: reflexivity.
+ unfold Qle; simpl. apply Pos2Z.pos_le_pos. apply Pos.le_max_r.
+ reflexivity.
+ rewrite <- (Qmult_lt_l _ _ (1#2)).
+ setoid_replace ((1 # 2) * (2 # k))%Q with (1#k)%Q. exact kmaj.
+ reflexivity. reflexivity. rewrite <- (Qmult_0_r (1#2)).
+ rewrite Qmult_lt_l. exact epsPos. reflexivity.
+ + rewrite <- (Qplus_lt_r _ _ (xn (Pos.to_nat (Pos.max 4 k)) - (n + 1 # 1) + (1#4))).
+ ring_simplify.
+ apply (Qle_lt_trans _ (Qabs (xn (Pos.to_nat (Pos.max 4 k)) - xn 4%nat))).
+ apply Qle_Qabs. apply limx.
+ rewrite Pos2Nat.inj_max. apply Nat.le_max_l. apply le_refl.
+ - apply (CReal_plus_lt_reg_l (-(2))). ring_simplify.
+ exists 4%positive. simpl.
+ rewrite <- Qinv_plus_distr.
+ rewrite <- (Qplus_lt_r _ _ ((n#1) - (1#2))). ring_simplify.
+ apply (Qle_lt_trans _ (xn 4%nat + (1 # 4)) _ H0).
+ unfold Pos.to_nat; simpl.
+ rewrite <- (Qplus_lt_r _ _ (-xn 4%nat)). ring_simplify.
+ reflexivity.
+Defined.
+
+Definition Rup_pos (x : CReal)
+ : { n : positive & x < inject_Q (Z.pos n # 1) }.
+Proof.
+ intros. destruct (CRealArchimedean x) as [p [maj _]].
+ destruct p.
+ - exists 1%positive. apply (CReal_lt_trans _ 0 _ maj). apply CRealLt_0_1.
+ - exists p. exact maj.
+ - exists 1%positive. apply (CReal_lt_trans _ (inject_Q (Z.neg p # 1)) _ maj).
+ apply (CReal_lt_trans _ 0). apply inject_Q_lt. reflexivity.
+ apply CRealLt_0_1.
+Qed.
+
+Lemma CRealLtDisjunctEpsilon : forall a b c d : CReal,
+ (CRealLtProp a b \/ CRealLtProp c d) -> CRealLt a b + CRealLt c d.
+Proof.
+ intros.
+ assert (exists n : nat, n <> O /\
+ (Qlt (2 # Pos.of_nat n) (proj1_sig b n - proj1_sig a n)
+ \/ Qlt (2 # Pos.of_nat n) (proj1_sig d n - proj1_sig c n))).
+ { destruct H. destruct H as [n maj]. exists (Pos.to_nat n). split.
+ intro abs. destruct (Pos2Nat.is_succ n). rewrite H in abs.
+ inversion abs. left. rewrite Pos2Nat.id. apply maj.
+ destruct H as [n maj]. exists (Pos.to_nat n). split.
+ intro abs. destruct (Pos2Nat.is_succ n). rewrite H in abs.
+ inversion abs. right. rewrite Pos2Nat.id. apply maj. }
+ apply constructive_indefinite_ground_description_nat in H0.
+ - destruct H0 as [n [nPos maj]].
+ destruct (Qlt_le_dec (2 # Pos.of_nat n)
+ (proj1_sig b n - proj1_sig a n)).
+ left. exists (Pos.of_nat n). rewrite Nat2Pos.id. apply q. apply nPos.
+ assert (2 # Pos.of_nat n < proj1_sig d n - proj1_sig c n)%Q.
+ destruct maj. exfalso.
+ apply (Qlt_not_le (2 # Pos.of_nat n) (proj1_sig b n - proj1_sig a n)); assumption.
+ assumption. clear maj. right. exists (Pos.of_nat n). rewrite Nat2Pos.id.
+ apply H0. apply nPos.
+ - clear H0. clear H. intro n. destruct n. right.
+ intros [abs _]. exact (abs (eq_refl O)).
+ destruct (Qlt_le_dec (2 # Pos.of_nat (S n)) (proj1_sig b (S n) - proj1_sig a (S n))).
+ left. split. discriminate. left. apply q.
+ destruct (Qlt_le_dec (2 # Pos.of_nat (S n)) (proj1_sig d (S n) - proj1_sig c (S n))).
+ left. split. discriminate. right. apply q0.
+ right. intros [_ [abs|abs]].
+ apply (Qlt_not_le (2 # Pos.of_nat (S n))
+ (proj1_sig b (S n) - proj1_sig a (S n))); assumption.
+ apply (Qlt_not_le (2 # Pos.of_nat (S n))
+ (proj1_sig d (S n) - proj1_sig c (S n))); assumption.
+Qed.
+
+Lemma CRealShiftReal : forall (x : CReal) (k : nat),
+ QCauchySeq (fun n => proj1_sig x (plus n k)) Pos.to_nat.
+Proof.
+ intros x k n p q H H0.
+ destruct x as [xn cau]; unfold proj1_sig.
+ destruct k. rewrite plus_0_r. rewrite plus_0_r. apply cau; assumption.
+ specialize (cau (n + Pos.of_nat (S k))%positive (p + S k)%nat (q + S k)%nat).
+ apply (Qlt_trans _ (1 # n + Pos.of_nat (S k))).
+ apply cau. rewrite Pos2Nat.inj_add. rewrite Nat2Pos.id.
+ apply Nat.add_le_mono_r. apply H. discriminate.
+ rewrite Pos2Nat.inj_add. rewrite Nat2Pos.id.
+ apply Nat.add_le_mono_r. apply H0. discriminate.
+ apply Pos2Nat.inj_lt; simpl. rewrite Pos2Nat.inj_add.
+ rewrite <- (plus_0_r (Pos.to_nat n)). rewrite <- plus_assoc.
+ apply Nat.add_lt_mono_l. apply Pos2Nat.is_pos.
+Qed.
+
+Lemma CRealShiftEqual : forall (x : CReal) (k : nat),
+ CRealEq x (exist _ (fun n => proj1_sig x (plus n k)) (CRealShiftReal x k)).
+Proof.
+ intros. split.
+ - intros [n maj]. destruct x as [xn cau]; simpl in maj.
+ specialize (cau n (Pos.to_nat n + k)%nat (Pos.to_nat n)).
+ apply Qlt_not_le in maj. apply maj. clear maj.
+ apply (Qle_trans _ (Qabs (xn (Pos.to_nat n + k)%nat - xn (Pos.to_nat n)))).
+ apply Qle_Qabs. apply (Qle_trans _ (1#n)). apply Zlt_le_weak.
+ apply cau. rewrite <- (plus_0_r (Pos.to_nat n)).
+ rewrite <- plus_assoc. apply Nat.add_le_mono_l. apply le_0_n.
+ apply le_refl. apply Z.mul_le_mono_pos_r. apply Pos2Z.is_pos.
+ discriminate.
+ - intros [n maj]. destruct x as [xn cau]; simpl in maj.
+ specialize (cau n (Pos.to_nat n) (Pos.to_nat n + k)%nat).
+ apply Qlt_not_le in maj. apply maj. clear maj.
+ apply (Qle_trans _ (Qabs (xn (Pos.to_nat n) - xn (Pos.to_nat n + k)%nat))).
+ apply Qle_Qabs. apply (Qle_trans _ (1#n)). apply Zlt_le_weak.
+ apply cau. apply le_refl. rewrite <- (plus_0_r (Pos.to_nat n)).
+ rewrite <- plus_assoc. apply Nat.add_le_mono_l. apply le_0_n.
+ apply Z.mul_le_mono_pos_r. apply Pos2Z.is_pos. discriminate.
+Qed.
+
+(* Find an equal negative real number, which rational sequence
+ stays below 0, so that it can be inversed. *)
+Definition CRealNegShift (x : CReal)
+ : CRealLt x (inject_Q 0)
+ -> { y : prod positive CReal | CRealEq x (snd y)
+ /\ forall n:nat, Qlt (proj1_sig (snd y) n) (-1 # fst y) }.
+Proof.
+ intro xNeg.
+ pose proof (CRealLt_aboveSig x (inject_Q 0)).
+ pose proof (CRealShiftReal x).
+ pose proof (CRealShiftEqual x).
+ destruct xNeg as [n maj], x as [xn cau]; simpl in maj.
+ specialize (H n maj); simpl in H.
+ destruct (Qarchimedean (/ (0 - xn (Pos.to_nat n) - (2 # n)))) as [a _].
+ remember (Pos.max n a~0) as k.
+ clear Heqk. clear maj. clear n.
+ exists (pair k
+ (exist _ (fun n => xn (plus n (Pos.to_nat k))) (H0 (Pos.to_nat k)))).
+ split. apply H1. intro n. simpl. apply Qlt_minus_iff.
+ destruct n.
+ - specialize (H k).
+ unfold Qminus in H. rewrite Qplus_0_l in H. apply Qlt_minus_iff in H.
+ unfold Qminus. rewrite Qplus_comm.
+ apply (Qlt_trans _ (- xn (Pos.to_nat k)%nat - (2 #k))). apply H.
+ unfold Qminus. simpl. apply Qplus_lt_r.
+ apply Z.mul_lt_mono_pos_r. simpl. apply Pos2Z.is_pos.
+ reflexivity. apply Pos.le_refl.
+ - apply (Qlt_trans _ (-(2 # k) - xn (S n + Pos.to_nat k)%nat)).
+ rewrite <- (Nat2Pos.id (S n)). rewrite <- Pos2Nat.inj_add.
+ specialize (H (Pos.of_nat (S n) + k)%positive).
+ unfold Qminus in H. rewrite Qplus_0_l in H. apply Qlt_minus_iff in H.
+ unfold Qminus. rewrite Qplus_comm. apply H. apply Pos2Nat.inj_le.
+ rewrite <- (plus_0_l (Pos.to_nat k)). rewrite Pos2Nat.inj_add.
+ apply Nat.add_le_mono_r. apply le_0_n. discriminate.
+ apply Qplus_lt_l.
+ apply Z.mul_lt_mono_pos_r. simpl. apply Pos2Z.is_pos.
+ reflexivity.
+Qed.
+
+Definition CRealPosShift (x : CReal)
+ : inject_Q 0 < x
+ -> { y : prod positive CReal | CRealEq x (snd y)
+ /\ forall n:nat, Qlt (1 # fst y) (proj1_sig (snd y) n) }.
+Proof.
+ intro xPos.
+ pose proof (CRealLt_aboveSig (inject_Q 0) x).
+ pose proof (CRealShiftReal x).
+ pose proof (CRealShiftEqual x).
+ destruct xPos as [n maj], x as [xn cau]; simpl in maj.
+ simpl in H. specialize (H n).
+ destruct (Qarchimedean (/ (xn (Pos.to_nat n) - 0 - (2 # n)))) as [a _].
+ specialize (H maj); simpl in H.
+ remember (Pos.max n a~0) as k.
+ clear Heqk. clear maj. clear n.
+ exists (pair k
+ (exist _ (fun n => xn (plus n (Pos.to_nat k))) (H0 (Pos.to_nat k)))).
+ split. apply H1. intro n. simpl. apply Qlt_minus_iff.
+ destruct n.
+ - specialize (H k).
+ unfold Qminus in H. rewrite Qplus_0_r in H.
+ simpl. rewrite <- Qlt_minus_iff.
+ apply (Qlt_trans _ (2 #k)).
+ apply Z.mul_lt_mono_pos_r. simpl. apply Pos2Z.is_pos.
+ reflexivity. apply H. apply Pos.le_refl.
+ - rewrite <- Qlt_minus_iff. apply (Qlt_trans _ (2 # k)).
+ apply Z.mul_lt_mono_pos_r. simpl. apply Pos2Z.is_pos.
+ reflexivity. specialize (H (Pos.of_nat (S n) + k)%positive).
+ unfold Qminus in H. rewrite Qplus_0_r in H.
+ rewrite Pos2Nat.inj_add in H. rewrite Nat2Pos.id in H.
+ apply H. apply Pos2Nat.inj_le.
+ rewrite <- (plus_0_l (Pos.to_nat k)). rewrite Pos2Nat.inj_add.
+ apply Nat.add_le_mono_r. apply le_0_n. discriminate.
+Qed.
+
+Lemma CReal_inv_neg : forall (yn : nat -> Q) (k : positive),
+ (QCauchySeq yn Pos.to_nat)
+ -> (forall n : nat, yn n < -1 # k)%Q
+ -> QCauchySeq (fun n : nat => / yn (Pos.to_nat k ^ 2 * n)%nat) Pos.to_nat.
+Proof.
+ (* Prove the inverse sequence is Cauchy *)
+ intros yn k cau maj n p q H0 H1.
+ setoid_replace (/ yn (Pos.to_nat k ^ 2 * p)%nat -
+ / yn (Pos.to_nat k ^ 2 * q)%nat)%Q
+ with ((yn (Pos.to_nat k ^ 2 * q)%nat -
+ yn (Pos.to_nat k ^ 2 * p)%nat)
+ / (yn (Pos.to_nat k ^ 2 * q)%nat *
+ yn (Pos.to_nat k ^ 2 * p)%nat)).
+ + apply (Qle_lt_trans _ (Qabs (yn (Pos.to_nat k ^ 2 * q)%nat
+ - yn (Pos.to_nat k ^ 2 * p)%nat)
+ / (1 # (k^2)))).
+ assert (1 # k ^ 2
+ < Qabs (yn (Pos.to_nat k ^ 2 * q)%nat * yn (Pos.to_nat k ^ 2 * p)%nat))%Q.
+ { rewrite Qabs_Qmult. unfold "^"%positive; simpl.
+ rewrite factorDenom. rewrite Pos.mul_1_r.
+ apply (Qlt_trans _ ((1#k) * Qabs (yn (Pos.to_nat k * (Pos.to_nat k * 1) * p)%nat))).
+ apply Qmult_lt_l. reflexivity. rewrite Qabs_neg.
+ specialize (maj (Pos.to_nat k * (Pos.to_nat k * 1) * p)%nat).
+ apply Qlt_minus_iff in maj. apply Qlt_minus_iff.
+ rewrite Qplus_comm. setoid_replace (-(1#k))%Q with (-1 # k)%Q. apply maj.
+ reflexivity. apply (Qle_trans _ (-1 # k)). apply Zlt_le_weak.
+ apply maj. discriminate.
+ apply Qmult_lt_r. apply (Qlt_trans 0 (1#k)). reflexivity.
+ rewrite Qabs_neg.
+ specialize (maj (Pos.to_nat k * (Pos.to_nat k * 1) * p)%nat).
+ apply Qlt_minus_iff in maj. apply Qlt_minus_iff.
+ rewrite Qplus_comm. setoid_replace (-(1#k))%Q with (-1 # k)%Q. apply maj.
+ reflexivity. apply (Qle_trans _ (-1 # k)). apply Zlt_le_weak.
+ apply maj. discriminate.
+ rewrite Qabs_neg.
+ specialize (maj (Pos.to_nat k * (Pos.to_nat k * 1) * q)%nat).
+ apply Qlt_minus_iff in maj. apply Qlt_minus_iff.
+ rewrite Qplus_comm. setoid_replace (-(1#k))%Q with (-1 # k)%Q. apply maj.
+ reflexivity. apply (Qle_trans _ (-1 # k)). apply Zlt_le_weak.
+ apply maj. discriminate. }
+ unfold Qdiv. rewrite Qabs_Qmult. rewrite Qabs_Qinv.
+ rewrite Qmult_comm. rewrite <- (Qmult_comm (/ (1 # k ^ 2))).
+ apply Qmult_le_compat_r. apply Qlt_le_weak.
+ rewrite <- Qmult_1_l. apply Qlt_shift_div_r.
+ apply (Qlt_trans 0 (1 # k ^ 2)). reflexivity. apply H.
+ rewrite Qmult_comm. apply Qlt_shift_div_l.
+ reflexivity. rewrite Qmult_1_l. apply H.
+ apply Qabs_nonneg. simpl in maj.
+ specialize (cau (n * (k^2))%positive
+ (Pos.to_nat k ^ 2 * q)%nat
+ (Pos.to_nat k ^ 2 * p)%nat).
+ apply Qlt_shift_div_r. reflexivity.
+ apply (Qlt_le_trans _ (1 # n * k ^ 2)). apply cau.
+ rewrite Pos2Nat.inj_mul. rewrite mult_comm.
+ unfold "^"%positive. simpl. rewrite Pos2Nat.inj_mul.
+ rewrite <- mult_assoc. rewrite <- mult_assoc.
+ apply Nat.mul_le_mono_nonneg_l. apply le_0_n.
+ rewrite (mult_1_r). rewrite Pos.mul_1_r.
+ apply Nat.mul_le_mono_nonneg_l. apply le_0_n.
+ apply (le_trans _ (q+0)). rewrite plus_0_r. assumption.
+ rewrite plus_0_r. apply le_refl.
+ rewrite Pos2Nat.inj_mul. rewrite mult_comm.
+ unfold "^"%positive; simpl. rewrite Pos2Nat.inj_mul.
+ rewrite <- mult_assoc. rewrite <- mult_assoc.
+ apply Nat.mul_le_mono_nonneg_l. apply le_0_n.
+ rewrite (mult_1_r). rewrite Pos.mul_1_r.
+ apply Nat.mul_le_mono_nonneg_l. apply le_0_n.
+ apply (le_trans _ (p+0)). rewrite plus_0_r. assumption.
+ rewrite plus_0_r. apply le_refl.
+ rewrite factorDenom. apply Qle_refl.
+ + field. split. intro abs.
+ specialize (maj (Pos.to_nat k ^ 2 * p)%nat).
+ rewrite abs in maj. inversion maj.
+ intro abs.
+ specialize (maj (Pos.to_nat k ^ 2 * q)%nat).
+ rewrite abs in maj. inversion maj.
+Qed.
+
+Lemma CReal_inv_pos : forall (yn : nat -> Q) (k : positive),
+ (QCauchySeq yn Pos.to_nat)
+ -> (forall n : nat, 1 # k < yn n)%Q
+ -> QCauchySeq (fun n : nat => / yn (Pos.to_nat k ^ 2 * n)%nat) Pos.to_nat.
+Proof.
+ intros yn k cau maj n p q H0 H1.
+ setoid_replace (/ yn (Pos.to_nat k ^ 2 * p)%nat -
+ / yn (Pos.to_nat k ^ 2 * q)%nat)%Q
+ with ((yn (Pos.to_nat k ^ 2 * q)%nat -
+ yn (Pos.to_nat k ^ 2 * p)%nat)
+ / (yn (Pos.to_nat k ^ 2 * q)%nat *
+ yn (Pos.to_nat k ^ 2 * p)%nat)).
+ + apply (Qle_lt_trans _ (Qabs (yn (Pos.to_nat k ^ 2 * q)%nat
+ - yn (Pos.to_nat k ^ 2 * p)%nat)
+ / (1 # (k^2)))).
+ assert (1 # k ^ 2
+ < Qabs (yn (Pos.to_nat k ^ 2 * q)%nat * yn (Pos.to_nat k ^ 2 * p)%nat))%Q.
+ { rewrite Qabs_Qmult. unfold "^"%positive; simpl.
+ rewrite factorDenom. rewrite Pos.mul_1_r.
+ apply (Qlt_trans _ ((1#k) * Qabs (yn (Pos.to_nat k * (Pos.to_nat k * 1) * p)%nat))).
+ apply Qmult_lt_l. reflexivity. rewrite Qabs_pos.
+ specialize (maj (Pos.to_nat k * (Pos.to_nat k * 1) * p)%nat).
+ apply maj. apply (Qle_trans _ (1 # k)).
+ discriminate. apply Zlt_le_weak. apply maj.
+ apply Qmult_lt_r. apply (Qlt_trans 0 (1#k)). reflexivity.
+ rewrite Qabs_pos.
+ specialize (maj (Pos.to_nat k * (Pos.to_nat k * 1) * p)%nat).
+ apply maj. apply (Qle_trans _ (1 # k)). discriminate.
+ apply Zlt_le_weak. apply maj.
+ rewrite Qabs_pos.
+ specialize (maj (Pos.to_nat k * (Pos.to_nat k * 1) * q)%nat).
+ apply maj. apply (Qle_trans _ (1 # k)). discriminate.
+ apply Zlt_le_weak. apply maj. }
+ unfold Qdiv. rewrite Qabs_Qmult. rewrite Qabs_Qinv.
+ rewrite Qmult_comm. rewrite <- (Qmult_comm (/ (1 # k ^ 2))).
+ apply Qmult_le_compat_r. apply Qlt_le_weak.
+ rewrite <- Qmult_1_l. apply Qlt_shift_div_r.
+ apply (Qlt_trans 0 (1 # k ^ 2)). reflexivity. apply H.
+ rewrite Qmult_comm. apply Qlt_shift_div_l.
+ reflexivity. rewrite Qmult_1_l. apply H.
+ apply Qabs_nonneg. simpl in maj.
+ specialize (cau (n * (k^2))%positive
+ (Pos.to_nat k ^ 2 * q)%nat
+ (Pos.to_nat k ^ 2 * p)%nat).
+ apply Qlt_shift_div_r. reflexivity.
+ apply (Qlt_le_trans _ (1 # n * k ^ 2)). apply cau.
+ rewrite Pos2Nat.inj_mul. rewrite mult_comm.
+ unfold "^"%positive. simpl. rewrite Pos2Nat.inj_mul.
+ rewrite <- mult_assoc. rewrite <- mult_assoc.
+ apply Nat.mul_le_mono_nonneg_l. apply le_0_n.
+ rewrite (mult_1_r). rewrite Pos.mul_1_r.
+ apply Nat.mul_le_mono_nonneg_l. apply le_0_n.
+ apply (le_trans _ (q+0)). rewrite plus_0_r. assumption.
+ rewrite plus_0_r. apply le_refl.
+ rewrite Pos2Nat.inj_mul. rewrite mult_comm.
+ unfold "^"%positive; simpl. rewrite Pos2Nat.inj_mul.
+ rewrite <- mult_assoc. rewrite <- mult_assoc.
+ apply Nat.mul_le_mono_nonneg_l. apply le_0_n.
+ rewrite (mult_1_r). rewrite Pos.mul_1_r.
+ apply Nat.mul_le_mono_nonneg_l. apply le_0_n.
+ apply (le_trans _ (p+0)). rewrite plus_0_r. assumption.
+ rewrite plus_0_r. apply le_refl.
+ rewrite factorDenom. apply Qle_refl.
+ + field. split. intro abs.
+ specialize (maj (Pos.to_nat k ^ 2 * p)%nat).
+ rewrite abs in maj. inversion maj.
+ intro abs.
+ specialize (maj (Pos.to_nat k ^ 2 * q)%nat).
+ rewrite abs in maj. inversion maj.
+Qed.
+
+Definition CReal_inv (x : CReal) (xnz : x # 0) : CReal.
+Proof.
+ destruct xnz as [xNeg | xPos].
+ - destruct (CRealNegShift x xNeg) as [[k y] [_ maj]].
+ destruct y as [yn cau]; unfold proj1_sig, snd, fst in maj.
+ exists (fun n => Qinv (yn (mult (Pos.to_nat k^2) n))).
+ apply (CReal_inv_neg yn). apply cau. apply maj.
+ - destruct (CRealPosShift x xPos) as [[k y] [_ maj]].
+ destruct y as [yn cau]; unfold proj1_sig, snd, fst in maj.
+ exists (fun n => Qinv (yn (mult (Pos.to_nat k^2) n))).
+ apply (CReal_inv_pos yn). apply cau. apply maj.
+Defined.
+
+Notation "/ x" := (CReal_inv x) (at level 35, right associativity) : CReal_scope.
+
+Lemma CReal_inv_0_lt_compat
+ : forall (r : CReal) (rnz : r # 0),
+ 0 < r -> 0 < ((/ r) rnz).
+Proof.
+ intros. unfold CReal_inv. simpl.
+ destruct rnz.
+ - exfalso. apply CRealLt_asym in H. contradiction.
+ - destruct (CRealPosShift r c) as [[k rpos] [req maj]].
+ clear req. destruct rpos as [rn cau]; simpl in maj.
+ unfold CRealLt; simpl.
+ destruct (Qarchimedean (rn 1%nat)) as [A majA].
+ exists (2 * (A + 1))%positive. unfold Qminus. rewrite Qplus_0_r.
+ rewrite <- (Qmult_1_l (Qinv (rn (Pos.to_nat k * (Pos.to_nat k * 1) * Pos.to_nat (2 * (A + 1)))%nat))).
+ apply Qlt_shift_div_l. apply (Qlt_trans 0 (1#k)). reflexivity.
+ apply maj. rewrite <- (Qmult_inv_r (Z.pos A + 1 # 1)).
+ setoid_replace (2 # 2 * (A + 1))%Q with (Qinv (Z.pos A + 1 # 1)).
+ 2: reflexivity.
+ rewrite Qmult_comm. apply Qmult_lt_r. reflexivity.
+ rewrite mult_1_r. rewrite <- Pos2Nat.inj_mul. rewrite <- Pos2Nat.inj_mul.
+ rewrite <- (Qplus_lt_l _ _ (- rn 1%nat)).
+ apply (Qle_lt_trans _ (Qabs (rn (Pos.to_nat (k * k * (2 * (A + 1)))) + - rn 1%nat))).
+ apply Qle_Qabs. apply (Qlt_le_trans _ 1). apply cau.
+ apply Pos2Nat.is_pos. apply le_refl.
+ rewrite <- Qinv_plus_distr. rewrite <- (Qplus_comm 1).
+ rewrite <- Qplus_0_r. rewrite <- Qplus_assoc. rewrite <- Qplus_assoc.
+ rewrite Qplus_le_r. rewrite Qplus_0_l. apply Qlt_le_weak.
+ apply Qlt_minus_iff in majA. apply majA.
+ intro abs. inversion abs.
+Qed.
+
+Lemma CReal_linear_shift : forall (x : CReal) (k : nat),
+ le 1 k -> QCauchySeq (fun n => proj1_sig x (k * n)%nat) Pos.to_nat.
+Proof.
+ intros [xn limx] k lek p n m H H0. unfold proj1_sig.
+ apply limx. apply (le_trans _ n). apply H.
+ rewrite <- (mult_1_l n). rewrite mult_assoc.
+ apply Nat.mul_le_mono_nonneg_r. apply le_0_n.
+ rewrite mult_1_r. apply lek. apply (le_trans _ m). apply H0.
+ rewrite <- (mult_1_l m). rewrite mult_assoc.
+ apply Nat.mul_le_mono_nonneg_r. apply le_0_n.
+ rewrite mult_1_r. apply lek.
+Qed.
+
+Lemma CReal_linear_shift_eq : forall (x : CReal) (k : nat) (kPos : le 1 k),
+ CRealEq x
+ (exist (fun n : nat -> Q => QCauchySeq n Pos.to_nat)
+ (fun n : nat => proj1_sig x (k * n)%nat) (CReal_linear_shift x k kPos)).
+Proof.
+ intros. apply CRealEq_diff. intro n.
+ destruct x as [xn limx]; unfold proj1_sig.
+ specialize (limx n (Pos.to_nat n) (k * Pos.to_nat n)%nat).
+ apply (Qle_trans _ (1 # n)). apply Qlt_le_weak. apply limx.
+ apply le_refl. rewrite <- (mult_1_l (Pos.to_nat n)).
+ rewrite mult_assoc. apply Nat.mul_le_mono_nonneg_r. apply le_0_n.
+ rewrite mult_1_r. apply kPos. apply Z.mul_le_mono_nonneg_r.
+ discriminate. discriminate.
+Qed.
+
+Lemma CReal_inv_l : forall (r:CReal) (rnz : r # 0),
+ ((/ r) rnz) * r == 1.
+Proof.
+ intros. unfold CReal_inv; simpl.
+ destruct rnz.
+ - (* r < 0 *) destruct (CRealNegShift r c) as [[k rneg] [req maj]].
+ simpl in req. apply CRealEq_diff. apply CRealEq_modindep.
+ apply (QSeqEquivEx_trans _
+ (proj1_sig (CReal_mult ((let
+ (yn, cau) as s
+ return ((forall n : nat, proj1_sig s n < -1 # k) -> CReal) := rneg in
+ fun maj0 : forall n : nat, yn n < -1 # k =>
+ exist (fun x : nat -> Q => QCauchySeq x Pos.to_nat)
+ (fun n : nat => Qinv (yn (Pos.to_nat k * (Pos.to_nat k * 1) * n))%nat)
+ (CReal_inv_neg yn k cau maj0)) maj) rneg)))%Q.
+ + apply CRealEq_modindep. apply CRealEq_diff.
+ apply CReal_mult_proper_l. apply req.
+ + assert (le 1 (Pos.to_nat k * (Pos.to_nat k * 1))%nat). rewrite mult_1_r.
+ rewrite <- Pos2Nat.inj_mul. apply Pos2Nat.is_pos.
+ apply (QSeqEquivEx_trans _
+ (proj1_sig (CReal_mult ((let
+ (yn, cau) as s
+ return ((forall n : nat, proj1_sig s n < -1 # k) -> CReal) := rneg in
+ fun maj0 : forall n : nat, yn n < -1 # k =>
+ exist (fun x : nat -> Q => QCauchySeq x Pos.to_nat)
+ (fun n : nat => Qinv (yn (Pos.to_nat k * (Pos.to_nat k * 1) * n)%nat))
+ (CReal_inv_neg yn k cau maj0)) maj)
+ (exist _ (fun n => proj1_sig rneg (Pos.to_nat k * (Pos.to_nat k * 1) * n)%nat) (CReal_linear_shift rneg _ H)))))%Q.
+ apply CRealEq_modindep. apply CRealEq_diff.
+ apply CReal_mult_proper_l. apply CReal_linear_shift_eq.
+ destruct r as [rn limr], rneg as [rnn limneg]; simpl.
+ destruct (QCauchySeq_bounded
+ (fun n : nat => Qinv (rnn (Pos.to_nat k * (Pos.to_nat k * 1) * n)%nat))
+ Pos.to_nat (CReal_inv_neg rnn k limneg maj)).
+ destruct (QCauchySeq_bounded
+ (fun n : nat => rnn (Pos.to_nat k * (Pos.to_nat k * 1) * n)%nat)
+ Pos.to_nat
+ (CReal_linear_shift
+ (exist (fun x0 : nat -> Q => QCauchySeq x0 Pos.to_nat) rnn limneg)
+ (Pos.to_nat k * (Pos.to_nat k * 1)) H)) ; simpl.
+ exists (fun n => 1%nat). intros p n m H0 H1. rewrite Qmult_comm.
+ rewrite Qmult_inv_r. unfold Qminus. rewrite Qplus_opp_r.
+ reflexivity. intro abs.
+ specialize (maj (Pos.to_nat k * (Pos.to_nat k * 1)
+ * (Pos.to_nat (Pos.max x x0)~0 * n))%nat).
+ simpl in maj. rewrite abs in maj. inversion maj.
+ - (* r > 0 *) destruct (CRealPosShift r c) as [[k rneg] [req maj]].
+ simpl in req. apply CRealEq_diff. apply CRealEq_modindep.
+ apply (QSeqEquivEx_trans _
+ (proj1_sig (CReal_mult ((let
+ (yn, cau) as s
+ return ((forall n : nat, 1 # k < proj1_sig s n) -> CReal) := rneg in
+ fun maj0 : forall n : nat, 1 # k < yn n =>
+ exist (fun x : nat -> Q => QCauchySeq x Pos.to_nat)
+ (fun n : nat => Qinv (yn (Pos.to_nat k * (Pos.to_nat k * 1) * n)%nat))
+ (CReal_inv_pos yn k cau maj0)) maj) rneg)))%Q.
+ + apply CRealEq_modindep. apply CRealEq_diff.
+ apply CReal_mult_proper_l. apply req.
+ + assert (le 1 (Pos.to_nat k * (Pos.to_nat k * 1))%nat). rewrite mult_1_r.
+ rewrite <- Pos2Nat.inj_mul. apply Pos2Nat.is_pos.
+ apply (QSeqEquivEx_trans _
+ (proj1_sig (CReal_mult ((let
+ (yn, cau) as s
+ return ((forall n : nat, 1 # k < proj1_sig s n) -> CReal) := rneg in
+ fun maj0 : forall n : nat, 1 # k < yn n =>
+ exist (fun x : nat -> Q => QCauchySeq x Pos.to_nat)
+ (fun n : nat => Qinv (yn (Pos.to_nat k * (Pos.to_nat k * 1) * n)%nat))
+ (CReal_inv_pos yn k cau maj0)) maj)
+ (exist _ (fun n => proj1_sig rneg (Pos.to_nat k * (Pos.to_nat k * 1) * n)%nat) (CReal_linear_shift rneg _ H)))))%Q.
+ apply CRealEq_modindep. apply CRealEq_diff.
+ apply CReal_mult_proper_l. apply CReal_linear_shift_eq.
+ destruct r as [rn limr], rneg as [rnn limneg]; simpl.
+ destruct (QCauchySeq_bounded
+ (fun n : nat => Qinv (rnn (Pos.to_nat k * (Pos.to_nat k * 1) * n)%nat))
+ Pos.to_nat (CReal_inv_pos rnn k limneg maj)).
+ destruct (QCauchySeq_bounded
+ (fun n : nat => rnn (Pos.to_nat k * (Pos.to_nat k * 1) * n)%nat)
+ Pos.to_nat
+ (CReal_linear_shift
+ (exist (fun x0 : nat -> Q => QCauchySeq x0 Pos.to_nat) rnn limneg)
+ (Pos.to_nat k * (Pos.to_nat k * 1)) H)) ; simpl.
+ exists (fun n => 1%nat). intros p n m H0 H1. rewrite Qmult_comm.
+ rewrite Qmult_inv_r. unfold Qminus. rewrite Qplus_opp_r.
+ reflexivity. intro abs.
+ specialize (maj (Pos.to_nat k * (Pos.to_nat k * 1)
+ * (Pos.to_nat (Pos.max x x0)~0 * n))%nat).
+ simpl in maj. rewrite abs in maj. inversion maj.
+Qed.
+
+Lemma CReal_inv_r : forall (r:CReal) (rnz : r # 0),
+ r * ((/ r) rnz) == 1.
+Proof.
+ intros. rewrite CReal_mult_comm, CReal_inv_l.
+ reflexivity.
+Qed.
+
+Lemma CReal_inv_1 : forall nz : 1 # 0, (/ 1) nz == 1.
+Proof.
+ intros. rewrite <- (CReal_mult_1_l ((/1) nz)). rewrite CReal_inv_r.
+ reflexivity.
+Qed.
+
+Lemma CReal_inv_mult_distr :
+ forall r1 r2 (r1nz : r1 # 0) (r2nz : r2 # 0) (rmnz : (r1*r2) # 0),
+ (/ (r1 * r2)) rmnz == (/ r1) r1nz * (/ r2) r2nz.
+Proof.
+ intros. apply (CReal_mult_eq_reg_l r1). exact r1nz.
+ rewrite <- CReal_mult_assoc. rewrite CReal_inv_r. rewrite CReal_mult_1_l.
+ apply (CReal_mult_eq_reg_l r2). exact r2nz.
+ rewrite CReal_inv_r. rewrite <- CReal_mult_assoc.
+ rewrite (CReal_mult_comm r2 r1). rewrite CReal_inv_r.
+ reflexivity.
+Qed.
+
+Lemma Rinv_eq_compat : forall x y (rxnz : x # 0) (rynz : y # 0),
+ x == y
+ -> (/ x) rxnz == (/ y) rynz.
+Proof.
+ intros. apply (CReal_mult_eq_reg_l x). exact rxnz.
+ rewrite CReal_inv_r, H, CReal_inv_r. reflexivity.
+Qed.
+
+Lemma CReal_mult_lt_reg_l : forall r r1 r2, 0 < r -> r * r1 < r * r2 -> r1 < r2.
+Proof.
+ intros z x y H H0.
+ apply (CReal_mult_lt_compat_l ((/z) (inr H))) in H0.
+ repeat rewrite <- CReal_mult_assoc in H0. rewrite CReal_inv_l in H0.
+ repeat rewrite CReal_mult_1_l in H0. apply H0.
+ apply CReal_inv_0_lt_compat. exact H.
+Qed.
+
+Lemma CReal_mult_lt_reg_r : forall r r1 r2, 0 < r -> r1 * r < r2 * r -> r1 < r2.
+Proof.
+ intros.
+ apply CReal_mult_lt_reg_l with r.
+ exact H.
+ now rewrite 2!(CReal_mult_comm r).
+Qed.
+
+Lemma CReal_mult_eq_reg_r : forall r r1 r2, r1 * r == r2 * r -> r # 0 -> r1 == r2.
+Proof.
+ intros. apply (CReal_mult_eq_reg_l r). exact H0.
+ now rewrite 2!(CReal_mult_comm r).
+Qed.
+
+Lemma CReal_mult_eq_compat_l : forall r r1 r2, r1 == r2 -> r * r1 == r * r2.
+Proof.
+ intros. rewrite H. reflexivity.
+Qed.
+
+Lemma CReal_mult_eq_compat_r : forall r r1 r2, r1 == r2 -> r1 * r == r2 * r.
+Proof.
+ intros. rewrite H. reflexivity.
+Qed.
+
+(* In particular x * y == 1 implies that 0 # x, 0 # y and
+ that x and y are inverses of each other. *)
+Lemma CReal_mult_pos_appart_zero : forall x y : CReal, 0 < x * y -> 0 # x.
+Proof.
+ intros. destruct (linear_order_T 0 x 1 (CRealLt_0_1)).
+ left. exact c. destruct (linear_order_T (CReal_opp 1) x 0).
+ rewrite <- CReal_opp_0. apply CReal_opp_gt_lt_contravar, CRealLt_0_1.
+ 2: right; exact c0.
+ pose proof (CRealLt_above _ _ H). destruct H0 as [k kmaj].
+ simpl in kmaj.
+ apply CRealLt_above in c. destruct c as [i imaj]. simpl in imaj.
+ apply CRealLt_above in c0. destruct c0 as [j jmaj]. simpl in jmaj.
+ pose proof (CReal_abs_appart_zero y).
+ destruct x as [xn xcau], y as [yn ycau]. simpl in kmaj.
+ destruct (QCauchySeq_bounded xn Pos.to_nat xcau) as [a amaj],
+ (QCauchySeq_bounded yn Pos.to_nat ycau) as [b bmaj]; simpl in kmaj.
+ clear amaj bmaj. simpl in imaj, jmaj. simpl in H0.
+ specialize (kmaj (Pos.max k (Pos.max i j)) (Pos.le_max_l _ _)).
+ destruct (H0 ((Pos.max a b)~0 * (Pos.max k (Pos.max i j)))%positive).
+ - apply (Qlt_trans _ (2#k)).
+ + unfold Qlt. rewrite <- Z.mul_lt_mono_pos_l. 2: reflexivity.
+ unfold Qden. apply Pos2Z.pos_lt_pos.
+ apply (Pos.le_lt_trans _ (1 * Pos.max k (Pos.max i j))).
+ rewrite Pos.mul_1_l. apply Pos.le_max_l.
+ apply Pos2Nat.inj_lt. do 2 rewrite Pos2Nat.inj_mul.
+ rewrite <- Nat.mul_lt_mono_pos_r. 2: apply Pos2Nat.is_pos.
+ fold (2*Pos.max a b)%positive. rewrite Pos2Nat.inj_mul.
+ apply Nat.lt_1_mul_pos. auto. apply Pos2Nat.is_pos.
+ + apply (Qlt_le_trans _ _ _ kmaj). unfold Qminus. rewrite Qplus_0_r.
+ rewrite <- (Qmult_1_l (Qabs (yn (Pos.to_nat ((Pos.max a b)~0 * Pos.max k (Pos.max i j)))))).
+ apply (Qle_trans _ _ _ (Qle_Qabs _)). rewrite Qabs_Qmult.
+ replace (Pos.to_nat (Pos.max a b)~0 * Pos.to_nat (Pos.max k (Pos.max i j)))%nat
+ with (Pos.to_nat ((Pos.max a b)~0 * Pos.max k (Pos.max i j))).
+ 2: apply Pos2Nat.inj_mul.
+ apply Qmult_le_compat_r. 2: apply Qabs_nonneg.
+ apply Qabs_Qle_condition. split.
+ apply Qlt_le_weak. apply Qlt_minus_iff, (Qlt_trans _ (2#j)).
+ reflexivity. apply jmaj.
+ apply (Pos.le_trans _ (1 * Pos.max k (Pos.max i j))).
+ rewrite Pos.mul_1_l.
+ apply (Pos.le_trans _ (Pos.max i j) _ (Pos.le_max_r _ _)).
+ apply Pos.le_max_r.
+ apply Pos2Nat.inj_le. do 2 rewrite Pos2Nat.inj_mul.
+ rewrite <- Nat.mul_le_mono_pos_r. 2: apply Pos2Nat.is_pos.
+ apply Pos2Nat.is_pos.
+ apply Qlt_le_weak. apply Qlt_minus_iff, (Qlt_trans _ (2#i)).
+ reflexivity. apply imaj.
+ apply (Pos.le_trans _ (1 * Pos.max k (Pos.max i j))).
+ rewrite Pos.mul_1_l.
+ apply (Pos.le_trans _ (Pos.max i j) _ (Pos.le_max_l _ _)).
+ apply Pos.le_max_r.
+ apply Pos2Nat.inj_le. do 2 rewrite Pos2Nat.inj_mul.
+ rewrite <- Nat.mul_le_mono_pos_r. 2: apply Pos2Nat.is_pos.
+ apply Pos2Nat.is_pos.
+ - left. apply (CReal_mult_lt_reg_r (exist _ yn ycau) _ _ c).
+ rewrite CReal_mult_0_l. exact H.
+ - right. apply (CReal_mult_lt_reg_r (CReal_opp (exist _ yn ycau))).
+ rewrite <- CReal_opp_0. apply CReal_opp_gt_lt_contravar. exact c.
+ rewrite CReal_mult_0_l, <- CReal_opp_0, <- CReal_opp_mult_distr_r.
+ apply CReal_opp_gt_lt_contravar. exact H.
+Qed.
+
+Fixpoint pow (r:CReal) (n:nat) : CReal :=
+ match n with
+ | O => 1
+ | S n => r * (pow r n)
+ end.
+
+
+Lemma CReal_mult_le_compat_l_half : forall r r1 r2,
+ 0 < r -> r1 <= r2 -> r * r1 <= r * r2.
+Proof.
+ intros. intro abs. apply (CReal_mult_lt_reg_l) in abs.
+ contradiction. apply H.
+Qed.
+
+Lemma CReal_invQ : forall (b : positive) (pos : Qlt 0 (Z.pos b # 1)),
+ CRealEq (CReal_inv (inject_Q (Z.pos b # 1)) (inr (CReal_injectQPos (Z.pos b # 1) pos)))
+ (inject_Q (1 # b)).
+Proof.
+ intros.
+ apply (CReal_mult_eq_reg_l (inject_Q (Z.pos b # 1))).
+ - right. apply CReal_injectQPos. exact pos.
+ - rewrite CReal_mult_comm, CReal_inv_l.
+ apply CRealEq_diff. intro n. simpl;
+ destruct (QCauchySeq_bounded (fun _ : nat => 1 # b)%Q Pos.to_nat (ConstCauchy (1 # b))),
+ (QCauchySeq_bounded (fun _ : nat => Z.pos b # 1)%Q Pos.to_nat (ConstCauchy (Z.pos b # 1))); simpl.
+ do 2 rewrite Pos.mul_1_r. rewrite Z.pos_sub_diag. discriminate.
+Qed.
+
+Definition CRealQ_dense (a b : CReal)
+ : a < b -> { q : Q & a < inject_Q q < b }.
+Proof.
+ (* Locate a and b at the index given by a<b,
+ and pick the middle rational number. *)
+ intros [p pmaj].
+ exists ((proj1_sig a (Pos.to_nat p) + proj1_sig b (Pos.to_nat p)) * (1#2))%Q.
+ split.
+ - apply (CReal_le_lt_trans _ _ _ (inject_Q_compare a p)). apply inject_Q_lt.
+ apply (Qmult_lt_l _ _ 2). reflexivity.
+ apply (Qplus_lt_l _ _ (-2*proj1_sig a (Pos.to_nat p))).
+ field_simplify. field_simplify in pmaj.
+ setoid_replace (-2#2)%Q with (-1)%Q. 2: reflexivity.
+ setoid_replace (2*(1#p))%Q with (2#p)%Q. 2: reflexivity.
+ rewrite Qplus_comm. exact pmaj.
+ - apply (CReal_plus_lt_reg_l (-b)).
+ rewrite CReal_plus_opp_l.
+ apply (CReal_plus_lt_reg_r
+ (-inject_Q ((proj1_sig a (Pos.to_nat p) + proj1_sig b (Pos.to_nat p)) * (1 # 2)))).
+ rewrite CReal_plus_assoc, CReal_plus_opp_r, CReal_plus_0_r, CReal_plus_0_l.
+ rewrite <- opp_inject_Q.
+ apply (CReal_le_lt_trans _ _ _ (inject_Q_compare (-b) p)). apply inject_Q_lt.
+ apply (Qmult_lt_l _ _ 2). reflexivity.
+ apply (Qplus_lt_l _ _ (2*proj1_sig b (Pos.to_nat p))).
+ destruct b as [bn bcau]; simpl. simpl in pmaj.
+ field_simplify. field_simplify in pmaj.
+ setoid_replace (-2#2)%Q with (-1)%Q. 2: reflexivity.
+ setoid_replace (2*(1#p))%Q with (2#p)%Q. 2: reflexivity. exact pmaj.
+Qed.
+
+Lemma inject_Q_mult : forall q r : Q,
+ inject_Q (q * r) == inject_Q q * inject_Q r.
+Proof.
+ split.
+ - intros [n maj]. simpl in maj.
+ destruct (QCauchySeq_bounded (fun _ : nat => q) Pos.to_nat (ConstCauchy q)).
+ destruct (QCauchySeq_bounded (fun _ : nat => r) Pos.to_nat (ConstCauchy r)).
+ simpl in maj. ring_simplify in maj. discriminate maj.
+ - intros [n maj]. simpl in maj.
+ destruct (QCauchySeq_bounded (fun _ : nat => q) Pos.to_nat (ConstCauchy q)).
+ destruct (QCauchySeq_bounded (fun _ : nat => r) Pos.to_nat (ConstCauchy r)).
+ simpl in maj. ring_simplify in maj. discriminate maj.
+Qed.
diff --git a/theories/Reals/ConstructiveRIneq.v b/theories/Reals/ConstructiveRIneq.v
index b53436be55..e0f08d2fbe 100644
--- a/theories/Reals/ConstructiveRIneq.v
+++ b/theories/Reals/ConstructiveRIneq.v
@@ -22,9 +22,8 @@
constructive reals, do not use ConstructiveCauchyReals
directly. *)
-Require Import ConstructiveCauchyReals.
+Require Import ConstructiveCauchyRealsMult.
Require Import ConstructiveRcomplete.
-Require Import ConstructiveRealsLUB.
Require Export ConstructiveReals.
Require Import Zpower.
Require Export ZArithRing.
@@ -37,11 +36,11 @@ Declare Scope R_scope_constr.
Local Open Scope Z_scope.
Local Open Scope R_scope_constr.
-Definition CR : ConstructiveReals.
+Definition CRealImplem : ConstructiveReals.
Proof.
assert (isLinearOrder CReal CRealLt) as lin.
{ repeat split. exact CRealLt_asym.
- exact CRealLt_trans.
+ exact CReal_lt_trans.
intros. destruct (CRealLt_dec x z y H).
left. exact c. right. exact c. }
apply (Build_ConstructiveReals
@@ -53,30 +52,25 @@ Proof.
CReal_plus_lt_compat_l CReal_plus_lt_reg_l
CReal_mult_lt_0_compat
CReal_inv CReal_inv_l CReal_inv_0_lt_compat
- CRealArchimedean).
+ inject_Q inject_Q_plus inject_Q_mult
+ inject_Q_one inject_Q_lt lt_inject_Q
+ CRealQ_dense Rup_pos).
- intros. destruct (Rcauchy_complete xn) as [l cv].
- intro n. apply (H (IQR (1#n))). apply IQR_pos. reflexivity.
- exists l. intros eps epsPos.
- destruct (Rup_nat ((/eps) (inr epsPos))) as [n nmaj].
- specialize (cv (Pos.of_nat (S n))) as [p pmaj].
- exists p. intros. specialize (pmaj i H0). unfold absSmall in pmaj.
- apply (CReal_mult_lt_compat_l eps) in nmaj.
- rewrite CReal_inv_r, CReal_mult_comm in nmaj.
- 2: apply epsPos. split.
- + apply (CRealLt_trans _ (-IQR (1 # Pos.of_nat (S n)))).
- 2: apply pmaj. clear pmaj.
- apply CReal_opp_gt_lt_contravar. unfold CRealGt, IQR.
- rewrite CReal_mult_1_l. apply (CReal_mult_lt_reg_l (IPR (Pos.of_nat (S n)))).
- apply IPR_pos. rewrite CReal_inv_r, <- INR_IPR, Nat2Pos.id.
- 2: discriminate. apply (CRealLt_trans _ (INR n * eps) _ nmaj).
- apply CReal_mult_lt_compat_r. exact epsPos. apply lt_INR, le_refl.
- + apply (CRealLt_trans _ (IQR (1 # Pos.of_nat (S n)))).
- apply pmaj. unfold IQR. rewrite CReal_mult_1_l.
- apply (CReal_mult_lt_reg_l (IPR (Pos.of_nat (S n)))).
- apply IPR_pos. rewrite CReal_inv_r, <- INR_IPR, Nat2Pos.id.
- 2: discriminate. apply (CRealLt_trans _ (INR n * eps) _ nmaj).
- apply CReal_mult_lt_compat_r. exact epsPos. apply lt_INR, le_refl.
- - exact sig_lub.
+ intro n. destruct (H n). exists x. intros.
+ specialize (a i j H0 H1) as [a b]. split. 2: exact b.
+ rewrite <- opp_inject_Q.
+ setoid_replace (-(1#n))%Q with (-1#n). exact a. reflexivity.
+ exists l. intros p. destruct (cv p).
+ exists x. intros. specialize (a i H0). split. 2: apply a.
+ unfold orderLe.
+ intro abs. setoid_replace (-1#p) with (-(1#p))%Q in abs.
+ rewrite opp_inject_Q in abs. destruct a. contradiction.
+ reflexivity.
+Defined.
+
+Definition CR : ConstructiveReals.
+Proof.
+ exact CRealImplem.
Qed. (* Keep it opaque to possibly change the implementation later *)
Definition R := CRcarrier CR.
@@ -1673,6 +1667,19 @@ Proof.
intro; destruct n. rewrite Rplus_0_l. reflexivity. reflexivity.
Qed.
+(**********)
+Lemma IZN : forall n:Z, (0 <= n)%Z -> { m : nat | n = Z.of_nat m }.
+Proof.
+ intros. exists (Z.to_nat n). rewrite Z2Nat.id. reflexivity. assumption.
+Qed.
+
+Lemma le_succ_r_T : forall n m : nat, (n <= S m)%nat -> {(n <= m)%nat} + {n = S m}.
+Proof.
+ intros. destruct (le_lt_dec n m). left. exact l.
+ right. apply Nat.le_succ_r in H. destruct H.
+ exfalso. apply (le_not_lt n m); assumption. exact H.
+Qed.
+
Lemma lt_INR : forall n m:nat, (n < m)%nat -> INR n < INR m.
Proof.
induction m.
@@ -2174,35 +2181,29 @@ Proof.
contradiction. apply H.
Qed.
-Lemma INR_gen_phiZ : forall (n : nat),
- gen_phiZ 0 1 Rplus Rmult Ropp (Z.of_nat n) == INR n.
+Lemma INR_CR_of_Q : forall (n : nat),
+ CR_of_Q CR (Z.of_nat n # 1) == INR n.
Proof.
induction n.
- - apply Req_refl.
- - replace (Z.of_nat (S n)) with (1 + Z.of_nat n)%Z.
- rewrite (gen_phiZ_add Req_rel (CRisRingExt CR) RisRing).
- rewrite IHn. clear IHn. simpl. rewrite (Rplus_comm 1).
- destruct n. rewrite Rplus_0_l. reflexivity. reflexivity.
+ - apply CR_of_Q_zero.
+ - transitivity (CR_of_Q CR (1 + (Z.of_nat n # 1))).
replace (S n) with (1 + n)%nat. 2: reflexivity.
- rewrite (Nat2Z.inj_add 1 n). reflexivity.
+ rewrite (Nat2Z.inj_add 1 n).
+ apply CR_of_Q_proper.
+ rewrite <- (Qinv_plus_distr (Z.of_nat 1) (Z.of_nat n) 1). reflexivity.
+ rewrite CR_of_Q_plus. rewrite IHn. clear IHn.
+ setoid_replace (INR (S n)) with (1 + INR n).
+ rewrite CR_of_Q_one. reflexivity.
+ simpl. destruct n. rewrite Rplus_0_r. reflexivity.
+ rewrite Rplus_comm. reflexivity.
Qed.
Definition Rup_nat (x : R)
: { n : nat & x < INR n }.
Proof.
- intros. destruct (CRarchimedean CR x) as [p maj].
- destruct p.
- - exists O. apply maj.
- - exists (Pos.to_nat p).
- rewrite <- positive_nat_Z, (INR_gen_phiZ (Pos.to_nat p)) in maj. exact maj.
- - exists O. apply (Rlt_trans _ _ _ maj). simpl.
- rewrite <- Ropp_0. apply Ropp_gt_lt_contravar.
- fold (gen_phiZ 0 1 Rplus Rmult Ropp (Z.pos p)).
- replace (gen_phiPOS 1 (CRplus CR) (CRmult CR) p)
- with (gen_phiZ 0 1 Rplus Rmult Ropp (Z.pos p)).
- 2: reflexivity.
- rewrite <- positive_nat_Z, (INR_gen_phiZ (Pos.to_nat p)).
- apply (lt_INR 0). apply Pos2Nat.is_pos.
+ intros. destruct (CR_archimedean CR x) as [p maj].
+ exists (Pos.to_nat p).
+ rewrite <- INR_CR_of_Q, positive_nat_Z. exact maj.
Qed.
Fixpoint Rarchimedean_ind (x:R) (n : Z) (p:nat) { struct p }
diff --git a/theories/Reals/ConstructiveRcomplete.v b/theories/Reals/ConstructiveRcomplete.v
index ce45bcd567..0a515672f2 100644
--- a/theories/Reals/ConstructiveRcomplete.v
+++ b/theories/Reals/ConstructiveRcomplete.v
@@ -11,227 +11,145 @@
Require Import QArith_base.
Require Import Qabs.
-Require Import ConstructiveCauchyReals.
+Require Import ConstructiveCauchyRealsMult.
Require Import Logic.ConstructiveEpsilon.
Local Open Scope CReal_scope.
+Definition absLe (a b : CReal) : Prop
+ := -b <= a <= b.
+
Lemma CReal_absSmall : forall (x y : CReal) (n : positive),
(Qlt (2 # n)
(proj1_sig x (Pos.to_nat n) - Qabs (proj1_sig y (Pos.to_nat n))))
- -> (CRealLt (CReal_opp x) y * CRealLt y x).
+ -> absLe y x.
Proof.
intros x y n maj. split.
- - exists n. destruct x as [xn caux], y as [yn cauy]; simpl.
+ - apply CRealLt_asym. exists n. destruct x as [xn caux], y as [yn cauy]; simpl.
simpl in maj. unfold Qminus. rewrite Qopp_involutive.
rewrite Qplus_comm.
apply (Qlt_le_trans _ (xn (Pos.to_nat n) - Qabs (yn (Pos.to_nat n)))).
apply maj. apply Qplus_le_r.
rewrite <- (Qopp_involutive (yn (Pos.to_nat n))).
apply Qopp_le_compat. rewrite Qabs_opp. apply Qle_Qabs.
- - exists n. destruct x as [xn caux], y as [yn cauy]; simpl.
+ - apply CRealLt_asym. exists n. destruct x as [xn caux], y as [yn cauy]; simpl.
simpl in maj.
apply (Qlt_le_trans _ (xn (Pos.to_nat n) - Qabs (yn (Pos.to_nat n)))).
apply maj. apply Qplus_le_r. apply Qopp_le_compat. apply Qle_Qabs.
Qed.
-Definition absSmall (a b : CReal) : Set
- := -b < a < b.
-
+(* We use absLe in sort Prop rather than Set,
+ to extract smaller programs. *)
Definition Un_cv_mod (un : nat -> CReal) (l : CReal) : Set
- := forall n : positive,
- { p : nat & forall i:nat, le p i -> absSmall (un i - l) (IQR (1#n)) }.
+ := forall p : positive,
+ { n : nat | forall i:nat, le n i -> absLe (un i - l) (inject_Q (1#p)) }.
Lemma Un_cv_mod_eq : forall (v u : nat -> CReal) (s : CReal),
(forall n:nat, u n == v n)
- -> Un_cv_mod u s -> Un_cv_mod v s.
+ -> Un_cv_mod u s
+ -> Un_cv_mod v s.
Proof.
intros v u s seq H1 p. specialize (H1 p) as [N H0].
- exists N. intros. unfold absSmall. split.
+ exists N. intros. split.
rewrite <- seq. apply H0. apply H.
rewrite <- seq. apply H0. apply H.
Qed.
Definition Un_cauchy_mod (un : nat -> CReal) : Set
- := forall n : positive,
- { p : nat & forall i j:nat, le p i
- -> le p j
- -> -IQR (1#n) < un i - un j < IQR (1#n) }.
+ := forall p : positive,
+ { n : nat | forall i j:nat, le n i -> le n j
+ -> absLe (un i - un j) (inject_Q (1#p)) }.
(* Sharpen the archimedean property : constructive versions of
- the usual floor and ceiling functions.
-
- n is a temporary parameter used for the recursion,
- look at Ffloor below. *)
-Fixpoint Rfloor_pos (a : CReal) (n : nat) { struct n }
- : 0 < a
- -> a < INR n
- -> { p : nat & INR p < a < INR p + 2 }.
-Proof.
- (* Decreasing loop on n, until it is the first integer above a. *)
- intros H H0. destruct n.
- - exfalso. apply (CRealLt_asym 0 a); assumption.
- - destruct n as [|p] eqn:des.
- + (* n = 1 *) exists O. split.
- apply H. rewrite CReal_plus_0_l. apply (CRealLt_trans a (1+0)).
- rewrite CReal_plus_comm, CReal_plus_0_l. apply H0.
- apply CReal_plus_le_lt_compat.
- apply CRealLe_refl. apply CRealLt_0_1.
- + (* n > 1 *)
- destruct (linear_order_T (INR p) a (INR (S p))).
- * rewrite <- CReal_plus_0_l, S_INR, CReal_plus_comm. apply CReal_plus_lt_compat_l.
- apply CRealLt_0_1.
- * exists p. split. exact c.
- rewrite S_INR, S_INR, CReal_plus_assoc in H0. exact H0.
- * apply (Rfloor_pos a n H). rewrite des. apply c.
-Qed.
-
+ the usual floor and ceiling functions. *)
Definition Rfloor (a : CReal)
- : { p : Z & IZR p < a < IZR p + 2 }.
+ : { p : Z & inject_Q (p#1) < a < inject_Q (p#1) + 2 }.
Proof.
- assert (forall x:CReal, 0 < x -> { n : nat & x < INR n }).
- { intros. pose proof (Rarchimedean x) as [n [maj _]].
- destruct n.
- + exfalso. apply (CRealLt_asym 0 x); assumption.
- + exists (Pos.to_nat p). rewrite INR_IPR. apply maj.
- + exfalso. apply (CRealLt_asym 0 x). apply H.
- apply (CRealLt_trans x (IZR (Z.neg p))). apply maj.
- apply (CReal_plus_lt_reg_l (-IZR (Z.neg p))).
- rewrite CReal_plus_comm, CReal_plus_opp_r. rewrite <- opp_IZR.
- rewrite CReal_plus_comm, CReal_plus_0_l.
- apply (IZR_lt 0). reflexivity. }
- destruct (linear_order_T 0 a 1 CRealLt_0_1).
- - destruct (H a c). destruct (Rfloor_pos a x c c0).
- exists (Z.of_nat x0). split; rewrite <- INR_IZR_INZ; apply p.
- - apply (CReal_plus_lt_compat_l (-a)) in c.
- rewrite CReal_plus_comm, CReal_plus_opp_r, CReal_plus_comm in c.
- destruct (H (1-a) c).
- destruct (Rfloor_pos (1-a) x c c0).
- exists (-(Z.of_nat x0 + 1))%Z. split; rewrite opp_IZR, plus_IZR.
- + rewrite <- (CReal_opp_involutive a). apply CReal_opp_gt_lt_contravar.
- destruct p as [_ a0]. apply (CReal_plus_lt_reg_r 1).
- rewrite CReal_plus_comm, CReal_plus_assoc. rewrite <- INR_IZR_INZ. apply a0.
- + destruct p as [a0 _]. apply (CReal_plus_lt_compat_l a) in a0.
- unfold CReal_minus in a0.
- rewrite <- (CReal_plus_comm (1+-a)), CReal_plus_assoc, CReal_plus_opp_l, CReal_plus_0_r in a0.
- rewrite <- INR_IZR_INZ.
- apply (CReal_plus_lt_reg_r (INR x0)). unfold IZR, IPR, IPR_2.
- ring_simplify. exact a0.
-Qed.
+ destruct (CRealArchimedean a) as [n [H H0]].
+ exists (n-2)%Z. split.
+ - setoid_replace (n - 2 # 1)%Q with ((n#1) + - 2)%Q.
+ rewrite inject_Q_plus, (opp_inject_Q 2).
+ apply (CReal_plus_lt_reg_r 2). ring_simplify.
+ rewrite CReal_plus_comm. exact H0.
+ rewrite Qinv_plus_distr. reflexivity.
+ - setoid_replace (n - 2 # 1)%Q with ((n#1) + - 2)%Q.
+ rewrite inject_Q_plus, (opp_inject_Q 2).
+ ring_simplify. exact H.
+ rewrite Qinv_plus_distr. reflexivity.
+Defined.
-Definition Rup_nat (x : CReal)
- : { n : nat & x < INR n }.
-Proof.
- intros. destruct (Rarchimedean x) as [p [maj _]].
- destruct p.
- - exists O. apply maj.
- - exists (Pos.to_nat p). rewrite INR_IPR. apply maj.
- - exists O. apply (CRealLt_trans _ (IZR (Z.neg p)) _ maj).
- apply (IZR_lt _ 0). reflexivity.
-Qed.
(* A point in an archimedean field is the limit of a
sequence of rational numbers (n maps to the q between
a and a+1/n). This will yield a maximum
archimedean field, which is the field of real numbers. *)
-Definition FQ_dense_pos (a b : CReal)
- : 0 < b
- -> a < b -> { q : Q & a < IQR q < b }.
-Proof.
- intros H H0.
- assert (0 < b - a) as epsPos.
- { apply (CReal_plus_lt_compat_l (-a)) in H0.
- rewrite CReal_plus_opp_l, CReal_plus_comm in H0.
- apply H0. }
- pose proof (Rup_nat ((/(b-a)) (inr epsPos)))
- as [n maj].
- destruct n as [|k].
- - exfalso.
- apply (CReal_mult_lt_compat_l (b-a)) in maj. 2: apply epsPos.
- rewrite CReal_mult_0_r in maj. rewrite CReal_inv_r in maj.
- apply (CRealLt_asym 0 1). apply CRealLt_0_1. apply maj.
- - (* 0 < n *)
- pose (Pos.of_nat (S k)) as n.
- destruct (Rfloor (IZR (2 * Z.pos n) * b)) as [p maj2].
- exists (p # (2*n))%Q. split.
- + apply (CRealLt_trans a (b - IQR (1 # n))).
- apply (CReal_plus_lt_reg_r (IQR (1#n))).
- unfold CReal_minus. rewrite CReal_plus_assoc. rewrite CReal_plus_opp_l.
- rewrite CReal_plus_0_r. apply (CReal_plus_lt_reg_l (-a)).
- rewrite <- CReal_plus_assoc, CReal_plus_opp_l, CReal_plus_0_l.
- rewrite CReal_plus_comm. unfold IQR.
- rewrite CReal_mult_1_l. apply (CReal_mult_lt_reg_l (IPR n)).
- apply IPR_pos. rewrite CReal_inv_r.
- apply (CReal_mult_lt_compat_l (b-a)) in maj.
- rewrite CReal_inv_r, CReal_mult_comm in maj.
- rewrite <- INR_IPR. unfold n. rewrite Nat2Pos.id.
- apply maj. discriminate. exact epsPos.
- apply (CReal_plus_lt_reg_r (IQR (1 # n))).
- unfold CReal_minus. rewrite CReal_plus_assoc, CReal_plus_opp_l.
- rewrite CReal_plus_0_r. rewrite <- plus_IQR.
- destruct maj2 as [_ maj2].
- setoid_replace ((p # 2 * n) + (1 # n))%Q
- with ((p + 2 # 2 * n))%Q. unfold IQR.
- apply (CReal_mult_lt_reg_r (IZR (Z.pos (2 * n)))).
- apply (IZR_lt 0). reflexivity. rewrite CReal_mult_assoc.
- rewrite CReal_inv_l. rewrite CReal_mult_1_r. rewrite CReal_mult_comm.
- rewrite plus_IZR. apply maj2.
- setoid_replace (1#n)%Q with (2#2*n)%Q. 2: reflexivity.
- apply Qinv_plus_distr.
- + destruct maj2 as [maj2 _]. unfold IQR.
- apply (CReal_mult_lt_reg_r (IZR (Z.pos (2 * n)))).
- apply (IZR_lt 0). apply Pos2Z.is_pos. rewrite CReal_mult_assoc, CReal_inv_l.
- rewrite CReal_mult_1_r, CReal_mult_comm. apply maj2.
-Qed.
-
Definition FQ_dense (a b : CReal)
- : a < b
- -> { q : Q & a < IQR q < b }.
+ : a < b -> { q : Q & a < inject_Q q < b }.
Proof.
- intros H. destruct (linear_order_T a 0 b). apply H.
- - destruct (FQ_dense_pos (-b) (-a)) as [q maj].
- apply (CReal_plus_lt_compat_l (-a)) in c. rewrite CReal_plus_opp_l in c.
- rewrite CReal_plus_0_r in c. apply c.
- apply (CReal_plus_lt_compat_l (-a)) in H.
+ intros H. assert (0 < b - a) as epsPos.
+ { apply (CReal_plus_lt_compat_l (-a)) in H.
rewrite CReal_plus_opp_l, CReal_plus_comm in H.
- apply (CReal_plus_lt_compat_l (-b)) in H. rewrite <- CReal_plus_assoc in H.
- rewrite CReal_plus_opp_l in H. rewrite CReal_plus_0_l in H.
- rewrite CReal_plus_0_r in H. apply H.
- exists (-q)%Q. split.
- + destruct maj as [_ maj].
- apply (CReal_plus_lt_compat_l (-IQR q)) in maj.
- rewrite CReal_plus_opp_l, <- opp_IQR, CReal_plus_comm in maj.
- apply (CReal_plus_lt_compat_l a) in maj. rewrite <- CReal_plus_assoc in maj.
- rewrite CReal_plus_opp_r, CReal_plus_0_l in maj.
- rewrite CReal_plus_0_r in maj. apply maj.
- + destruct maj as [maj _].
- apply (CReal_plus_lt_compat_l (-IQR q)) in maj.
- rewrite CReal_plus_opp_l, <- opp_IQR, CReal_plus_comm in maj.
- apply (CReal_plus_lt_compat_l b) in maj. rewrite <- CReal_plus_assoc in maj.
- rewrite CReal_plus_opp_r in maj. rewrite CReal_plus_0_l in maj.
- rewrite CReal_plus_0_r in maj. apply maj.
- - apply FQ_dense_pos. apply c. apply H.
+ apply H. }
+ pose proof (Rup_pos ((/(b-a)) (inr epsPos)))
+ as [n maj].
+ destruct (Rfloor (inject_Q (2 * Z.pos n # 1) * b)) as [p maj2].
+ exists (p # (2*n))%Q. split.
+ - apply (CReal_lt_trans a (b - inject_Q (1 # n))).
+ apply (CReal_plus_lt_reg_r (inject_Q (1#n))).
+ unfold CReal_minus. rewrite CReal_plus_assoc. rewrite CReal_plus_opp_l.
+ rewrite CReal_plus_0_r. apply (CReal_plus_lt_reg_l (-a)).
+ rewrite <- CReal_plus_assoc, CReal_plus_opp_l, CReal_plus_0_l.
+ rewrite CReal_plus_comm.
+ apply (CReal_mult_lt_reg_l (inject_Q (Z.pos n # 1))).
+ apply inject_Q_lt. reflexivity. rewrite <- inject_Q_mult.
+ setoid_replace ((Z.pos n # 1) * (1 # n))%Q with 1%Q.
+ apply (CReal_mult_lt_compat_l (b-a)) in maj.
+ rewrite CReal_inv_r, CReal_mult_comm in maj. exact maj.
+ exact epsPos. unfold Qeq; simpl. do 2 rewrite Pos.mul_1_r. reflexivity.
+ apply (CReal_plus_lt_reg_r (inject_Q (1 # n))).
+ unfold CReal_minus. rewrite CReal_plus_assoc, CReal_plus_opp_l.
+ rewrite CReal_plus_0_r. rewrite <- inject_Q_plus.
+ destruct maj2 as [_ maj2].
+ setoid_replace ((p # 2 * n) + (1 # n))%Q
+ with ((p + 2 # 2 * n))%Q.
+ apply (CReal_mult_lt_reg_r (inject_Q (Z.pos (2 * n) # 1))).
+ apply inject_Q_lt. reflexivity. rewrite <- inject_Q_mult.
+ setoid_replace ((p + 2 # 2 * n) * (Z.pos (2 * n) # 1))%Q
+ with ((p#1) + 2)%Q.
+ rewrite inject_Q_plus. rewrite Pos2Z.inj_mul.
+ rewrite CReal_mult_comm. exact maj2.
+ unfold Qeq; simpl. rewrite Pos.mul_1_r, Z.mul_1_r. ring.
+ setoid_replace (1#n)%Q with (2#2*n)%Q. 2: reflexivity.
+ apply Qinv_plus_distr.
+ - destruct maj2 as [maj2 _].
+ apply (CReal_mult_lt_reg_r (inject_Q (Z.pos (2 * n) # 1))).
+ apply inject_Q_lt. reflexivity.
+ rewrite <- inject_Q_mult.
+ setoid_replace ((p # 2 * n) * (Z.pos (2 * n) # 1))%Q
+ with ((p#1))%Q.
+ rewrite CReal_mult_comm. exact maj2.
+ unfold Qeq; simpl. rewrite Pos.mul_1_r, Z.mul_1_r. reflexivity.
Qed.
Definition RQ_limit : forall (x : CReal) (n:nat),
- { q:Q & x < IQR q < x + IQR (1 # Pos.of_nat n) }.
+ { q:Q & x < inject_Q q < x + inject_Q (1 # Pos.of_nat n) }.
Proof.
- intros x n. apply (FQ_dense x (x + IQR (1 # Pos.of_nat n))).
+ intros x n. apply (FQ_dense x (x + inject_Q (1 # Pos.of_nat n))).
rewrite <- (CReal_plus_0_r x). rewrite CReal_plus_assoc.
- apply CReal_plus_lt_compat_l. rewrite CReal_plus_0_l. apply IQR_pos.
+ apply CReal_plus_lt_compat_l. rewrite CReal_plus_0_l. apply inject_Q_lt.
reflexivity.
Qed.
Definition Un_cauchy_Q (xn : nat -> Q) : Set
:= forall n : positive,
{ k : nat | forall p q : nat, le k p -> le k q
- -> Qlt (-(1#n)) (xn p - xn q)
- /\ Qlt (xn p - xn q) (1#n) }.
+ -> Qle (-(1#n)) (xn p - xn q)
+ /\ Qle (xn p - xn q) (1#n) }.
Lemma Rdiag_cauchy_sequence : forall (xn : nat -> CReal),
Un_cauchy_mod xn
- -> Un_cauchy_Q (fun n => let (l,_) := RQ_limit (xn n) n in l).
+ -> Un_cauchy_Q (fun n:nat => let (l,_) := RQ_limit (xn n) n in l).
Proof.
intros xn H p. specialize (H (2 * p)%positive) as [k cv].
exists (max k (2 * Pos.to_nat p)). intros.
@@ -241,67 +159,69 @@ Proof.
apply (le_trans _ (Init.Nat.max k (2 * Pos.to_nat p))).
apply Nat.le_max_l. apply H0.
split.
- - apply lt_IQR. unfold Qminus.
- apply (CRealLt_trans _ (xn p0 - (xn q + IQR (1 # 2 * p)))).
- + unfold CReal_minus. rewrite CReal_opp_plus_distr. unfold CReal_minus.
+ - apply le_inject_Q. unfold Qminus.
+ apply (CReal_le_trans _ (xn p0 - (xn q + inject_Q (1 # 2 * p)))).
+ + unfold CReal_minus. rewrite CReal_opp_plus_distr.
rewrite <- CReal_plus_assoc.
- apply (CReal_plus_lt_reg_r (IQR (1 # 2 * p))).
+ apply (CReal_plus_le_reg_r (inject_Q (1 # 2 * p))).
rewrite CReal_plus_assoc. rewrite CReal_plus_opp_l. rewrite CReal_plus_0_r.
- rewrite <- plus_IQR.
+ rewrite <- inject_Q_plus.
setoid_replace (- (1 # p) + (1 # 2 * p))%Q with (- (1 # 2 * p))%Q.
- rewrite opp_IQR. exact c.
+ rewrite opp_inject_Q. exact H1.
rewrite Qplus_comm.
setoid_replace (1#p)%Q with (2 # 2 *p)%Q. rewrite Qinv_minus_distr.
reflexivity. reflexivity.
- + rewrite plus_IQR. apply CReal_plus_le_lt_compat.
+ + rewrite inject_Q_plus. apply CReal_plus_le_compat.
apply CRealLt_asym.
destruct (RQ_limit (xn p0) p0); simpl. apply p1.
+ apply CRealLt_asym.
destruct (RQ_limit (xn q) q); unfold proj1_sig.
- rewrite opp_IQR. apply CReal_opp_gt_lt_contravar.
- apply (CRealLt_Le_trans _ (xn q + IQR (1 # Pos.of_nat q))).
- apply p1. apply CReal_plus_le_compat_l. apply IQR_le.
+ rewrite opp_inject_Q. apply CReal_opp_gt_lt_contravar.
+ apply (CReal_lt_le_trans _ (xn q + inject_Q (1 # Pos.of_nat q))).
+ apply p1. apply CReal_plus_le_compat_l. apply inject_Q_le.
apply Z2Nat.inj_le. discriminate. discriminate.
simpl. assert ((Pos.to_nat p~0 <= q)%nat).
{ apply (le_trans _ (Init.Nat.max k (2 * Pos.to_nat p))).
2: apply H0. replace (p~0)%positive with (2*p)%positive.
2: reflexivity. rewrite Pos2Nat.inj_mul.
apply Nat.le_max_r. }
- rewrite Nat2Pos.id. apply H1. intro abs. subst q.
- inversion H1. pose proof (Pos2Nat.is_pos (p~0)).
- rewrite H3 in H2. inversion H2.
- - apply lt_IQR. unfold Qminus.
- apply (CRealLt_trans _ (xn p0 + IQR (1 # 2 * p) - xn q)).
- + rewrite plus_IQR. apply CReal_plus_le_lt_compat.
+ rewrite Nat2Pos.id. apply H3. intro abs. subst q.
+ inversion H3. pose proof (Pos2Nat.is_pos (p~0)).
+ rewrite H5 in H4. inversion H4.
+ - apply le_inject_Q. unfold Qminus.
+ apply (CReal_le_trans _ (xn p0 + inject_Q (1 # 2 * p) - xn q)).
+ + rewrite inject_Q_plus. apply CReal_plus_le_compat.
apply CRealLt_asym.
destruct (RQ_limit (xn p0) p0); unfold proj1_sig.
- apply (CRealLt_Le_trans _ (xn p0 + IQR (1 # Pos.of_nat p0))).
- apply p1. apply CReal_plus_le_compat_l. apply IQR_le.
+ apply (CReal_lt_le_trans _ (xn p0 + inject_Q (1 # Pos.of_nat p0))).
+ apply p1. apply CReal_plus_le_compat_l. apply inject_Q_le.
apply Z2Nat.inj_le. discriminate. discriminate.
simpl. assert ((Pos.to_nat p~0 <= p0)%nat).
{ apply (le_trans _ (Init.Nat.max k (2 * Pos.to_nat p))).
2: apply H. replace (p~0)%positive with (2*p)%positive.
2: reflexivity. rewrite Pos2Nat.inj_mul.
apply Nat.le_max_r. }
- rewrite Nat2Pos.id. apply H1. intro abs. subst p0.
- inversion H1. pose proof (Pos2Nat.is_pos (p~0)).
- rewrite H3 in H2. inversion H2.
- rewrite opp_IQR. apply CReal_opp_gt_lt_contravar.
+ rewrite Nat2Pos.id. apply H3. intro abs. subst p0.
+ inversion H3. pose proof (Pos2Nat.is_pos (p~0)).
+ rewrite H5 in H4. inversion H4.
+ apply CRealLt_asym.
+ rewrite opp_inject_Q. apply CReal_opp_gt_lt_contravar.
destruct (RQ_limit (xn q) q); simpl. apply p1.
+ unfold CReal_minus. rewrite (CReal_plus_comm (xn p0)).
rewrite CReal_plus_assoc.
- apply (CReal_plus_lt_reg_l (- IQR (1 # 2 * p))).
+ apply (CReal_plus_le_reg_l (- inject_Q (1 # 2 * p))).
rewrite <- CReal_plus_assoc. rewrite CReal_plus_opp_l. rewrite CReal_plus_0_l.
- rewrite <- opp_IQR. rewrite <- plus_IQR.
+ rewrite <- opp_inject_Q. rewrite <- inject_Q_plus.
setoid_replace (- (1 # 2 * p) + (1 # p))%Q with (1 # 2 * p)%Q.
- exact c0. rewrite Qplus_comm.
+ exact H2. rewrite Qplus_comm.
setoid_replace (1#p)%Q with (2 # 2*p)%Q. rewrite Qinv_minus_distr.
reflexivity. reflexivity.
Qed.
-Lemma doubleLtCovariant : forall a b c d e f : CReal,
+Lemma doubleLeCovariant : forall a b c d e f : CReal,
a == b -> c == d -> e == f
- -> (a < c < e)
- -> (b < d < f).
+ -> (a <= c <= e)
+ -> (b <= d <= f).
Proof.
split. rewrite <- H. rewrite <- H0. apply H2.
rewrite <- H0. rewrite <- H1. apply H2.
@@ -311,15 +231,13 @@ Qed.
show that it converges to itself in CReal. *)
Lemma CReal_cv_self : forall (qn : nat -> Q) (x : CReal) (cvmod : positive -> nat),
QSeqEquiv qn (fun n => proj1_sig x n) cvmod
- -> Un_cv_mod (fun n => IQR (qn n)) x.
+ -> Un_cv_mod (fun n => inject_Q (qn n)) x.
Proof.
intros qn x cvmod H p.
specialize (H (2*p)%positive). exists (cvmod (2*p)%positive).
- intros p0 H0. unfold absSmall, CReal_minus.
- apply (doubleLtCovariant (-inject_Q (1#p)) _ (inject_Q (qn p0) - x) _ (inject_Q (1#p))).
- rewrite FinjectQ_CReal. reflexivity.
- rewrite FinjectQ_CReal. reflexivity.
- rewrite FinjectQ_CReal. reflexivity.
+ intros p0 H0. unfold absLe, CReal_minus.
+ apply (doubleLeCovariant (-inject_Q (1#p)) _ (inject_Q (qn p0) - x) _ (inject_Q (1#p))).
+ reflexivity. reflexivity. reflexivity.
apply (CReal_absSmall _ _ (Pos.max (4 * p)%positive (Pos.of_nat (cvmod (2 * p)%positive)))).
setoid_replace (proj1_sig (inject_Q (1 # p)) (Pos.to_nat (Pos.max (4 * p) (Pos.of_nat (cvmod (2 * p)%positive)))))
with (1 # p)%Q.
@@ -353,7 +271,7 @@ Lemma Un_cv_extens : forall (xn yn : nat -> CReal) (l : CReal),
-> Un_cv_mod yn l.
Proof.
intros. intro p. destruct (H p) as [n cv]. exists n.
- intros. unfold absSmall, CReal_minus.
+ intros. unfold absLe, CReal_minus.
split; rewrite <- (H0 i); apply cv; apply H1.
Qed.
@@ -362,29 +280,28 @@ Qed.
The biggest computable such field has all rational limits. *)
Lemma R_has_all_rational_limits : forall qn : nat -> Q,
Un_cauchy_Q qn
- -> { r : CReal & Un_cv_mod (fun n => IQR (qn n)) r }.
+ -> { r : CReal & Un_cv_mod (fun n:nat => inject_Q (qn n)) r }.
Proof.
- (* qn is an element of CReal. Show that IQR qn
+ (* qn is an element of CReal. Show that inject_Q qn
converges to it in CReal. *)
intros.
- destruct (standard_modulus qn (fun p => proj1_sig (H p))).
- - intros p n k H0 H1. destruct (H p); simpl in H0,H1.
- specialize (a n k H0 H1). apply Qabs_case.
- intros _. apply a. intros _.
- apply (Qplus_lt_r _ _ (qn n -qn k-(1#p))). ring_simplify.
- destruct a. ring_simplify in H2. exact H2.
+ destruct (standard_modulus qn (fun p => proj1_sig (H (Pos.succ p)))).
+ - intros p n k H0 H1. destruct (H (Pos.succ p)%positive) as [x a]; simpl in H0,H1.
+ specialize (a n k H0 H1).
+ apply (Qle_lt_trans _ (1#Pos.succ p)).
+ apply Qabs_Qle_condition. exact a.
+ apply Pos2Z.pos_lt_pos. simpl. apply Pos.lt_succ_diag_r.
- exists (exist _ (fun n : nat =>
- qn (increasing_modulus (fun p : positive => proj1_sig (H p)) n)) H0).
- apply (Un_cv_extens (fun n : nat => IQR (qn n))).
+ qn (increasing_modulus (fun p : positive => proj1_sig (H (Pos.succ p))) n)) H0).
apply (CReal_cv_self qn (exist _ (fun n : nat =>
- qn (increasing_modulus (fun p : positive => proj1_sig (H p)) n)) H0)
- (fun p : positive => Init.Nat.max (proj1_sig (H p)) (Pos.to_nat p))).
- apply H1. intro n. reflexivity.
+ qn (increasing_modulus (fun p : positive => proj1_sig (H (Pos.succ p))) n)) H0)
+ (fun p : positive => Init.Nat.max (proj1_sig (H (Pos.succ p))) (Pos.to_nat p))).
+ apply H1.
Qed.
Lemma Rcauchy_complete : forall (xn : nat -> CReal),
Un_cauchy_mod xn
- -> { l : CReal & Un_cv_mod xn l }.
+ -> { l : CReal & Un_cv_mod xn l }.
Proof.
intros xn cau.
destruct (R_has_all_rational_limits (fun n => let (l,_) := RQ_limit (xn n) n in l)
@@ -396,21 +313,21 @@ Proof.
apply Nat.le_max_l. apply H.
destruct (RQ_limit (xn p0) p0) as [q maj]; unfold proj1_sig in H0,H1.
split.
- - apply (CRealLt_trans _ (IQR q - IQR (1 # 2 * p) - l)).
- + unfold CReal_minus. rewrite (CReal_plus_comm (IQR q)).
- apply (CReal_plus_lt_reg_l (IQR (1 # 2 * p))).
- ring_simplify. unfold CReal_minus. rewrite <- opp_IQR. rewrite <- plus_IQR.
+ - apply (CReal_le_trans _ (inject_Q q - inject_Q (1 # 2 * p) - l)).
+ + unfold CReal_minus. rewrite (CReal_plus_comm (inject_Q q)).
+ apply (CReal_plus_le_reg_l (inject_Q (1 # 2 * p))).
+ ring_simplify. unfold CReal_minus. rewrite <- opp_inject_Q. rewrite <- inject_Q_plus.
setoid_replace ((1 # 2 * p) + - (1 # p))%Q with (-(1#2*p))%Q.
- rewrite opp_IQR. apply H0.
+ rewrite opp_inject_Q. apply H0.
setoid_replace (1#p)%Q with (2 # 2*p)%Q.
rewrite Qinv_minus_distr. reflexivity. reflexivity.
+ unfold CReal_minus.
- do 2 rewrite <- (CReal_plus_comm (-l)). apply CReal_plus_lt_compat_l.
- apply (CReal_plus_lt_reg_r (IQR (1 # 2 * p))).
+ do 2 rewrite <- (CReal_plus_comm (-l)). apply CReal_plus_le_compat_l.
+ apply (CReal_plus_le_reg_r (inject_Q (1 # 2 * p))).
ring_simplify. rewrite CReal_plus_comm.
- apply (CRealLt_Le_trans _ (xn p0 + IQR (1 # Pos.of_nat p0))).
- apply maj. apply CReal_plus_le_compat_l.
- apply IQR_le.
+ apply (CReal_le_trans _ (xn p0 + inject_Q (1 # Pos.of_nat p0))).
+ apply CRealLt_asym, maj. apply CReal_plus_le_compat_l.
+ apply inject_Q_le.
apply Z2Nat.inj_le. discriminate. discriminate.
simpl. assert ((Pos.to_nat p~0 <= p0)%nat).
{ apply (le_trans _ (Init.Nat.max k (2 * Pos.to_nat p))).
@@ -420,13 +337,13 @@ Proof.
rewrite Nat2Pos.id. apply H2. intro abs. subst p0.
inversion H2. pose proof (Pos2Nat.is_pos (p~0)).
rewrite H4 in H3. inversion H3.
- - apply (CRealLt_trans _ (IQR q - l)).
+ - apply (CReal_le_trans _ (inject_Q q - l)).
+ unfold CReal_minus. do 2 rewrite <- (CReal_plus_comm (-l)).
- apply CReal_plus_lt_compat_l. apply maj.
- + apply (CRealLt_trans _ (IQR (1 # 2 * p))).
- apply H1. apply IQR_lt.
+ apply CReal_plus_le_compat_l. apply CRealLt_asym, maj.
+ + apply (CReal_le_trans _ (inject_Q (1 # 2 * p))).
+ apply H1. apply inject_Q_le.
rewrite <- Qplus_0_r.
setoid_replace (1#p)%Q with ((1#2*p)+(1#2*p))%Q.
- apply Qplus_lt_r. reflexivity.
+ apply Qplus_le_r. discriminate.
rewrite Qinv_plus_distr. reflexivity.
Qed.
diff --git a/theories/Reals/ConstructiveReals.v b/theories/Reals/ConstructiveReals.v
index fc3d6afe15..25242f5ea9 100644
--- a/theories/Reals/ConstructiveReals.v
+++ b/theories/Reals/ConstructiveReals.v
@@ -9,10 +9,10 @@
(************************************************************************)
(************************************************************************)
-(* An interface for constructive and computable real numbers.
- All of its instances are isomorphic, for example it contains
- the Cauchy reals implemented in file ConstructivecauchyReals
- and the sumbool-based Dedekind reals defined by
+(** An interface for constructive and computable real numbers.
+ All of its instances are isomorphic (see file ConstructiveRealsMorphisms).
+ For example it contains the Cauchy reals implemented in file
+ ConstructivecauchyReals and the sumbool-based Dedekind reals defined by
Structure R := {
(* The cuts are represented as propositional functions, rather than subsets,
@@ -41,7 +41,22 @@ Structure R := {
see github.com/andrejbauer/dedekind-reals for the Prop-based
version of those Dedekind reals (although Prop fails to make
- them an instance of ConstructiveReals). *)
+ them an instance of ConstructiveReals).
+
+ Any computation about constructive reals, can be worked
+ in the fastest instance for it; we then transport the results
+ to all other instances by the isomorphisms. This way of working
+ is different from the usual interfaces, where we would rather
+ prove things abstractly, by quantifying universally on the instance.
+
+ The functions of ConstructiveReals do not have a direct impact
+ on performance, because algorithms will be extracted from instances,
+ and because fast ConstructiveReals morphisms should be coded
+ manually. However, since instances are forced to implement
+ those functions, it is probable that they will also use them
+ in their algorithms. So those functions hint at what we think
+ will yield fast and small extracted programs. *)
+
Require Import QArith.
@@ -56,6 +71,9 @@ Definition orderEq (X : Set) (Xlt : X -> X -> Set) (x y : X) : Prop
Definition orderAppart (X : Set) (Xlt : X -> X -> Set) (x y : X) : Set
:= Xlt x y + Xlt y x.
+Definition orderLe (X : Set) (Xlt : X -> X -> Set) (x y : X) : Prop
+ := Xlt y x -> False.
+
Definition sig_forall_dec_T : Type
:= forall (P : nat -> Prop), (forall n, {P n} + {~P n})
-> {n | ~P n} + {forall n, P n}.
@@ -65,9 +83,17 @@ Definition sig_not_dec_T : Type := forall P : Prop, { ~~P } + { ~P }.
Record ConstructiveReals : Type :=
{
CRcarrier : Set;
+
+ (* Put this order relation in sort Set rather than Prop,
+ to allow the definition of fast ConstructiveReals morphisms.
+ For example, the Cauchy reals do store information in
+ the proofs of CRlt, which is used in algorithms in sort Set. *)
CRlt : CRcarrier -> CRcarrier -> Set;
CRltLinear : isLinearOrder CRcarrier CRlt;
+ (* The propositional truncation of CRlt. It facilitates proofs
+ when computations are not considered important, for example in
+ classical reals with extra logical axioms. *)
CRltProp : CRcarrier -> CRcarrier -> Prop;
(* This choice algorithm can be slow, keep it for the classical
quotient of the reals, where computations are blocked by
@@ -114,36 +140,696 @@ Record ConstructiveReals : Type :=
CRinv_0_lt_compat : forall (r : CRcarrier) (rnz : orderAppart _ CRlt r CRzero),
CRlt CRzero r -> CRlt CRzero (CRinv r rnz);
- CRarchimedean : forall x : CRcarrier,
- { k : Z & CRlt x (gen_phiZ CRzero CRone CRplus CRmult CRopp k) };
+ (* The initial field morphism (in characteristic zero).
+ The abstract definition by iteration of addition is
+ probably the slowest. Let each instance implement
+ a faster (and often simpler) version. *)
+ CR_of_Q : Q -> CRcarrier;
+ CR_of_Q_plus : forall q r : Q, orderEq _ CRlt (CR_of_Q (q+r))
+ (CRplus (CR_of_Q q) (CR_of_Q r));
+ CR_of_Q_mult : forall q r : Q, orderEq _ CRlt (CR_of_Q (q*r))
+ (CRmult (CR_of_Q q) (CR_of_Q r));
+ CR_of_Q_one : orderEq _ CRlt (CR_of_Q 1) CRone;
+ CR_of_Q_lt : forall q r : Q,
+ Qlt q r -> CRlt (CR_of_Q q) (CR_of_Q r);
+ lt_CR_of_Q : forall q r : Q,
+ CRlt (CR_of_Q q) (CR_of_Q r) -> Qlt q r;
+
+ (* This function is very fast in both the Cauchy and Dedekind
+ instances, because this rational number q is almost what
+ the proof of CRlt x y contains.
+ This function is also the heart of the computation of
+ constructive real numbers : it approximates x to any
+ requested precision y. *)
+ CR_Q_dense : forall x y : CRcarrier, CRlt x y ->
+ { q : Q & prod (CRlt x (CR_of_Q q))
+ (CRlt (CR_of_Q q) y) };
+ CR_archimedean : forall x : CRcarrier,
+ { n : positive & CRlt x (CR_of_Q (Z.pos n # 1)) };
CRminus (x y : CRcarrier) : CRcarrier
:= CRplus x (CRopp y);
+
+ (* Definitions of convergence and Cauchy-ness. The formulas
+ with orderLe or CRlt are logically equivalent, the choice of
+ orderLe in sort Prop is a question of performance.
+ It is very rare to turn back to the strict order to
+ define functions in sort Set, so we prefer to discard
+ those proofs during extraction. And even in those rare cases,
+ it is easy to divide epsilon by 2 for example. *)
CR_cv (un : nat -> CRcarrier) (l : CRcarrier) : Set
- := forall eps:CRcarrier,
- CRlt CRzero eps
- -> { p : nat & forall i:nat, le p i -> CRlt (CRopp eps) (CRminus (un i) l)
- * CRlt (CRminus (un i) l) eps };
+ := forall p:positive,
+ { n : nat | forall i:nat, le n i
+ -> orderLe _ CRlt (CR_of_Q (-1#p)) (CRminus (un i) l)
+ /\ orderLe _ CRlt (CRminus (un i) l) (CR_of_Q (1#p)) };
CR_cauchy (un : nat -> CRcarrier) : Set
- := forall eps:CRcarrier,
- CRlt CRzero eps
- -> { p : nat & forall i j:nat, le p i -> le p j ->
- CRlt (CRopp eps) (CRminus (un i) (un j))
- * CRlt (CRminus (un i) (un j)) eps };
+ := forall p : positive,
+ { n : nat | forall i j:nat, le n i -> le n j
+ -> orderLe _ CRlt (CR_of_Q (-1#p)) (CRminus (un i) (un j))
+ /\ orderLe _ CRlt (CRminus (un i) (un j)) (CR_of_Q (1#p)) };
+ (* For the Cauchy reals, this algorithm consists in building
+ a Cauchy sequence of rationals un : nat -> Q that has
+ the same limit as xn. For each n:nat, un n is a 1/n
+ rational approximation of a point of xn that has converged
+ within 1/n. *)
CR_complete :
- forall xn : nat -> CRcarrier, CR_cauchy xn -> { l : CRcarrier & CR_cv xn l };
-
- (* Those are redundant, they could be proved from the previous hypotheses *)
- CRis_upper_bound (E:CRcarrier -> Prop) (m:CRcarrier)
- := forall x:CRcarrier, E x -> CRlt m x -> False;
-
- CR_sig_lub :
- forall (E:CRcarrier -> Prop),
- sig_forall_dec_T
- -> sig_not_dec_T
- -> (exists x : CRcarrier, E x)
- -> (exists x : CRcarrier, CRis_upper_bound E x)
- -> { u : CRcarrier | CRis_upper_bound E u /\
- forall y:CRcarrier, CRis_upper_bound E y -> CRlt y u -> False };
+ forall xn : (nat -> CRcarrier),
+ CR_cauchy xn -> { l : CRcarrier & CR_cv xn l };
}.
+
+Lemma CRlt_asym : forall (R : ConstructiveReals) (x y : CRcarrier R),
+ CRlt R x y -> CRlt R y x -> False.
+Proof.
+ intros. destruct (CRltLinear R), p.
+ apply (f x y); assumption.
+Qed.
+
+Lemma CRlt_proper
+ : forall R : ConstructiveReals,
+ CMorphisms.Proper
+ (CMorphisms.respectful (orderEq _ (CRlt R))
+ (CMorphisms.respectful (orderEq _ (CRlt R)) CRelationClasses.iffT)) (CRlt R).
+Proof.
+ intros R x y H x0 y0 H0. destruct H, H0.
+ destruct (CRltLinear R). split.
+ - intro. destruct (s x y x0). assumption.
+ contradiction. destruct (s y y0 x0).
+ assumption. assumption. contradiction.
+ - intro. destruct (s y x y0). assumption.
+ contradiction. destruct (s x x0 y0).
+ assumption. assumption. contradiction.
+Qed.
+
+Lemma CRle_refl : forall (R : ConstructiveReals) (x : CRcarrier R),
+ CRlt R x x -> False.
+Proof.
+ intros. destruct (CRltLinear R), p.
+ exact (f x x H H).
+Qed.
+
+Lemma CRle_lt_trans : forall (R : ConstructiveReals) (r1 r2 r3 : CRcarrier R),
+ (CRlt R r2 r1 -> False) -> CRlt R r2 r3 -> CRlt R r1 r3.
+Proof.
+ intros. destruct (CRltLinear R).
+ destruct (s r2 r1 r3 H0). contradiction. apply c.
+Qed.
+
+Lemma CRlt_le_trans : forall (R : ConstructiveReals) (r1 r2 r3 : CRcarrier R),
+ CRlt R r1 r2 -> (CRlt R r3 r2 -> False) -> CRlt R r1 r3.
+Proof.
+ intros. destruct (CRltLinear R).
+ destruct (s r1 r3 r2 H). apply c. contradiction.
+Qed.
+
+Lemma CRle_trans : forall (R : ConstructiveReals) (x y z : CRcarrier R),
+ orderLe _ (CRlt R) x y -> orderLe _ (CRlt R) y z -> orderLe _ (CRlt R) x z.
+Proof.
+ intros. intro abs. apply H0.
+ apply (CRlt_le_trans _ _ x); assumption.
+Qed.
+
+Lemma CRlt_trans : forall (R : ConstructiveReals) (x y z : CRcarrier R),
+ CRlt R x y -> CRlt R y z -> CRlt R x z.
+Proof.
+ intros. apply (CRlt_le_trans R _ y _ H).
+ apply CRlt_asym. exact H0.
+Defined.
+
+Lemma CRlt_trans_flip : forall (R : ConstructiveReals) (x y z : CRcarrier R),
+ CRlt R y z -> CRlt R x y -> CRlt R x z.
+Proof.
+ intros. apply (CRlt_le_trans R _ y). exact H0.
+ apply CRlt_asym. exact H.
+Defined.
+
+Lemma CReq_refl : forall (R : ConstructiveReals) (x : CRcarrier R),
+ orderEq _ (CRlt R) x x.
+Proof.
+ split; apply CRle_refl.
+Qed.
+
+Lemma CReq_sym : forall (R : ConstructiveReals) (x y : CRcarrier R),
+ orderEq _ (CRlt R) x y
+ -> orderEq _ (CRlt R) y x.
+Proof.
+ intros. destruct H. split; intro abs; contradiction.
+Qed.
+
+Lemma CReq_trans : forall (R : ConstructiveReals) (x y z : CRcarrier R),
+ orderEq _ (CRlt R) x y
+ -> orderEq _ (CRlt R) y z
+ -> orderEq _ (CRlt R) x z.
+Proof.
+ intros. destruct H,H0. destruct (CRltLinear R), p. split.
+ - intro abs. destruct (s _ y _ abs); contradiction.
+ - intro abs. destruct (s _ y _ abs); contradiction.
+Qed.
+
+Lemma CR_setoid : forall R : ConstructiveReals,
+ Setoid_Theory (CRcarrier R) (orderEq _ (CRlt R)).
+Proof.
+ split. intro x. apply CReq_refl.
+ intros x y. apply CReq_sym.
+ intros x y z. apply CReq_trans.
+Qed.
+
+Lemma CRplus_0_r : forall (R : ConstructiveReals) (x : CRcarrier R),
+ orderEq _ (CRlt R) (CRplus R x (CRzero R)) x.
+Proof.
+ intros. destruct (CRisRing R).
+ apply (CReq_trans R _ (CRplus R (CRzero R) x)).
+ apply Radd_comm. apply Radd_0_l.
+Qed.
+
+Lemma CRmult_1_r : forall (R : ConstructiveReals) (x : CRcarrier R),
+ orderEq _ (CRlt R) (CRmult R x (CRone R)) x.
+Proof.
+ intros. destruct (CRisRing R).
+ apply (CReq_trans R _ (CRmult R (CRone R) x)).
+ apply Rmul_comm. apply Rmul_1_l.
+Qed.
+
+Lemma CRplus_opp_l : forall (R : ConstructiveReals) (x : CRcarrier R),
+ orderEq _ (CRlt R) (CRplus R (CRopp R x) x) (CRzero R).
+Proof.
+ intros. destruct (CRisRing R).
+ apply (CReq_trans R _ (CRplus R x (CRopp R x))).
+ apply Radd_comm. apply Ropp_def.
+Qed.
+
+Lemma CRplus_lt_compat_r : forall (R : ConstructiveReals) (r r1 r2 : CRcarrier R),
+ CRlt R r1 r2 -> CRlt R (CRplus R r1 r) (CRplus R r2 r).
+Proof.
+ intros. destruct (CRisRing R).
+ apply (CRlt_proper R (CRplus R r r1) (CRplus R r1 r) (Radd_comm _ _)
+ (CRplus R r2 r) (CRplus R r2 r)).
+ apply CReq_refl.
+ apply (CRlt_proper R _ _ (CReq_refl _ _) _ (CRplus R r r2)).
+ apply Radd_comm. apply CRplus_lt_compat_l. exact H.
+Qed.
+
+Lemma CRplus_lt_reg_r : forall (R : ConstructiveReals) (r r1 r2 : CRcarrier R),
+ CRlt R (CRplus R r1 r) (CRplus R r2 r) -> CRlt R r1 r2.
+Proof.
+ intros. destruct (CRisRing R).
+ apply (CRlt_proper R (CRplus R r r1) (CRplus R r1 r) (Radd_comm _ _)
+ (CRplus R r2 r) (CRplus R r2 r)) in H.
+ 2: apply CReq_refl.
+ apply (CRlt_proper R _ _ (CReq_refl _ _) _ (CRplus R r r2)) in H.
+ apply CRplus_lt_reg_l in H. exact H.
+ apply Radd_comm.
+Qed.
+
+Lemma CRplus_le_compat_l : forall (R : ConstructiveReals) (r r1 r2 : CRcarrier R),
+ orderLe _ (CRlt R) r1 r2
+ -> orderLe _ (CRlt R) (CRplus R r r1) (CRplus R r r2).
+Proof.
+ intros. intros abs. apply CRplus_lt_reg_l in abs. apply H. exact abs.
+Qed.
+
+Lemma CRplus_le_compat_r : forall (R : ConstructiveReals) (r r1 r2 : CRcarrier R),
+ orderLe _ (CRlt R) r1 r2
+ -> orderLe _ (CRlt R) (CRplus R r1 r) (CRplus R r2 r).
+Proof.
+ intros. intros abs. apply CRplus_lt_reg_r in abs. apply H. exact abs.
+Qed.
+
+Lemma CRplus_le_reg_l : forall (R : ConstructiveReals) (r r1 r2 : CRcarrier R),
+ orderLe _ (CRlt R) (CRplus R r r1) (CRplus R r r2)
+ -> orderLe _ (CRlt R) r1 r2.
+Proof.
+ intros. intro abs. apply H. clear H.
+ apply CRplus_lt_compat_l. exact abs.
+Qed.
+
+Lemma CRplus_le_reg_r : forall (R : ConstructiveReals) (r r1 r2 : CRcarrier R),
+ orderLe _ (CRlt R) (CRplus R r1 r) (CRplus R r2 r)
+ -> orderLe _ (CRlt R) r1 r2.
+Proof.
+ intros. intro abs. apply H. clear H.
+ apply CRplus_lt_compat_r. exact abs.
+Qed.
+
+Lemma CRplus_lt_le_compat :
+ forall (R : ConstructiveReals) (r1 r2 r3 r4 : CRcarrier R),
+ CRlt R r1 r2
+ -> (CRlt R r4 r3 -> False)
+ -> CRlt R (CRplus R r1 r3) (CRplus R r2 r4).
+Proof.
+ intros. apply (CRlt_le_trans R _ (CRplus R r2 r3)).
+ apply CRplus_lt_compat_r. exact H. intro abs.
+ apply CRplus_lt_reg_l in abs. contradiction.
+Qed.
+
+Lemma CRplus_eq_reg_l : forall (R : ConstructiveReals) (r r1 r2 : CRcarrier R),
+ orderEq _ (CRlt R) (CRplus R r r1) (CRplus R r r2)
+ -> orderEq _ (CRlt R) r1 r2.
+Proof.
+ intros.
+ destruct (CRisRingExt R). clear Rmul_ext Ropp_ext.
+ pose proof (Radd_ext
+ (CRopp R r) (CRopp R r) (CReq_refl _ _)
+ _ _ H).
+ destruct (CRisRing R).
+ apply (CReq_trans _ r1) in H0.
+ apply (CReq_trans R _ _ _ H0).
+ apply (CReq_trans R _ (CRplus R (CRplus R (CRopp R r) r) r2)).
+ apply Radd_assoc.
+ apply (CReq_trans R _ (CRplus R (CRzero R) r2)).
+ apply Radd_ext. apply CRplus_opp_l. apply CReq_refl.
+ apply Radd_0_l. apply CReq_sym.
+ apply (CReq_trans R _ (CRplus R (CRplus R (CRopp R r) r) r1)).
+ apply Radd_assoc.
+ apply (CReq_trans R _ (CRplus R (CRzero R) r1)).
+ apply Radd_ext. apply CRplus_opp_l. apply CReq_refl.
+ apply Radd_0_l.
+Qed.
+
+Lemma CRplus_eq_reg_r : forall (R : ConstructiveReals) (r r1 r2 : CRcarrier R),
+ orderEq _ (CRlt R) (CRplus R r1 r) (CRplus R r2 r)
+ -> orderEq _ (CRlt R) r1 r2.
+Proof.
+ intros. apply (CRplus_eq_reg_l R r).
+ apply (CReq_trans R _ (CRplus R r1 r)). apply (Radd_comm (CRisRing R)).
+ apply (CReq_trans R _ (CRplus R r2 r)).
+ exact H. apply (Radd_comm (CRisRing R)).
+Qed.
+
+Lemma CRopp_involutive : forall (R : ConstructiveReals) (r : CRcarrier R),
+ orderEq _ (CRlt R) (CRopp R (CRopp R r)) r.
+Proof.
+ intros. apply (CRplus_eq_reg_l R (CRopp R r)).
+ apply (CReq_trans R _ (CRzero R)). apply CRisRing.
+ apply CReq_sym, (CReq_trans R _ (CRplus R r (CRopp R r))).
+ apply CRisRing. apply CRisRing.
+Qed.
+
+Lemma CRopp_gt_lt_contravar
+ : forall (R : ConstructiveReals) (r1 r2 : CRcarrier R),
+ CRlt R r2 r1 -> CRlt R (CRopp R r1) (CRopp R r2).
+Proof.
+ intros. apply (CRplus_lt_reg_l R r1).
+ destruct (CRisRing R).
+ apply (CRle_lt_trans R _ (CRzero R)). apply Ropp_def.
+ apply (CRplus_lt_compat_l R (CRopp R r2)) in H.
+ apply (CRle_lt_trans R _ (CRplus R (CRopp R r2) r2)).
+ apply (CRle_trans R _ (CRplus R r2 (CRopp R r2))).
+ destruct (Ropp_def r2). exact H0.
+ destruct (Radd_comm r2 (CRopp R r2)). exact H1.
+ apply (CRlt_le_trans R _ _ _ H).
+ destruct (Radd_comm r1 (CRopp R r2)). exact H0.
+Qed.
+
+Lemma CRopp_lt_cancel : forall (R : ConstructiveReals) (r1 r2 : CRcarrier R),
+ CRlt R (CRopp R r2) (CRopp R r1) -> CRlt R r1 r2.
+Proof.
+ intros. apply (CRplus_lt_compat_r R r1) in H.
+ destruct (CRplus_opp_l R r1) as [_ H1].
+ apply (CRlt_le_trans R _ _ _ H) in H1. clear H.
+ apply (CRplus_lt_compat_l R r2) in H1.
+ destruct (CRplus_0_r R r2) as [_ H0].
+ apply (CRlt_le_trans R _ _ _ H1) in H0. clear H1.
+ destruct (Radd_assoc (CRisRing R) r2 (CRopp R r2) r1) as [H _].
+ apply (CRle_lt_trans R _ _ _ H) in H0. clear H.
+ apply (CRle_lt_trans R _ (CRplus R (CRzero R) r1)).
+ apply (Radd_0_l (CRisRing R)).
+ apply (CRle_lt_trans R _ (CRplus R (CRplus R r2 (CRopp R r2)) r1)).
+ 2: exact H0. apply CRplus_le_compat_r.
+ destruct (Ropp_def (CRisRing R) r2). exact H.
+Qed.
+
+Lemma CRopp_plus_distr : forall (R : ConstructiveReals) (r1 r2 : CRcarrier R),
+ orderEq _ (CRlt R) (CRopp R (CRplus R r1 r2)) (CRplus R (CRopp R r1) (CRopp R r2)).
+Proof.
+ intros. destruct (CRisRing R), (CRisRingExt R).
+ apply (CRplus_eq_reg_l R (CRplus R r1 r2)).
+ apply (CReq_trans R _ (CRzero R)). apply Ropp_def.
+ apply (CReq_trans R _ (CRplus R (CRplus R r2 r1) (CRplus R (CRopp R r1) (CRopp R r2)))).
+ apply (CReq_trans R _ (CRplus R r2 (CRplus R r1 (CRplus R (CRopp R r1) (CRopp R r2))))).
+ apply (CReq_trans R _ (CRplus R r2 (CRopp R r2))).
+ apply CReq_sym. apply Ropp_def. apply Radd_ext.
+ apply CReq_refl.
+ apply (CReq_trans R _ (CRplus R (CRzero R) (CRopp R r2))).
+ apply CReq_sym, Radd_0_l.
+ apply (CReq_trans R _ (CRplus R (CRplus R r1 (CRopp R r1)) (CRopp R r2))).
+ apply Radd_ext. 2: apply CReq_refl. apply CReq_sym, Ropp_def.
+ apply CReq_sym, Radd_assoc. apply Radd_assoc.
+ apply Radd_ext. 2: apply CReq_refl. apply Radd_comm.
+Qed.
+
+Lemma CRmult_plus_distr_l : forall (R : ConstructiveReals) (r1 r2 r3 : CRcarrier R),
+ orderEq _ (CRlt R) (CRmult R r1 (CRplus R r2 r3))
+ (CRplus R (CRmult R r1 r2) (CRmult R r1 r3)).
+Proof.
+ intros. destruct (CRisRing R).
+ apply (CReq_trans R _ (CRmult R (CRplus R r2 r3) r1)).
+ apply Rmul_comm.
+ apply (CReq_trans R _ (CRplus R (CRmult R r2 r1) (CRmult R r3 r1))).
+ apply Rdistr_l.
+ apply (CReq_trans R _ (CRplus R (CRmult R r1 r2) (CRmult R r3 r1))).
+ destruct (CRisRingExt R). apply Radd_ext.
+ apply Rmul_comm. apply CReq_refl.
+ destruct (CRisRingExt R). apply Radd_ext.
+ apply CReq_refl. apply Rmul_comm.
+Qed.
+
+(* x == x+x -> x == 0 *)
+Lemma CRzero_double : forall (R : ConstructiveReals) (x : CRcarrier R),
+ orderEq _ (CRlt R) x (CRplus R x x)
+ -> orderEq _ (CRlt R) x (CRzero R).
+Proof.
+ intros.
+ apply (CRplus_eq_reg_l R x), CReq_sym, (CReq_trans R _ x).
+ apply CRplus_0_r. exact H.
+Qed.
+
+Lemma CRmult_0_r : forall (R : ConstructiveReals) (x : CRcarrier R),
+ orderEq _ (CRlt R) (CRmult R x (CRzero R)) (CRzero R).
+Proof.
+ intros. apply CRzero_double.
+ apply (CReq_trans R _ (CRmult R x (CRplus R (CRzero R) (CRzero R)))).
+ destruct (CRisRingExt R). apply Rmul_ext. apply CReq_refl.
+ apply CReq_sym, CRplus_0_r.
+ destruct (CRisRing R). apply CRmult_plus_distr_l.
+Qed.
+
+Lemma CRopp_mult_distr_r : forall (R : ConstructiveReals) (r1 r2 : CRcarrier R),
+ orderEq _ (CRlt R) (CRopp R (CRmult R r1 r2))
+ (CRmult R r1 (CRopp R r2)).
+Proof.
+ intros. apply (CRplus_eq_reg_l R (CRmult R r1 r2)).
+ destruct (CRisRing R).
+ apply (CReq_trans R _ (CRzero R)). apply Ropp_def.
+ apply (CReq_trans R _ (CRmult R r1 (CRplus R r2 (CRopp R r2)))).
+ 2: apply CRmult_plus_distr_l.
+ apply (CReq_trans R _ (CRmult R r1 (CRzero R))).
+ apply CReq_sym, CRmult_0_r.
+ destruct (CRisRingExt R). apply Rmul_ext. apply CReq_refl.
+ apply CReq_sym, Ropp_def.
+Qed.
+
+Lemma CRopp_mult_distr_l : forall (R : ConstructiveReals) (r1 r2 : CRcarrier R),
+ orderEq _ (CRlt R) (CRopp R (CRmult R r1 r2))
+ (CRmult R (CRopp R r1) r2).
+Proof.
+ intros. apply (CReq_trans R _ (CRmult R r2 (CRopp R r1))).
+ apply (CReq_trans R _ (CRopp R (CRmult R r2 r1))).
+ apply (Ropp_ext (CRisRingExt R)).
+ apply CReq_sym, (Rmul_comm (CRisRing R)).
+ apply CRopp_mult_distr_r.
+ apply CReq_sym, (Rmul_comm (CRisRing R)).
+Qed.
+
+Lemma CRmult_lt_compat_r : forall (R : ConstructiveReals) (r r1 r2 : CRcarrier R),
+ CRlt R (CRzero R) r
+ -> CRlt R r1 r2
+ -> CRlt R (CRmult R r1 r) (CRmult R r2 r).
+Proof.
+ intros. apply (CRplus_lt_reg_r R (CRopp R (CRmult R r1 r))).
+ apply (CRle_lt_trans R _ (CRzero R)).
+ apply (Ropp_def (CRisRing R)).
+ apply (CRlt_le_trans R _ (CRplus R (CRmult R r2 r) (CRmult R (CRopp R r1) r))).
+ apply (CRlt_le_trans R _ (CRmult R (CRplus R r2 (CRopp R r1)) r)).
+ apply CRmult_lt_0_compat. 2: exact H.
+ apply (CRplus_lt_reg_r R r1).
+ apply (CRle_lt_trans R _ r1). apply (Radd_0_l (CRisRing R)).
+ apply (CRlt_le_trans R _ r2 _ H0).
+ apply (CRle_trans R _ (CRplus R r2 (CRplus R (CRopp R r1) r1))).
+ apply (CRle_trans R _ (CRplus R r2 (CRzero R))).
+ destruct (CRplus_0_r R r2). exact H1.
+ apply CRplus_le_compat_l. destruct (CRplus_opp_l R r1). exact H1.
+ destruct (Radd_assoc (CRisRing R) r2 (CRopp R r1) r1). exact H2.
+ destruct (CRisRing R).
+ destruct (Rdistr_l r2 (CRopp R r1) r). exact H2.
+ apply CRplus_le_compat_l. destruct (CRopp_mult_distr_l R r1 r).
+ exact H1.
+Qed.
+
+Lemma CRinv_r : forall (R : ConstructiveReals) (r:CRcarrier R)
+ (rnz : orderAppart _ (CRlt R) r (CRzero R)),
+ orderEq _ (CRlt R) (CRmult R r (CRinv R r rnz)) (CRone R).
+Proof.
+ intros. apply (CReq_trans R _ (CRmult R (CRinv R r rnz) r)).
+ apply (CRisRing R). apply CRinv_l.
+Qed.
+
+Lemma CRmult_lt_reg_r : forall (R : ConstructiveReals) (r r1 r2 : CRcarrier R),
+ CRlt R (CRzero R) r
+ -> CRlt R (CRmult R r1 r) (CRmult R r2 r)
+ -> CRlt R r1 r2.
+Proof.
+ intros. apply (CRmult_lt_compat_r R (CRinv R r (inr H))) in H0.
+ 2: apply CRinv_0_lt_compat, H.
+ apply (CRle_lt_trans R _ (CRmult R (CRmult R r1 r) (CRinv R r (inr H)))).
+ - clear H0. apply (CRle_trans R _ (CRmult R r1 (CRone R))).
+ destruct (CRmult_1_r R r1). exact H0.
+ apply (CRle_trans R _ (CRmult R r1 (CRmult R r (CRinv R r (inr H))))).
+ destruct (Rmul_ext (CRisRingExt R) r1 r1 (CReq_refl R r1)
+ (CRmult R r (CRinv R r (inr H))) (CRone R)).
+ apply CRinv_r. exact H0.
+ destruct (Rmul_assoc (CRisRing R) r1 r (CRinv R r (inr H))). exact H1.
+ - apply (CRlt_le_trans R _ (CRmult R (CRmult R r2 r) (CRinv R r (inr H)))).
+ exact H0. clear H0.
+ apply (CRle_trans R _ (CRmult R r2 (CRone R))).
+ 2: destruct (CRmult_1_r R r2); exact H1.
+ apply (CRle_trans R _ (CRmult R r2 (CRmult R r (CRinv R r (inr H))))).
+ destruct (Rmul_assoc (CRisRing R) r2 r (CRinv R r (inr H))). exact H0.
+ destruct (Rmul_ext (CRisRingExt R) r2 r2 (CReq_refl R r2)
+ (CRmult R r (CRinv R r (inr H))) (CRone R)).
+ apply CRinv_r. exact H1.
+Qed.
+
+Lemma CRmult_lt_reg_l : forall (R : ConstructiveReals) (r r1 r2 : CRcarrier R),
+ CRlt R (CRzero R) r
+ -> CRlt R (CRmult R r r1) (CRmult R r r2)
+ -> CRlt R r1 r2.
+Proof.
+ intros.
+ destruct (Rmul_comm (CRisRing R) r r1) as [H1 _].
+ apply (CRle_lt_trans R _ _ _ H1) in H0. clear H1.
+ destruct (Rmul_comm (CRisRing R) r r2) as [_ H1].
+ apply (CRlt_le_trans R _ _ _ H0) in H1. clear H0.
+ apply CRmult_lt_reg_r in H1.
+ exact H1. exact H.
+Qed.
+
+Lemma CRmult_le_compat_l : forall (R : ConstructiveReals) (r r1 r2 : CRcarrier R),
+ CRlt R (CRzero R) r
+ -> orderLe _ (CRlt R) r1 r2
+ -> orderLe _ (CRlt R) (CRmult R r r1) (CRmult R r r2).
+Proof.
+ intros. intro abs. apply CRmult_lt_reg_l in abs.
+ contradiction. exact H.
+Qed.
+
+Lemma CRmult_le_compat_r : forall (R : ConstructiveReals) (r r1 r2 : CRcarrier R),
+ CRlt R (CRzero R) r
+ -> orderLe _ (CRlt R) r1 r2
+ -> orderLe _ (CRlt R) (CRmult R r1 r) (CRmult R r2 r).
+Proof.
+ intros. intro abs. apply CRmult_lt_reg_r in abs.
+ contradiction. exact H.
+Qed.
+
+Lemma CRmult_eq_reg_r : forall (R : ConstructiveReals) (r r1 r2 : CRcarrier R),
+ orderAppart _ (CRlt R) (CRzero R) r
+ -> orderEq _ (CRlt R) (CRmult R r1 r) (CRmult R r2 r)
+ -> orderEq _ (CRlt R) r1 r2.
+Proof.
+ intros. destruct H0,H.
+ - split.
+ + intro abs. apply H0. apply CRmult_lt_compat_r.
+ exact c. exact abs.
+ + intro abs. apply H1. apply CRmult_lt_compat_r.
+ exact c. exact abs.
+ - split.
+ + intro abs. apply H1. apply CRopp_lt_cancel.
+ apply (CRle_lt_trans R _ (CRmult R r1 (CRopp R r))).
+ apply CRopp_mult_distr_r.
+ apply (CRlt_le_trans R _ (CRmult R r2 (CRopp R r))).
+ 2: apply CRopp_mult_distr_r.
+ apply CRmult_lt_compat_r. 2: exact abs.
+ apply (CRplus_lt_reg_r R r). apply (CRle_lt_trans R _ r).
+ apply (Radd_0_l (CRisRing R)).
+ apply (CRlt_le_trans R _ (CRzero R) _ c).
+ apply CRplus_opp_l.
+ + intro abs. apply H0. apply CRopp_lt_cancel.
+ apply (CRle_lt_trans R _ (CRmult R r2 (CRopp R r))).
+ apply CRopp_mult_distr_r.
+ apply (CRlt_le_trans R _ (CRmult R r1 (CRopp R r))).
+ 2: apply CRopp_mult_distr_r.
+ apply CRmult_lt_compat_r. 2: exact abs.
+ apply (CRplus_lt_reg_r R r). apply (CRle_lt_trans R _ r).
+ apply (Radd_0_l (CRisRing R)).
+ apply (CRlt_le_trans R _ (CRzero R) _ c).
+ apply CRplus_opp_l.
+Qed.
+
+Lemma CR_of_Q_proper : forall (R : ConstructiveReals) (q r : Q),
+ q == r -> orderEq _ (CRlt R) (CR_of_Q R q) (CR_of_Q R r).
+Proof.
+ split.
+ - intro abs. apply lt_CR_of_Q in abs. rewrite H in abs.
+ exact (Qlt_not_le r r abs (Qle_refl r)).
+ - intro abs. apply lt_CR_of_Q in abs. rewrite H in abs.
+ exact (Qlt_not_le r r abs (Qle_refl r)).
+Qed.
+
+Lemma CR_of_Q_zero : forall (R : ConstructiveReals),
+ orderEq _ (CRlt R) (CR_of_Q R 0) (CRzero R).
+Proof.
+ intros. apply CRzero_double.
+ apply (CReq_trans R _ (CR_of_Q R (0+0))). apply CR_of_Q_proper.
+ reflexivity. apply CR_of_Q_plus.
+Qed.
+
+Lemma CR_of_Q_opp : forall (R : ConstructiveReals) (q : Q),
+ orderEq _ (CRlt R) (CR_of_Q R (-q)) (CRopp R (CR_of_Q R q)).
+Proof.
+ intros. apply (CRplus_eq_reg_l R (CR_of_Q R q)).
+ apply (CReq_trans R _ (CRzero R)).
+ apply (CReq_trans R _ (CR_of_Q R (q-q))).
+ apply CReq_sym, CR_of_Q_plus.
+ apply (CReq_trans R _ (CR_of_Q R 0)).
+ apply CR_of_Q_proper. ring. apply CR_of_Q_zero.
+ apply CReq_sym. apply (CRisRing R).
+Qed.
+
+Lemma CR_of_Q_le : forall (R : ConstructiveReals) (r q : Q),
+ Qle r q
+ -> orderLe _ (CRlt R) (CR_of_Q R r) (CR_of_Q R q).
+Proof.
+ intros. intro abs. apply lt_CR_of_Q in abs.
+ exact (Qlt_not_le _ _ abs H).
+Qed.
+
+Lemma CR_of_Q_pos : forall (R : ConstructiveReals) (q:Q),
+ Qlt 0 q -> CRlt R (CRzero R) (CR_of_Q R q).
+Proof.
+ intros. apply (CRle_lt_trans R _ (CR_of_Q R 0)).
+ apply CR_of_Q_zero. apply CR_of_Q_lt. exact H.
+Qed.
+
+Lemma CR_cv_above_rat
+ : forall (R : ConstructiveReals) (xn : nat -> Q) (x : CRcarrier R) (q : Q),
+ CR_cv R (fun n : nat => CR_of_Q R (xn n)) x
+ -> CRlt R (CR_of_Q R q) x
+ -> { n : nat | forall p:nat, le n p -> Qlt q (xn p) }.
+Proof.
+ intros.
+ destruct (CR_Q_dense R _ _ H0) as [r [H1 H2]].
+ apply lt_CR_of_Q in H1. clear H0.
+ destruct (Qarchimedean (/(r-q))) as [p pmaj].
+ destruct (H p) as [n nmaj].
+ exists n. intros k lenk. specialize (nmaj k lenk) as [H3 _].
+ apply (lt_CR_of_Q R), (CRlt_le_trans R _ (CRplus R x (CR_of_Q R (-1#p)))).
+ apply (CRlt_trans R _ (CRplus R (CR_of_Q R r) (CR_of_Q R (-1#p)))).
+ 2: apply CRplus_lt_compat_r, H2.
+ apply (CRlt_le_trans R _ (CR_of_Q R (r+(-1#p)))).
+ - apply CR_of_Q_lt.
+ apply (Qplus_lt_l _ _ (-(-1#p)-q)). field_simplify.
+ setoid_replace (-1*(-1#p)) with (1#p). 2: reflexivity.
+ apply (Qmult_lt_l _ _ (r-q)) in pmaj.
+ rewrite Qmult_inv_r in pmaj. apply Qlt_shift_div_r in pmaj.
+ 2: reflexivity. setoid_replace (-1*q + r) with (r-q). exact pmaj.
+ ring. intro abs. apply Qlt_minus_iff in H1.
+ rewrite abs in H1. inversion H1.
+ apply Qlt_minus_iff in H1. exact H1.
+ - apply CR_of_Q_plus.
+ - apply (CRplus_le_reg_r R (CRopp R x)).
+ apply (CRle_trans R _ (CR_of_Q R (-1#p))). 2: exact H3. clear H3.
+ apply (CRle_trans R _ (CRplus R (CRopp R x) (CRplus R x (CR_of_Q R (-1 # p))))).
+ exact (proj1 (Radd_comm (CRisRing R) _ _)).
+ apply (CRle_trans R _ (CRplus R (CRplus R (CRopp R x) x) (CR_of_Q R (-1 # p)))).
+ exact (proj2 (Radd_assoc (CRisRing R) _ _ _)).
+ apply (CRle_trans R _ (CRplus R (CRzero R) (CR_of_Q R (-1 # p)))).
+ apply CRplus_le_compat_r. exact (proj2 (CRplus_opp_l R _)).
+ exact (proj2 (Radd_0_l (CRisRing R) _)).
+Qed.
+
+Lemma CR_cv_below_rat
+ : forall (R : ConstructiveReals) (xn : nat -> Q) (x : CRcarrier R) (q : Q),
+ CR_cv R (fun n : nat => CR_of_Q R (xn n)) x
+ -> CRlt R x (CR_of_Q R q)
+ -> { n : nat | forall p:nat, le n p -> Qlt (xn p) q }.
+Proof.
+ intros.
+ destruct (CR_Q_dense R _ _ H0) as [r [H1 H2]].
+ apply lt_CR_of_Q in H2. clear H0.
+ destruct (Qarchimedean (/(q-r))) as [p pmaj].
+ destruct (H p) as [n nmaj].
+ exists n. intros k lenk. specialize (nmaj k lenk) as [_ H4].
+ apply (lt_CR_of_Q R), (CRle_lt_trans R _ (CRplus R x (CR_of_Q R (1#p)))).
+ - apply (CRplus_le_reg_r R (CRopp R x)).
+ apply (CRle_trans R _ (CR_of_Q R (1#p))). exact H4. clear H4.
+ apply (CRle_trans R _ (CRplus R (CRopp R x) (CRplus R x (CR_of_Q R (1 # p))))).
+ 2: exact (proj1 (Radd_comm (CRisRing R) _ _)).
+ apply (CRle_trans R _ (CRplus R (CRplus R (CRopp R x) x) (CR_of_Q R (1 # p)))).
+ 2: exact (proj1 (Radd_assoc (CRisRing R) _ _ _)).
+ apply (CRle_trans R _ (CRplus R (CRzero R) (CR_of_Q R (1 # p)))).
+ exact (proj1 (Radd_0_l (CRisRing R) _)).
+ apply CRplus_le_compat_r. exact (proj1 (CRplus_opp_l R _)).
+ - apply (CRlt_trans R _ (CRplus R (CR_of_Q R r) (CR_of_Q R (1 # p)))).
+ apply CRplus_lt_compat_r. exact H1.
+ apply (CRle_lt_trans R _ (CR_of_Q R (r + (1#p)))).
+ apply CR_of_Q_plus. apply CR_of_Q_lt.
+ apply (Qmult_lt_l _ _ (q-r)) in pmaj.
+ rewrite Qmult_inv_r in pmaj. apply Qlt_shift_div_r in pmaj.
+ apply (Qplus_lt_l _ _ (-r)). field_simplify.
+ setoid_replace (-1*r + q) with (q-r). exact pmaj.
+ ring. reflexivity. intro abs. apply Qlt_minus_iff in H2.
+ rewrite abs in H2. inversion H2.
+ apply Qlt_minus_iff in H2. exact H2.
+Qed.
+
+Lemma CR_cv_const : forall (R : ConstructiveReals) (x y : CRcarrier R),
+ CR_cv R (fun _ => x) y -> orderEq _ (CRlt R) x y.
+Proof.
+ intros. destruct (CRisRing R). split.
+ - intro abs.
+ destruct (CR_Q_dense R x y abs) as [q [H0 H1]].
+ destruct (CR_Q_dense R _ _ H1) as [r [H2 H3]].
+ apply lt_CR_of_Q in H2.
+ destruct (Qarchimedean (/(r-q))) as [p pmaj].
+ destruct (H p) as [n nmaj]. specialize (nmaj n (le_refl n)) as [nmaj _].
+ apply nmaj. clear nmaj.
+ apply (CRlt_trans R _ (CR_of_Q R (q-r))).
+ apply (CRlt_le_trans R _ (CRplus R (CR_of_Q R q) (CRopp R (CR_of_Q R r)))).
+ + apply CRplus_lt_le_compat. exact H0.
+ intro H4. apply CRopp_lt_cancel in H4. exact (CRlt_asym R _ _ H4 H3).
+ + apply (CRle_trans R _ (CRplus R (CR_of_Q R q) (CR_of_Q R (-r)))).
+ apply CRplus_le_compat_l. exact (proj1 (CR_of_Q_opp R r)).
+ exact (proj1 (CR_of_Q_plus R _ _)).
+ + apply CR_of_Q_lt.
+ apply (Qplus_lt_l _ _ (-(-1#p)+r-q)). field_simplify.
+ setoid_replace (-1*(-1#p)) with (1#p). 2: reflexivity.
+ apply (Qmult_lt_l _ _ (r-q)) in pmaj.
+ rewrite Qmult_inv_r in pmaj. apply Qlt_shift_div_r in pmaj.
+ 2: reflexivity. setoid_replace (-1*q + r) with (r-q). exact pmaj.
+ ring. intro H4. apply Qlt_minus_iff in H2.
+ rewrite H4 in H2. inversion H2.
+ apply Qlt_minus_iff in H2. exact H2.
+ - intro abs.
+ destruct (CR_Q_dense R _ _ abs) as [q [H0 H1]].
+ destruct (CR_Q_dense R _ _ H0) as [r [H2 H3]].
+ apply lt_CR_of_Q in H3.
+ destruct (Qarchimedean (/(q-r))) as [p pmaj].
+ destruct (H p) as [n nmaj]. specialize (nmaj n (le_refl n)) as [_ nmaj].
+ apply nmaj. clear nmaj.
+ apply (CRlt_trans R _ (CR_of_Q R (q-r))).
+ + apply CR_of_Q_lt.
+ apply (Qmult_lt_l _ _ (q-r)) in pmaj.
+ rewrite Qmult_inv_r in pmaj. apply Qlt_shift_div_r in pmaj.
+ exact pmaj. reflexivity.
+ intro H4. apply Qlt_minus_iff in H3.
+ rewrite H4 in H3. inversion H3.
+ apply Qlt_minus_iff in H3. exact H3.
+ + apply (CRle_lt_trans R _ (CRplus R (CR_of_Q R q) (CR_of_Q R (-r)))).
+ apply CR_of_Q_plus.
+ apply (CRle_lt_trans R _ (CRplus R (CR_of_Q R q) (CRopp R (CR_of_Q R r)))).
+ apply CRplus_le_compat_l. exact (proj2 (CR_of_Q_opp R r)).
+ apply CRplus_lt_le_compat. exact H1.
+ intro H4. apply CRopp_lt_cancel in H4.
+ exact (CRlt_asym R _ _ H4 H2).
+Qed.
diff --git a/theories/Reals/ConstructiveRealsLUB.v b/theories/Reals/ConstructiveRealsLUB.v
index f5c447f7db..3a26b8cefb 100644
--- a/theories/Reals/ConstructiveRealsLUB.v
+++ b/theories/Reals/ConstructiveRealsLUB.v
@@ -15,7 +15,9 @@
Require Import QArith_base.
Require Import Qabs.
-Require Import ConstructiveCauchyReals.
+Require Import ConstructiveReals.
+Require Import ConstructiveCauchyRealsMult.
+Require Import ConstructiveRealsMorphisms.
Require Import ConstructiveRcomplete.
Require Import Logic.ConstructiveEpsilon.
@@ -54,14 +56,15 @@ Lemma is_upper_bound_epsilon :
sig_forall_dec_T
-> sig_not_dec_T
-> (exists x:CReal, is_upper_bound E x)
- -> { n:nat | is_upper_bound E (INR n) }.
+ -> { n:nat | is_upper_bound E (inject_Q (Z.of_nat n # 1)) }.
Proof.
intros E lpo sig_not_dec Ebound.
apply constructive_indefinite_ground_description_nat.
- intro n. apply is_upper_bound_dec. exact lpo. exact sig_not_dec.
- - destruct Ebound as [x H]. destruct (Rup_nat x). exists x0.
+ - destruct Ebound as [x H]. destruct (Rup_pos x). exists (Pos.to_nat x0).
intros y ey. specialize (H y ey).
- apply CRealLt_asym. apply (CRealLe_Lt_trans _ x); assumption.
+ apply CRealLt_asym. apply (CReal_le_lt_trans _ x).
+ exact H. rewrite positive_nat_Z. exact c.
Qed.
Lemma is_upper_bound_not_epsilon :
@@ -69,15 +72,16 @@ Lemma is_upper_bound_not_epsilon :
sig_forall_dec_T
-> sig_not_dec_T
-> (exists x : CReal, E x)
- -> { m:nat | ~is_upper_bound E (-INR m) }.
+ -> { m:nat | ~is_upper_bound E (-inject_Q (Z.of_nat m # 1)) }.
Proof.
intros E lpo sig_not_dec H.
apply constructive_indefinite_ground_description_nat.
- - intro n. destruct (is_upper_bound_dec E (-INR n) lpo sig_not_dec).
+ - intro n. destruct (is_upper_bound_dec E (-inject_Q (Z.of_nat n # 1)) lpo sig_not_dec).
right. intro abs. contradiction. left. exact n0.
- - destruct H as [x H]. destruct (Rup_nat (-x)) as [n H0].
- exists n. intro abs. specialize (abs x H).
- apply abs. apply (CReal_plus_lt_reg_l (INR n-x)).
+ - destruct H as [x H]. destruct (Rup_pos (-x)) as [n H0].
+ exists (Pos.to_nat n). intro abs. specialize (abs x H).
+ apply abs. rewrite positive_nat_Z.
+ apply (CReal_plus_lt_reg_l (inject_Q (Z.pos n # 1)-x)).
ring_simplify. exact H0.
Qed.
@@ -140,8 +144,8 @@ Proof.
Qed.
Lemma glb_dec_Q : forall upcut : DedekindDecCut,
- { x : CReal | forall r:Q, (x < IQR r -> DDupcut upcut r)
- /\ (IQR r < x -> ~DDupcut upcut r) }.
+ { x : CReal | forall r:Q, (x < inject_Q r -> DDupcut upcut r)
+ /\ (inject_Q r < x -> ~DDupcut upcut r) }.
Proof.
intros.
assert (forall a b : Q, Qle a b -> Qle (-b) (-a)).
@@ -175,7 +179,7 @@ Proof.
pose (exist (fun qn => QSeqEquiv qn qn Pos.to_nat) _ H0) as l.
exists l. split.
- intros. (* find an upper point between the limit and r *)
- rewrite FinjectQ_CReal in H1. destruct H1 as [p pmaj].
+ destruct H1 as [p pmaj].
unfold l,proj1_sig in pmaj.
destruct (DDcut_limit upcut (1 # Pos.of_nat (Pos.to_nat p)) eq_refl) as [q qmaj]
; simpl in pmaj.
@@ -184,8 +188,7 @@ Proof.
apply (Qle_trans _ ((2#p) + q)).
apply (Qplus_le_l _ _ (-q)). ring_simplify. discriminate.
apply Qlt_le_weak. exact pmaj.
- - intros H1 abs.
- rewrite FinjectQ_CReal in H1. destruct H1 as [p pmaj].
+ - intros [p pmaj] abs.
unfold l,proj1_sig in pmaj.
destruct (DDcut_limit upcut (1 # Pos.of_nat (Pos.to_nat p)) eq_refl) as [q qmaj]
; simpl in pmaj.
@@ -205,26 +208,24 @@ Lemma is_upper_bound_glb :
-> sig_forall_dec_T
-> (exists x : CReal, E x)
-> (exists x : CReal, is_upper_bound E x)
- -> { x : CReal | forall r:Q, (x < IQR r -> is_upper_bound E (IQR r))
- /\ (IQR r < x -> ~is_upper_bound E (IQR r)) }.
+ -> { x : CReal | forall r:Q, (x < inject_Q r -> is_upper_bound E (inject_Q r))
+ /\ (inject_Q r < x -> ~is_upper_bound E (inject_Q r)) }.
Proof.
intros E sig_not_dec lpo Einhab Ebound.
destruct (is_upper_bound_epsilon E lpo sig_not_dec Ebound) as [a luba].
destruct (is_upper_bound_not_epsilon E lpo sig_not_dec Einhab) as [b glbb].
- pose (fun q => is_upper_bound E (IQR q)) as upcut.
+ pose (fun q => is_upper_bound E (inject_Q q)) as upcut.
assert (forall q:Q, { upcut q } + { ~upcut q } ).
{ intro q. apply is_upper_bound_dec. exact lpo. exact sig_not_dec. }
assert (forall q r : Q, (q <= r)%Q -> upcut q -> upcut r).
{ intros. intros x Ex. specialize (H1 x Ex). intro abs.
- apply H1. apply (CRealLe_Lt_trans _ (IQR r)). 2: exact abs.
- apply IQR_le. exact H0. }
+ apply H1. apply (CReal_le_lt_trans _ (inject_Q r)). 2: exact abs.
+ apply inject_Q_le. exact H0. }
assert (upcut (Z.of_nat a # 1)%Q).
- { intros x Ex. unfold IQR. rewrite CReal_inv_1, CReal_mult_1_r.
- specialize (luba x Ex). rewrite <- INR_IZR_INZ. exact luba. }
+ { intros x Ex. exact (luba x Ex). }
assert (~upcut (- Z.of_nat b # 1)%Q).
{ intros abs. apply glbb. intros x Ex.
- specialize (abs x Ex). unfold IQR in abs.
- rewrite CReal_inv_1, CReal_mult_1_r, opp_IZR, <- INR_IZR_INZ in abs.
+ specialize (abs x Ex). rewrite <- opp_inject_Q.
exact abs. }
assert (forall q r : Q, (q == r)%Q -> upcut q -> upcut r).
{ intros. intros x Ex. specialize (H4 x Ex). rewrite <- H3. exact H4. }
@@ -257,7 +258,7 @@ Proof.
intro abs. destruct (FQ_dense b x abs) as [q [qmaj H0]].
specialize (a q) as [_ a]. apply a. exact H0.
intros y Ey. specialize (H y Ey). intro abs2.
- apply H. exact (CRealLt_trans _ (IQR q) _ qmaj abs2).
+ apply H. exact (CReal_lt_trans _ (inject_Q q) _ qmaj abs2).
Qed.
Lemma sig_lub :
@@ -274,3 +275,44 @@ Proof.
E sig_not_dec sig_forall_dec Einhab Ebound); simpl in H.
exists x. exact H.
Qed.
+
+Definition CRis_upper_bound (R : ConstructiveReals) (E:CRcarrier R -> Prop) (m:CRcarrier R)
+ := forall x:CRcarrier R, E x -> CRlt R m x -> False.
+
+Lemma CR_sig_lub :
+ forall (R : ConstructiveReals) (E:CRcarrier R -> Prop),
+ (forall x y : CRcarrier R, orderEq _ (CRlt R) x y -> (E x <-> E y))
+ -> sig_forall_dec_T
+ -> sig_not_dec_T
+ -> (exists x : CRcarrier R, E x)
+ -> (exists x : CRcarrier R, CRis_upper_bound R E x)
+ -> { u : CRcarrier R | CRis_upper_bound R E u /\
+ forall y:CRcarrier R, CRis_upper_bound R E y -> CRlt R y u -> False }.
+Proof.
+ intros. destruct (sig_lub (fun x:CReal => E (CauchyMorph R x)) X X0) as [u ulub].
+ - destruct H0. exists (CauchyMorph_inv R x).
+ specialize (H (CauchyMorph R (CauchyMorph_inv R x)) x
+ (CauchyMorph_surject R x)) as [_ H].
+ exact (H H0).
+ - destruct H1. exists (CauchyMorph_inv R x).
+ intros y Ey. specialize (H1 (CauchyMorph R y) Ey).
+ intros abs. apply H1.
+ apply (CauchyMorph_increasing R) in abs.
+ apply (CRle_lt_trans R _ (CauchyMorph R (CauchyMorph_inv R x))).
+ 2: exact abs. apply (CauchyMorph_surject R x).
+ - exists (CauchyMorph R u). destruct ulub. split.
+ + intros y Ey abs. specialize (H2 (CauchyMorph_inv R y)).
+ simpl in H2.
+ specialize (H (CauchyMorph R (CauchyMorph_inv R y)) y
+ (CauchyMorph_surject R y)) as [_ H].
+ specialize (H2 (H Ey)). apply H2.
+ apply CauchyMorph_inv_increasing in abs.
+ rewrite CauchyMorph_inject in abs. exact abs.
+ + intros. apply (H3 (CauchyMorph_inv R y)).
+ intros z Ez abs. specialize (H4 (CauchyMorph R z)).
+ apply (H4 Ez). apply (CauchyMorph_increasing R) in abs.
+ apply (CRle_lt_trans R _ (CauchyMorph R (CauchyMorph_inv R y))).
+ 2: exact abs. apply (CauchyMorph_surject R y).
+ apply CauchyMorph_inv_increasing in H5.
+ rewrite CauchyMorph_inject in H5. exact H5.
+Qed.
diff --git a/theories/Reals/ConstructiveRealsMorphisms.v b/theories/Reals/ConstructiveRealsMorphisms.v
new file mode 100644
index 0000000000..0d3027d475
--- /dev/null
+++ b/theories/Reals/ConstructiveRealsMorphisms.v
@@ -0,0 +1,1158 @@
+(************************************************************************)
+(* * The Coq Proof Assistant / The Coq Development Team *)
+(* v * INRIA, CNRS and contributors - Copyright 1999-2019 *)
+(* <O___,, * (see CREDITS file for the list of authors) *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(* * (see LICENSE file for the text of the license) *)
+(************************************************************************)
+(************************************************************************)
+
+(** Morphisms used to transport results from any instance of
+ ConstructiveReals to any other.
+ Between any two constructive reals structures R1 and R2,
+ all morphisms R1 -> R2 are extensionally equal. We will
+ further show that they exist, and so are isomorphisms.
+ The difference between two morphisms R1 -> R2 is therefore
+ the speed of computation.
+
+ The canonical isomorphisms we provide here are often very slow,
+ when a new implementation of constructive reals is added,
+ it should define its own ad hoc isomorphisms for better speed.
+
+ Apart from the speed, those unique isomorphisms also serve as
+ sanity checks of the interface ConstructiveReals :
+ it captures a concept with a strong notion of uniqueness. *)
+
+Require Import QArith.
+Require Import Qabs.
+Require Import ConstructiveReals.
+Require Import ConstructiveCauchyRealsMult.
+Require Import ConstructiveRIneq.
+
+
+Record ConstructiveRealsMorphism (R1 R2 : ConstructiveReals) : Set :=
+ {
+ CRmorph : CRcarrier R1 -> CRcarrier R2;
+ CRmorph_rat : forall q : Q,
+ orderEq _ (CRlt R2) (CRmorph (CR_of_Q R1 q)) (CR_of_Q R2 q);
+ CRmorph_increasing : forall x y : CRcarrier R1,
+ CRlt R1 x y -> CRlt R2 (CRmorph x) (CRmorph y);
+ }.
+
+
+Lemma CRmorph_increasing_inv
+ : forall (R1 R2 : ConstructiveReals)
+ (f : ConstructiveRealsMorphism R1 R2)
+ (x y : CRcarrier R1),
+ CRlt R2 (CRmorph _ _ f x) (CRmorph _ _ f y)
+ -> CRlt R1 x y.
+Proof.
+ intros. destruct (CR_Q_dense R2 _ _ H) as [q [H0 H1]].
+ destruct (CR_Q_dense R2 _ _ H0) as [r [H2 H3]].
+ apply lt_CR_of_Q, (CR_of_Q_lt R1) in H3.
+ destruct (CRltLinear R1).
+ destruct (s _ x _ H3).
+ - exfalso. apply (CRmorph_increasing _ _ f) in c.
+ destruct (CRmorph_rat _ _ f r) as [H4 _].
+ apply (CRle_lt_trans R2 _ _ _ H4) in c. clear H4.
+ exact (CRlt_asym R2 _ _ c H2).
+ - clear H2 H3 r. apply (CRlt_trans R1 _ _ _ c). clear c.
+ destruct (CR_Q_dense R2 _ _ H1) as [t [H2 H3]].
+ apply lt_CR_of_Q, (CR_of_Q_lt R1) in H2.
+ destruct (s _ y _ H2). exact c.
+ exfalso. apply (CRmorph_increasing _ _ f) in c.
+ destruct (CRmorph_rat _ _ f t) as [_ H4].
+ apply (CRlt_le_trans R2 _ _ _ c) in H4. clear c.
+ exact (CRlt_asym R2 _ _ H4 H3).
+Qed.
+
+Lemma CRmorph_unique : forall (R1 R2 : ConstructiveReals)
+ (f g : ConstructiveRealsMorphism R1 R2)
+ (x : CRcarrier R1),
+ orderEq _ (CRlt R2) (CRmorph _ _ f x) (CRmorph _ _ g x).
+Proof.
+ split.
+ - intro abs. destruct (CR_Q_dense R2 _ _ abs) as [q [H H0]].
+ destruct (CRmorph_rat _ _ f q) as [H1 _].
+ apply (CRlt_le_trans R2 _ _ _ H) in H1. clear H.
+ apply CRmorph_increasing_inv in H1.
+ destruct (CRmorph_rat _ _ g q) as [_ H2].
+ apply (CRle_lt_trans R2 _ _ _ H2) in H0. clear H2.
+ apply CRmorph_increasing_inv in H0.
+ exact (CRlt_asym R1 _ _ H0 H1).
+ - intro abs. destruct (CR_Q_dense R2 _ _ abs) as [q [H H0]].
+ destruct (CRmorph_rat _ _ f q) as [_ H1].
+ apply (CRle_lt_trans R2 _ _ _ H1) in H0. clear H1.
+ apply CRmorph_increasing_inv in H0.
+ destruct (CRmorph_rat _ _ g q) as [H2 _].
+ apply (CRlt_le_trans R2 _ _ _ H) in H2. clear H.
+ apply CRmorph_increasing_inv in H2.
+ exact (CRlt_asym R1 _ _ H0 H2).
+Qed.
+
+
+(* The identity is the only endomorphism of constructive reals.
+ For any ConstructiveReals R1, R2 and any morphisms
+ f : R1 -> R2 and g : R2 -> R1,
+ f and g are isomorphisms and are inverses of each other. *)
+Lemma Endomorph_id : forall (R : ConstructiveReals) (f : ConstructiveRealsMorphism R R)
+ (x : CRcarrier R),
+ orderEq _ (CRlt R) (CRmorph _ _ f x) x.
+Proof.
+ split.
+ - intro abs. destruct (CR_Q_dense R _ _ abs) as [q [H0 H1]].
+ destruct (CRmorph_rat _ _ f q) as [H _].
+ apply (CRlt_le_trans R _ _ _ H0) in H. clear H0.
+ apply CRmorph_increasing_inv in H.
+ exact (CRlt_asym R _ _ H1 H).
+ - intro abs. destruct (CR_Q_dense R _ _ abs) as [q [H0 H1]].
+ destruct (CRmorph_rat _ _ f q) as [_ H].
+ apply (CRle_lt_trans R _ _ _ H) in H1. clear H.
+ apply CRmorph_increasing_inv in H1.
+ exact (CRlt_asym R _ _ H1 H0).
+Qed.
+
+Lemma CRmorph_proper : forall (R1 R2 : ConstructiveReals)
+ (f : ConstructiveRealsMorphism R1 R2)
+ (x y : CRcarrier R1),
+ orderEq _ (CRlt R1) x y
+ -> orderEq _ (CRlt R2) (CRmorph _ _ f x) (CRmorph _ _ f y).
+Proof.
+ split.
+ - intro abs. apply CRmorph_increasing_inv in abs.
+ destruct H. contradiction.
+ - intro abs. apply CRmorph_increasing_inv in abs.
+ destruct H. contradiction.
+Qed.
+
+Definition CRmorph_compose (R1 R2 R3 : ConstructiveReals)
+ (f : ConstructiveRealsMorphism R1 R2)
+ (g : ConstructiveRealsMorphism R2 R3)
+ : ConstructiveRealsMorphism R1 R3.
+Proof.
+ apply (Build_ConstructiveRealsMorphism
+ R1 R3 (fun x:CRcarrier R1 => CRmorph _ _ g (CRmorph _ _ f x))).
+ - intro q. apply (CReq_trans R3 _ (CRmorph R2 R3 g (CR_of_Q R2 q))).
+ apply CRmorph_proper. apply CRmorph_rat. apply CRmorph_rat.
+ - intros. apply CRmorph_increasing. apply CRmorph_increasing. exact H.
+Defined.
+
+Lemma CRmorph_le : forall (R1 R2 : ConstructiveReals)
+ (f : ConstructiveRealsMorphism R1 R2)
+ (x y : CRcarrier R1),
+ orderLe _ (CRlt R1) x y
+ -> orderLe _ (CRlt R2) (CRmorph _ _ f x) (CRmorph _ _ f y).
+Proof.
+ intros. intro abs. apply CRmorph_increasing_inv in abs. contradiction.
+Qed.
+
+Lemma CRmorph_le_inv : forall (R1 R2 : ConstructiveReals)
+ (f : ConstructiveRealsMorphism R1 R2)
+ (x y : CRcarrier R1),
+ orderLe _ (CRlt R2) (CRmorph _ _ f x) (CRmorph _ _ f y)
+ -> orderLe _ (CRlt R1) x y.
+Proof.
+ intros. intro abs. apply (CRmorph_increasing _ _ f) in abs. contradiction.
+Qed.
+
+Lemma CRmorph_zero : forall (R1 R2 : ConstructiveReals)
+ (f : ConstructiveRealsMorphism R1 R2),
+ orderEq _ (CRlt R2) (CRmorph _ _ f (CRzero R1)) (CRzero R2).
+Proof.
+ intros. apply (CReq_trans R2 _ (CRmorph _ _ f (CR_of_Q R1 0))).
+ apply CRmorph_proper. apply CReq_sym, CR_of_Q_zero.
+ apply (CReq_trans R2 _ (CR_of_Q R2 0)).
+ apply CRmorph_rat. apply CR_of_Q_zero.
+Qed.
+
+Lemma CRmorph_one : forall (R1 R2 : ConstructiveReals)
+ (f : ConstructiveRealsMorphism R1 R2),
+ orderEq _ (CRlt R2) (CRmorph _ _ f (CRone R1)) (CRone R2).
+Proof.
+ intros. apply (CReq_trans R2 _ (CRmorph _ _ f (CR_of_Q R1 1))).
+ apply CRmorph_proper. apply CReq_sym, CR_of_Q_one.
+ apply (CReq_trans R2 _ (CR_of_Q R2 1)).
+ apply CRmorph_rat. apply CR_of_Q_one.
+Qed.
+
+Lemma CRmorph_opp : forall (R1 R2 : ConstructiveReals)
+ (f : ConstructiveRealsMorphism R1 R2)
+ (x : CRcarrier R1),
+ orderEq _ (CRlt R2) (CRmorph _ _ f (CRopp R1 x))
+ (CRopp R2 (CRmorph _ _ f x)).
+Proof.
+ split.
+ - intro abs.
+ destruct (CR_Q_dense R2 _ _ abs) as [q [H H0]]. clear abs.
+ destruct (CRmorph_rat R1 R2 f q) as [H1 _].
+ apply (CRlt_le_trans R2 _ _ _ H) in H1. clear H.
+ apply CRmorph_increasing_inv in H1.
+ apply CRopp_gt_lt_contravar in H0.
+ destruct (CR_of_Q_opp R2 q) as [H2 _].
+ apply (CRlt_le_trans R2 _ _ _ H0) in H2. clear H0.
+ pose proof (CRopp_involutive R2 (CRmorph R1 R2 f x)) as [H _].
+ apply (CRle_lt_trans R2 _ _ _ H) in H2. clear H.
+ destruct (CRmorph_rat R1 R2 f (-q)) as [H _].
+ apply (CRlt_le_trans R2 _ _ _ H2) in H. clear H2.
+ apply CRmorph_increasing_inv in H.
+ destruct (CR_of_Q_opp R1 q) as [_ H2].
+ apply (CRlt_le_trans R1 _ _ _ H) in H2. clear H.
+ apply CRopp_gt_lt_contravar in H2.
+ pose proof (CRopp_involutive R1 (CR_of_Q R1 q)) as [H _].
+ apply (CRle_lt_trans R1 _ _ _ H) in H2. clear H.
+ exact (CRlt_asym R1 _ _ H1 H2).
+ - intro abs.
+ destruct (CR_Q_dense R2 _ _ abs) as [q [H H0]]. clear abs.
+ destruct (CRmorph_rat R1 R2 f q) as [_ H1].
+ apply (CRle_lt_trans R2 _ _ _ H1) in H0. clear H1.
+ apply CRmorph_increasing_inv in H0.
+ apply CRopp_gt_lt_contravar in H.
+ pose proof (CRopp_involutive R2 (CRmorph R1 R2 f x)) as [_ H1].
+ apply (CRlt_le_trans R2 _ _ _ H) in H1. clear H.
+ destruct (CR_of_Q_opp R2 q) as [_ H2].
+ apply (CRle_lt_trans R2 _ _ _ H2) in H1. clear H2.
+ destruct (CRmorph_rat R1 R2 f (-q)) as [_ H].
+ apply (CRle_lt_trans R2 _ _ _ H) in H1. clear H.
+ apply CRmorph_increasing_inv in H1.
+ destruct (CR_of_Q_opp R1 q) as [H2 _].
+ apply (CRle_lt_trans R1 _ _ _ H2) in H1. clear H2.
+ apply CRopp_gt_lt_contravar in H1.
+ pose proof (CRopp_involutive R1 (CR_of_Q R1 q)) as [_ H].
+ apply (CRlt_le_trans R1 _ _ _ H1) in H. clear H1.
+ exact (CRlt_asym R1 _ _ H0 H).
+Qed.
+
+Lemma CRplus_pos_rat_lt : forall (R : ConstructiveReals) (x : CRcarrier R) (q : Q),
+ Qlt 0 q -> CRlt R x (CRplus R x (CR_of_Q R q)).
+Proof.
+ intros.
+ apply (CRle_lt_trans R _ (CRplus R x (CRzero R))). apply CRplus_0_r.
+ apply CRplus_lt_compat_l.
+ apply (CRle_lt_trans R _ (CR_of_Q R 0)). apply CR_of_Q_zero.
+ apply CR_of_Q_lt. exact H.
+Defined.
+
+Lemma CRplus_neg_rat_lt : forall (R : ConstructiveReals) (x : CRcarrier R) (q : Q),
+ Qlt q 0 -> CRlt R (CRplus R x (CR_of_Q R q)) x.
+Proof.
+ intros.
+ apply (CRlt_le_trans R _ (CRplus R x (CRzero R))). 2: apply CRplus_0_r.
+ apply CRplus_lt_compat_l.
+ apply (CRlt_le_trans R _ (CR_of_Q R 0)).
+ apply CR_of_Q_lt. exact H. apply CR_of_Q_zero.
+Qed.
+
+Lemma CRmorph_plus_rat : forall (R1 R2 : ConstructiveReals)
+ (f : ConstructiveRealsMorphism R1 R2)
+ (x : CRcarrier R1) (q : Q),
+ orderEq _ (CRlt R2) (CRmorph _ _ f (CRplus R1 x (CR_of_Q R1 q)))
+ (CRplus R2 (CRmorph _ _ f x) (CR_of_Q R2 q)).
+Proof.
+ split.
+ - intro abs.
+ destruct (CR_Q_dense R2 _ _ abs) as [r [H H0]]. clear abs.
+ destruct (CRmorph_rat _ _ f r) as [H1 _].
+ apply (CRlt_le_trans R2 _ _ _ H) in H1. clear H.
+ apply CRmorph_increasing_inv in H1.
+ apply (CRlt_asym R1 _ _ H1). clear H1.
+ apply (CRplus_lt_reg_r R1 (CRopp R1 (CR_of_Q R1 q))).
+ apply (CRlt_le_trans R1 _ x).
+ apply (CRle_lt_trans R1 _ (CR_of_Q R1 (r-q))).
+ apply (CRle_trans R1 _ (CRplus R1 (CR_of_Q R1 r) (CR_of_Q R1 (-q)))).
+ apply CRplus_le_compat_l. destruct (CR_of_Q_opp R1 q). exact H.
+ destruct (CR_of_Q_plus R1 r (-q)). exact H.
+ apply (CRmorph_increasing_inv _ _ f).
+ apply (CRle_lt_trans R2 _ (CR_of_Q R2 (r - q))).
+ apply CRmorph_rat.
+ apply (CRplus_lt_reg_r R2 (CR_of_Q R2 q)).
+ apply (CRle_lt_trans R2 _ (CR_of_Q R2 r)). 2: exact H0.
+ intro H.
+ destruct (CR_of_Q_plus R2 (r-q) q) as [H1 _].
+ apply (CRlt_le_trans R2 _ _ _ H) in H1. clear H.
+ apply lt_CR_of_Q in H1. ring_simplify in H1.
+ exact (Qlt_not_le _ _ H1 (Qle_refl _)).
+ destruct (CRisRing R1).
+ apply (CRle_trans R1 _ (CRplus R1 x (CRplus R1 (CR_of_Q R1 q) (CRopp R1 (CR_of_Q R1 q))))).
+ apply (CRle_trans R1 _ (CRplus R1 x (CRzero R1))).
+ destruct (CRplus_0_r R1 x). exact H.
+ apply CRplus_le_compat_l. destruct (Ropp_def (CR_of_Q R1 q)). exact H.
+ destruct (Radd_assoc x (CR_of_Q R1 q) (CRopp R1 (CR_of_Q R1 q))).
+ exact H1.
+ - intro abs.
+ destruct (CR_Q_dense R2 _ _ abs) as [r [H H0]]. clear abs.
+ destruct (CRmorph_rat _ _ f r) as [_ H1].
+ apply (CRle_lt_trans R2 _ _ _ H1) in H0. clear H1.
+ apply CRmorph_increasing_inv in H0.
+ apply (CRlt_asym R1 _ _ H0). clear H0.
+ apply (CRplus_lt_reg_r R1 (CRopp R1 (CR_of_Q R1 q))).
+ apply (CRle_lt_trans R1 _ x).
+ destruct (CRisRing R1).
+ apply (CRle_trans R1 _ (CRplus R1 x (CRplus R1 (CR_of_Q R1 q) (CRopp R1 (CR_of_Q R1 q))))).
+ destruct (Radd_assoc x (CR_of_Q R1 q) (CRopp R1 (CR_of_Q R1 q))).
+ exact H0.
+ apply (CRle_trans R1 _ (CRplus R1 x (CRzero R1))).
+ apply CRplus_le_compat_l. destruct (Ropp_def (CR_of_Q R1 q)). exact H1.
+ destruct (CRplus_0_r R1 x). exact H1.
+ apply (CRlt_le_trans R1 _ (CR_of_Q R1 (r-q))).
+ apply (CRmorph_increasing_inv _ _ f).
+ apply (CRlt_le_trans R2 _ (CR_of_Q R2 (r - q))).
+ apply (CRplus_lt_reg_r R2 (CR_of_Q R2 q)).
+ apply (CRlt_le_trans R2 _ _ _ H).
+ 2: apply CRmorph_rat.
+ apply (CRle_trans R2 _ (CR_of_Q R2 (r-q+q))).
+ intro abs. apply lt_CR_of_Q in abs. ring_simplify in abs.
+ exact (Qlt_not_le _ _ abs (Qle_refl _)).
+ destruct (CR_of_Q_plus R2 (r-q) q). exact H1.
+ apply (CRle_trans R1 _ (CRplus R1 (CR_of_Q R1 r) (CR_of_Q R1 (-q)))).
+ destruct (CR_of_Q_plus R1 r (-q)). exact H1.
+ apply CRplus_le_compat_l. destruct (CR_of_Q_opp R1 q). exact H1.
+Qed.
+
+Lemma CRmorph_plus : forall (R1 R2 : ConstructiveReals)
+ (f : ConstructiveRealsMorphism R1 R2)
+ (x y : CRcarrier R1),
+ orderEq _ (CRlt R2) (CRmorph _ _ f (CRplus R1 x y))
+ (CRplus R2 (CRmorph _ _ f x) (CRmorph _ _ f y)).
+Proof.
+ intros R1 R2 f.
+ assert (forall (x y : CRcarrier R1),
+ orderLe _ (CRlt R2) (CRplus R2 (CRmorph R1 R2 f x) (CRmorph R1 R2 f y))
+ (CRmorph R1 R2 f (CRplus R1 x y))).
+ { intros x y abs. destruct (CR_Q_dense R2 _ _ abs) as [r [H H0]]. clear abs.
+ destruct (CRmorph_rat _ _ f r) as [H1 _].
+ apply (CRlt_le_trans R2 _ _ _ H) in H1. clear H.
+ apply CRmorph_increasing_inv in H1.
+ apply (CRlt_asym R1 _ _ H1). clear H1.
+ destruct (CR_Q_dense R2 _ _ H0) as [q [H2 H3]].
+ apply lt_CR_of_Q in H2.
+ assert (Qlt (r-q) 0) as epsNeg.
+ { apply (Qplus_lt_r _ _ q). ring_simplify. exact H2. }
+ destruct (CR_Q_dense R1 _ _ (CRplus_neg_rat_lt R1 x (r-q) epsNeg))
+ as [s [H4 H5]].
+ apply (CRlt_trans R1 _ (CRplus R1 (CR_of_Q R1 s) y)).
+ 2: apply CRplus_lt_compat_r, H5.
+ apply (CRmorph_increasing_inv _ _ f).
+ apply (CRlt_le_trans R2 _ (CRplus R2 (CR_of_Q R2 s) (CRmorph _ _ f y))).
+ apply (CRmorph_increasing _ _ f) in H4.
+ destruct (CRmorph_plus_rat _ _ f x (r-q)) as [H _].
+ apply (CRle_lt_trans R2 _ _ _ H) in H4. clear H.
+ destruct (CRmorph_rat _ _ f s) as [_ H1].
+ apply (CRlt_le_trans R2 _ _ _ H4) in H1. clear H4.
+ apply (CRlt_trans R2 _ (CRplus R2 (CRplus R2 (CRmorph R1 R2 f x) (CR_of_Q R2 (r - q)))
+ (CRmorph R1 R2 f y))).
+ 2: apply CRplus_lt_compat_r, H1.
+ apply (CRlt_le_trans R2 _ (CRplus R2 (CRplus R2 (CR_of_Q R2 (r - q)) (CRmorph R1 R2 f x))
+ (CRmorph R1 R2 f y))).
+ apply (CRlt_le_trans R2 _ (CRplus R2 (CR_of_Q R2 (r - q))
+ (CRplus R2 (CRmorph R1 R2 f x) (CRmorph R1 R2 f y)))).
+ apply (CRle_lt_trans R2 _ (CRplus R2 (CR_of_Q R2 (r - q)) (CR_of_Q R2 q))).
+ 2: apply CRplus_lt_compat_l, H3.
+ intro abs.
+ destruct (CR_of_Q_plus R2 (r-q) q) as [_ H4].
+ apply (CRle_lt_trans R2 _ _ _ H4) in abs. clear H4.
+ destruct (CRmorph_rat _ _ f r) as [_ H4].
+ apply (CRlt_le_trans R2 _ _ _ abs) in H4. clear abs.
+ apply lt_CR_of_Q in H4. ring_simplify in H4.
+ exact (Qlt_not_le _ _ H4 (Qle_refl _)).
+ destruct (CRisRing R2); apply Radd_assoc.
+ apply CRplus_le_compat_r. destruct (CRisRing R2).
+ destruct (Radd_comm (CRmorph R1 R2 f x) (CR_of_Q R2 (r - q))).
+ exact H.
+ intro abs.
+ destruct (CRmorph_plus_rat _ _ f y s) as [H _]. apply H. clear H.
+ apply (CRlt_le_trans R2 _ (CRplus R2 (CR_of_Q R2 s) (CRmorph R1 R2 f y))).
+ apply (CRle_lt_trans R2 _ (CRmorph R1 R2 f (CRplus R1 (CR_of_Q R1 s) y))).
+ apply CRmorph_proper. destruct (CRisRing R1); apply Radd_comm.
+ exact abs. destruct (CRisRing R2); apply Radd_comm. }
+ split.
+ - apply H.
+ - specialize (H (CRplus R1 x y) (CRopp R1 y)).
+ intro abs. apply H. clear H.
+ apply (CRle_lt_trans R2 _ (CRmorph R1 R2 f x)).
+ apply CRmorph_proper. destruct (CRisRing R1).
+ apply (CReq_trans R1 _ (CRplus R1 x (CRplus R1 y (CRopp R1 y)))).
+ apply CReq_sym, Radd_assoc.
+ apply (CReq_trans R1 _ (CRplus R1 x (CRzero R1))). 2: apply CRplus_0_r.
+ destruct (CRisRingExt R1). apply Radd_ext.
+ apply CReq_refl. apply Ropp_def.
+ apply (CRplus_lt_reg_r R2 (CRmorph R1 R2 f y)).
+ apply (CRlt_le_trans R2 _ _ _ abs). clear abs.
+ apply (CRle_trans R2 _ (CRplus R2 (CRmorph R1 R2 f (CRplus R1 x y)) (CRzero R2))).
+ destruct (CRplus_0_r R2 (CRmorph R1 R2 f (CRplus R1 x y))). exact H.
+ apply (CRle_trans R2 _ (CRplus R2 (CRmorph R1 R2 f (CRplus R1 x y))
+ (CRplus R2 (CRmorph R1 R2 f (CRopp R1 y)) (CRmorph R1 R2 f y)))).
+ apply CRplus_le_compat_l.
+ apply (CRle_trans R2 _ (CRplus R2 (CRopp R2 (CRmorph R1 R2 f y)) (CRmorph R1 R2 f y))).
+ destruct (CRplus_opp_l R2 (CRmorph R1 R2 f y)). exact H.
+ apply CRplus_le_compat_r. destruct (CRmorph_opp _ _ f y). exact H.
+ destruct (CRisRing R2).
+ destruct (Radd_assoc (CRmorph R1 R2 f (CRplus R1 x y))
+ (CRmorph R1 R2 f (CRopp R1 y)) (CRmorph R1 R2 f y)).
+ exact H0.
+Qed.
+
+Lemma CRmorph_mult_pos : forall (R1 R2 : ConstructiveReals)
+ (f : ConstructiveRealsMorphism R1 R2)
+ (x : CRcarrier R1) (n : nat),
+ orderEq _ (CRlt R2) (CRmorph _ _ f (CRmult R1 x (CR_of_Q R1 (Z.of_nat n # 1))))
+ (CRmult R2 (CRmorph _ _ f x) (CR_of_Q R2 (Z.of_nat n # 1))).
+Proof.
+ induction n.
+ - simpl. destruct (CRisRingExt R1).
+ apply (CReq_trans R2 _ (CRzero R2)).
+ + apply (CReq_trans R2 _ (CRmorph R1 R2 f (CRzero R1))).
+ 2: apply CRmorph_zero. apply CRmorph_proper.
+ apply (CReq_trans R1 _ (CRmult R1 x (CRzero R1))).
+ 2: apply CRmult_0_r. apply Rmul_ext. apply CReq_refl. apply CR_of_Q_zero.
+ + apply (CReq_trans R2 _ (CRmult R2 (CRmorph R1 R2 f x) (CRzero R2))).
+ apply CReq_sym, CRmult_0_r. destruct (CRisRingExt R2).
+ apply Rmul_ext0. apply CReq_refl. apply CReq_sym, CR_of_Q_zero.
+ - destruct (CRisRingExt R1), (CRisRingExt R2).
+ apply (CReq_trans
+ R2 _ (CRmorph R1 R2 f (CRplus R1 x (CRmult R1 x (CR_of_Q R1 (Z.of_nat n # 1)))))).
+ apply CRmorph_proper.
+ apply (CReq_trans R1 _ (CRmult R1 x (CRplus R1 (CRone R1) (CR_of_Q R1 (Z.of_nat n # 1))))).
+ apply Rmul_ext. apply CReq_refl.
+ apply (CReq_trans R1 _ (CR_of_Q R1 (1 + (Z.of_nat n # 1)))).
+ apply CR_of_Q_proper. rewrite Nat2Z.inj_succ. unfold Z.succ.
+ rewrite Z.add_comm. rewrite Qinv_plus_distr. reflexivity.
+ apply (CReq_trans R1 _ (CRplus R1 (CR_of_Q R1 1) (CR_of_Q R1 (Z.of_nat n # 1)))).
+ apply CR_of_Q_plus. apply Radd_ext. apply CR_of_Q_one. apply CReq_refl.
+ apply (CReq_trans R1 _ (CRplus R1 (CRmult R1 x (CRone R1))
+ (CRmult R1 x (CR_of_Q R1 (Z.of_nat n # 1))))).
+ apply CRmult_plus_distr_l. apply Radd_ext. apply CRmult_1_r. apply CReq_refl.
+ apply (CReq_trans R2 _ (CRplus R2 (CRmorph R1 R2 f x)
+ (CRmorph R1 R2 f (CRmult R1 x (CR_of_Q R1 (Z.of_nat n # 1)))))).
+ apply CRmorph_plus.
+ apply (CReq_trans R2 _ (CRplus R2 (CRmorph R1 R2 f x)
+ (CRmult R2 (CRmorph R1 R2 f x) (CR_of_Q R2 (Z.of_nat n # 1))))).
+ apply Radd_ext0. apply CReq_refl. exact IHn.
+ apply (CReq_trans R2 _ (CRmult R2 (CRmorph R1 R2 f x) (CRplus R2 (CRone R2) (CR_of_Q R2 (Z.of_nat n # 1))))).
+ apply (CReq_trans R2 _ (CRplus R2 (CRmult R2 (CRmorph R1 R2 f x) (CRone R2))
+ (CRmult R2 (CRmorph _ _ f x) (CR_of_Q R2 (Z.of_nat n # 1))))).
+ apply Radd_ext0. 2: apply CReq_refl. apply CReq_sym, CRmult_1_r.
+ apply CReq_sym, CRmult_plus_distr_l.
+ apply Rmul_ext0. apply CReq_refl.
+ apply (CReq_trans R2 _ (CR_of_Q R2 (1 + (Z.of_nat n # 1)))).
+ apply (CReq_trans R2 _ (CRplus R2 (CR_of_Q R2 1) (CR_of_Q R2 (Z.of_nat n # 1)))).
+ apply Radd_ext0. apply CReq_sym, CR_of_Q_one. apply CReq_refl.
+ apply CReq_sym, CR_of_Q_plus.
+ apply CR_of_Q_proper. rewrite Nat2Z.inj_succ. unfold Z.succ.
+ rewrite Z.add_comm. rewrite Qinv_plus_distr. reflexivity.
+Qed.
+
+Lemma NatOfZ : forall n : Z, { p : nat | n = Z.of_nat p \/ n = Z.opp (Z.of_nat p) }.
+Proof.
+ intros [|p|n].
+ - exists O. left. reflexivity.
+ - exists (Pos.to_nat p). left. rewrite positive_nat_Z. reflexivity.
+ - exists (Pos.to_nat n). right. rewrite positive_nat_Z. reflexivity.
+Qed.
+
+Lemma CRmorph_mult_int : forall (R1 R2 : ConstructiveReals)
+ (f : ConstructiveRealsMorphism R1 R2)
+ (x : CRcarrier R1) (n : Z),
+ orderEq _ (CRlt R2) (CRmorph _ _ f (CRmult R1 x (CR_of_Q R1 (n # 1))))
+ (CRmult R2 (CRmorph _ _ f x) (CR_of_Q R2 (n # 1))).
+Proof.
+ intros. destruct (NatOfZ n) as [p [pos|neg]].
+ - subst n. apply CRmorph_mult_pos.
+ - subst n.
+ apply (CReq_trans R2 _ (CRopp R2 (CRmorph R1 R2 f (CRmult R1 x (CR_of_Q R1 (Z.of_nat p # 1)))))).
+ + apply (CReq_trans R2 _ (CRmorph R1 R2 f (CRopp R1 (CRmult R1 x (CR_of_Q R1 (Z.of_nat p # 1)))))).
+ 2: apply CRmorph_opp. apply CRmorph_proper.
+ apply (CReq_trans R1 _ (CRmult R1 x (CR_of_Q R1 (- (Z.of_nat p # 1))))).
+ destruct (CRisRingExt R1). apply Rmul_ext. apply CReq_refl.
+ apply CR_of_Q_proper. reflexivity.
+ apply (CReq_trans R1 _ (CRmult R1 x (CRopp R1 (CR_of_Q R1 (Z.of_nat p # 1))))).
+ destruct (CRisRingExt R1). apply Rmul_ext. apply CReq_refl.
+ apply CR_of_Q_opp. apply CReq_sym, CRopp_mult_distr_r.
+ + apply (CReq_trans R2 _ (CRopp R2 (CRmult R2 (CRmorph R1 R2 f x) (CR_of_Q R2 (Z.of_nat p # 1))))).
+ destruct (CRisRingExt R2). apply Ropp_ext. apply CRmorph_mult_pos.
+ apply (CReq_trans R2 _ (CRmult R2 (CRmorph R1 R2 f x) (CRopp R2 (CR_of_Q R2 (Z.of_nat p # 1))))).
+ apply CRopp_mult_distr_r. destruct (CRisRingExt R2).
+ apply Rmul_ext. apply CReq_refl.
+ apply (CReq_trans R2 _ (CR_of_Q R2 (- (Z.of_nat p # 1)))).
+ apply CReq_sym, CR_of_Q_opp. apply CR_of_Q_proper. reflexivity.
+Qed.
+
+Lemma CRmorph_mult_inv : forall (R1 R2 : ConstructiveReals)
+ (f : ConstructiveRealsMorphism R1 R2)
+ (x : CRcarrier R1) (p : positive),
+ orderEq _ (CRlt R2) (CRmorph _ _ f (CRmult R1 x (CR_of_Q R1 (1 # p))))
+ (CRmult R2 (CRmorph _ _ f x) (CR_of_Q R2 (1 # p))).
+Proof.
+ intros. apply (CRmult_eq_reg_r R2 (CR_of_Q R2 (Z.pos p # 1))).
+ left. apply (CRle_lt_trans R2 _ (CR_of_Q R2 0)).
+ apply CR_of_Q_zero. apply CR_of_Q_lt. reflexivity.
+ apply (CReq_trans R2 _ (CRmorph _ _ f x)).
+ - apply (CReq_trans
+ R2 _ (CRmorph R1 R2 f (CRmult R1 (CRmult R1 x (CR_of_Q R1 (1 # p)))
+ (CR_of_Q R1 (Z.pos p # 1))))).
+ apply CReq_sym, CRmorph_mult_int. apply CRmorph_proper.
+ apply (CReq_trans
+ R1 _ (CRmult R1 x (CRmult R1 (CR_of_Q R1 (1 # p))
+ (CR_of_Q R1 (Z.pos p # 1))))).
+ destruct (CRisRing R1). apply CReq_sym, Rmul_assoc.
+ apply (CReq_trans R1 _ (CRmult R1 x (CRone R1))).
+ apply (Rmul_ext (CRisRingExt R1)). apply CReq_refl.
+ apply (CReq_trans R1 _ (CR_of_Q R1 ((1#p) * (Z.pos p # 1)))).
+ apply CReq_sym, CR_of_Q_mult.
+ apply (CReq_trans R1 _ (CR_of_Q R1 1)).
+ apply CR_of_Q_proper. reflexivity. apply CR_of_Q_one.
+ apply CRmult_1_r.
+ - apply (CReq_trans R2 _ (CRmult R2 (CRmorph R1 R2 f x)
+ (CRmult R2 (CR_of_Q R2 (1 # p)) (CR_of_Q R2 (Z.pos p # 1))))).
+ 2: apply (Rmul_assoc (CRisRing R2)).
+ apply (CReq_trans R2 _ (CRmult R2 (CRmorph R1 R2 f x) (CRone R2))).
+ apply CReq_sym, CRmult_1_r.
+ apply (Rmul_ext (CRisRingExt R2)). apply CReq_refl.
+ apply (CReq_trans R2 _ (CR_of_Q R2 1)).
+ apply CReq_sym, CR_of_Q_one.
+ apply (CReq_trans R2 _ (CR_of_Q R2 ((1#p)*(Z.pos p # 1)))).
+ apply CR_of_Q_proper. reflexivity. apply CR_of_Q_mult.
+Qed.
+
+Lemma CRmorph_mult_rat : forall (R1 R2 : ConstructiveReals)
+ (f : ConstructiveRealsMorphism R1 R2)
+ (x : CRcarrier R1) (q : Q),
+ orderEq _ (CRlt R2) (CRmorph _ _ f (CRmult R1 x (CR_of_Q R1 q)))
+ (CRmult R2 (CRmorph _ _ f x) (CR_of_Q R2 q)).
+Proof.
+ intros. destruct q as [a b].
+ apply (CReq_trans R2 _ (CRmult R2 (CRmorph _ _ f (CRmult R1 x (CR_of_Q R1 (a # 1))))
+ (CR_of_Q R2 (1 # b)))).
+ - apply (CReq_trans
+ R2 _ (CRmorph R1 R2 f (CRmult R1 (CRmult R1 x (CR_of_Q R1 (a # 1)))
+ (CR_of_Q R1 (1 # b))))).
+ 2: apply CRmorph_mult_inv. apply CRmorph_proper.
+ apply (CReq_trans R1 _ (CRmult R1 x (CRmult R1 (CR_of_Q R1 (a # 1))
+ (CR_of_Q R1 (1 # b))))).
+ apply (Rmul_ext (CRisRingExt R1)). apply CReq_refl.
+ apply (CReq_trans R1 _ (CR_of_Q R1 ((a#1)*(1#b)))).
+ apply CR_of_Q_proper. unfold Qeq; simpl. rewrite Z.mul_1_r. reflexivity.
+ apply CR_of_Q_mult.
+ apply (Rmul_assoc (CRisRing R1)).
+ - apply (CReq_trans R2 _ (CRmult R2 (CRmult R2 (CRmorph _ _ f x) (CR_of_Q R2 (a # 1)))
+ (CR_of_Q R2 (1 # b)))).
+ apply (Rmul_ext (CRisRingExt R2)). apply CRmorph_mult_int.
+ apply CReq_refl.
+ apply (CReq_trans R2 _ (CRmult R2 (CRmorph R1 R2 f x)
+ (CRmult R2 (CR_of_Q R2 (a # 1)) (CR_of_Q R2 (1 # b))))).
+ apply CReq_sym, (Rmul_assoc (CRisRing R2)).
+ apply (Rmul_ext (CRisRingExt R2)). apply CReq_refl.
+ apply (CReq_trans R2 _ (CR_of_Q R2 ((a#1)*(1#b)))).
+ apply CReq_sym, CR_of_Q_mult.
+ apply CR_of_Q_proper. unfold Qeq; simpl. rewrite Z.mul_1_r. reflexivity.
+Qed.
+
+Lemma CRmorph_mult_pos_pos_le : forall (R1 R2 : ConstructiveReals)
+ (f : ConstructiveRealsMorphism R1 R2)
+ (x y : CRcarrier R1),
+ CRlt R1 (CRzero R1) y
+ -> orderLe _ (CRlt R2) (CRmult R2 (CRmorph _ _ f x) (CRmorph _ _ f y))
+ (CRmorph _ _ f (CRmult R1 x y)).
+Proof.
+ intros. intro abs. destruct (CR_Q_dense R2 _ _ abs) as [q [H1 H2]].
+ destruct (CRmorph_rat _ _ f q) as [H3 _].
+ apply (CRlt_le_trans R2 _ _ _ H1) in H3. clear H1.
+ apply CRmorph_increasing_inv in H3.
+ apply (CRlt_asym R1 _ _ H3). clear H3.
+ destruct (CR_Q_dense R2 _ _ H2) as [r [H1 H3]].
+ apply lt_CR_of_Q in H1.
+ destruct (CR_archimedean R1 y) as [A Amaj].
+ assert (/ ((r - q) * (1 # A)) * (q - r) == - (Z.pos A # 1)) as diveq.
+ { rewrite Qinv_mult_distr. setoid_replace (q-r) with (-1*(r-q)).
+ field_simplify. reflexivity. 2: field.
+ split. intro H4. inversion H4. intro H4.
+ apply Qlt_minus_iff in H1. rewrite H4 in H1. inversion H1. }
+ destruct (CR_Q_dense R1 (CRplus R1 x (CR_of_Q R1 ((q-r) * (1#A)))) x)
+ as [s [H4 H5]].
+ - apply (CRlt_le_trans R1 _ (CRplus R1 x (CRzero R1))).
+ 2: apply CRplus_0_r. apply CRplus_lt_compat_l.
+ apply (CRplus_lt_reg_l R1 (CR_of_Q R1 ((r-q) * (1#A)))).
+ apply (CRle_lt_trans R1 _ (CRzero R1)).
+ apply (CRle_trans R1 _ (CR_of_Q R1 ((r-q)*(1#A) + (q-r)*(1#A)))).
+ destruct (CR_of_Q_plus R1 ((r-q)*(1#A)) ((q-r)*(1#A))).
+ exact H0. apply (CRle_trans R1 _ (CR_of_Q R1 0)).
+ 2: destruct (CR_of_Q_zero R1); exact H4.
+ intro H4. apply lt_CR_of_Q in H4. ring_simplify in H4.
+ inversion H4.
+ apply (CRlt_le_trans R1 _ (CR_of_Q R1 ((r - q) * (1 # A)))).
+ 2: apply CRplus_0_r.
+ apply (CRle_lt_trans R1 _ (CR_of_Q R1 0)).
+ apply CR_of_Q_zero. apply CR_of_Q_lt.
+ rewrite <- (Qmult_0_r (r-q)). apply Qmult_lt_l.
+ apply Qlt_minus_iff in H1. exact H1. reflexivity.
+ - apply (CRmorph_increasing _ _ f) in H4.
+ destruct (CRmorph_plus _ _ f x (CR_of_Q R1 ((q-r) * (1#A)))) as [H6 _].
+ apply (CRle_lt_trans R2 _ _ _ H6) in H4. clear H6.
+ destruct (CRmorph_rat _ _ f s) as [_ H6].
+ apply (CRlt_le_trans R2 _ _ _ H4) in H6. clear H4.
+ apply (CRmult_lt_compat_r R2 (CRmorph _ _ f y)) in H6.
+ destruct (Rdistr_l (CRisRing R2) (CRmorph _ _ f x)
+ (CRmorph R1 R2 f (CR_of_Q R1 ((q-r) * (1#A))))
+ (CRmorph _ _ f y)) as [H4 _].
+ apply (CRle_lt_trans R2 _ _ _ H4) in H6. clear H4.
+ apply (CRle_lt_trans R1 _ (CRmult R1 (CR_of_Q R1 s) y)).
+ 2: apply CRmult_lt_compat_r. 2: exact H. 2: exact H5.
+ apply (CRmorph_le_inv _ _ f).
+ apply (CRle_trans R2 _ (CR_of_Q R2 q)).
+ destruct (CRmorph_rat _ _ f q). exact H4.
+ apply (CRle_trans R2 _ (CRmult R2 (CR_of_Q R2 s) (CRmorph _ _ f y))).
+ apply (CRle_trans R2 _ (CRplus R2 (CRmult R2 (CRmorph _ _ f x) (CRmorph _ _ f y))
+ (CR_of_Q R2 (q-r)))).
+ apply (CRle_trans R2 _ (CRplus R2 (CR_of_Q R2 r) (CR_of_Q R2 (q - r)))).
+ + apply (CRle_trans R2 _ (CR_of_Q R2 (r + (q-r)))).
+ intro H4. apply lt_CR_of_Q in H4. ring_simplify in H4.
+ exact (Qlt_not_le q q H4 (Qle_refl q)).
+ destruct (CR_of_Q_plus R2 r (q-r)). exact H4.
+ + apply CRplus_le_compat_r. intro H4.
+ apply (CRlt_asym R2 _ _ H3). exact H4.
+ + intro H4. apply (CRlt_asym R2 _ _ H4). clear H4.
+ apply (CRlt_trans_flip R2 _ _ _ H6). clear H6.
+ apply CRplus_lt_compat_l.
+ apply (CRlt_le_trans R2 _ (CRmult R2 (CR_of_Q R2 ((q - r) * (1 # A))) (CRmorph R1 R2 f y))).
+ apply (CRmult_lt_reg_l R2 (CR_of_Q R2 (/((r-q)*(1#A))))).
+ apply (CRle_lt_trans R2 _ (CR_of_Q R2 0)). apply CR_of_Q_zero.
+ apply CR_of_Q_lt, Qinv_lt_0_compat.
+ rewrite <- (Qmult_0_r (r-q)). apply Qmult_lt_l.
+ apply Qlt_minus_iff in H1. exact H1. reflexivity.
+ apply (CRle_lt_trans R2 _ (CRopp R2 (CR_of_Q R2 (Z.pos A # 1)))).
+ apply (CRle_trans R2 _ (CR_of_Q R2 (-(Z.pos A # 1)))).
+ apply (CRle_trans R2 _ (CR_of_Q R2 ((/ ((r - q) * (1 # A))) * (q - r)))).
+ destruct (CR_of_Q_mult R2 (/ ((r - q) * (1 # A))) (q - r)).
+ exact H0. destruct (CR_of_Q_proper R2 (/ ((r - q) * (1 # A)) * (q - r))
+ (-(Z.pos A # 1))).
+ exact diveq. intro H7. apply lt_CR_of_Q in H7.
+ rewrite diveq in H7. exact (Qlt_not_le _ _ H7 (Qle_refl _)).
+ destruct (CR_of_Q_opp R2 (Z.pos A # 1)). exact H4.
+ apply (CRlt_le_trans R2 _ (CRopp R2 (CRmorph _ _ f y))).
+ apply CRopp_gt_lt_contravar.
+ apply (CRlt_le_trans R2 _ (CRmorph _ _ f (CR_of_Q R1 (Z.pos A # 1)))).
+ apply CRmorph_increasing. exact Amaj.
+ destruct (CRmorph_rat _ _ f (Z.pos A # 1)). exact H4.
+ apply (CRle_trans R2 _ (CRmult R2 (CRopp R2 (CRone R2)) (CRmorph _ _ f y))).
+ apply (CRle_trans R2 _ (CRopp R2 (CRmult R2 (CRone R2) (CRmorph R1 R2 f y)))).
+ destruct (Ropp_ext (CRisRingExt R2) (CRmorph _ _ f y)
+ (CRmult R2 (CRone R2) (CRmorph R1 R2 f y))).
+ apply CReq_sym, (Rmul_1_l (CRisRing R2)). exact H4.
+ destruct (CRopp_mult_distr_l R2 (CRone R2) (CRmorph _ _ f y)). exact H4.
+ apply (CRle_trans R2 _ (CRmult R2 (CRmult R2 (CR_of_Q R2 (/ ((r - q) * (1 # A))))
+ (CR_of_Q R2 ((q - r) * (1 # A))))
+ (CRmorph R1 R2 f y))).
+ apply CRmult_le_compat_r.
+ apply (CRle_lt_trans R2 _ (CRmorph _ _ f (CRzero R1))).
+ apply CRmorph_zero. apply CRmorph_increasing. exact H.
+ apply (CRle_trans R2 _ (CR_of_Q R2 ((/ ((r - q) * (1 # A)))
+ * ((q - r) * (1 # A))))).
+ apply (CRle_trans R2 _ (CR_of_Q R2 (-1))).
+ apply (CRle_trans R2 _ (CRopp R2 (CR_of_Q R2 1))).
+ destruct (Ropp_ext (CRisRingExt R2) (CRone R2) (CR_of_Q R2 1)).
+ apply CReq_sym, CR_of_Q_one. exact H4.
+ destruct (CR_of_Q_opp R2 1). exact H0.
+ destruct (CR_of_Q_proper R2 (-1) (/ ((r - q) * (1 # A)) * ((q - r) * (1 # A)))).
+ field. split.
+ intro H4. inversion H4. intro H4. apply Qlt_minus_iff in H1.
+ rewrite H4 in H1. inversion H1. exact H4.
+ destruct (CR_of_Q_mult R2 (/ ((r - q) * (1 # A))) ((q - r) * (1 # A))).
+ exact H4.
+ destruct (Rmul_assoc (CRisRing R2) (CR_of_Q R2 (/ ((r - q) * (1 # A))))
+ (CR_of_Q R2 ((q - r) * (1 # A)))
+ (CRmorph R1 R2 f y)).
+ exact H0.
+ apply CRmult_le_compat_r.
+ apply (CRle_lt_trans R2 _ (CRmorph _ _ f (CRzero R1))).
+ apply CRmorph_zero. apply CRmorph_increasing. exact H.
+ destruct (CRmorph_rat _ _ f ((q - r) * (1 # A))). exact H0.
+ + apply (CRle_trans R2 _ (CRmorph _ _ f (CRmult R1 y (CR_of_Q R1 s)))).
+ apply (CRle_trans R2 _ (CRmult R2 (CRmorph R1 R2 f y) (CR_of_Q R2 s))).
+ destruct (Rmul_comm (CRisRing R2) (CRmorph R1 R2 f y) (CR_of_Q R2 s)).
+ exact H0.
+ destruct (CRmorph_mult_rat _ _ f y s). exact H0.
+ destruct (CRmorph_proper _ _ f (CRmult R1 y (CR_of_Q R1 s))
+ (CRmult R1 (CR_of_Q R1 s) y)).
+ apply (Rmul_comm (CRisRing R1)). exact H4.
+ + apply (CRle_lt_trans R2 _ (CRmorph _ _ f (CRzero R1))).
+ apply CRmorph_zero. apply CRmorph_increasing. exact H.
+Qed.
+
+Lemma CRmorph_mult_pos_pos : forall (R1 R2 : ConstructiveReals)
+ (f : ConstructiveRealsMorphism R1 R2)
+ (x y : CRcarrier R1),
+ CRlt R1 (CRzero R1) y
+ -> orderEq _ (CRlt R2) (CRmorph _ _ f (CRmult R1 x y))
+ (CRmult R2 (CRmorph _ _ f x) (CRmorph _ _ f y)).
+Proof.
+ split. apply CRmorph_mult_pos_pos_le. exact H.
+ intro abs. destruct (CR_Q_dense R2 _ _ abs) as [q [H1 H2]].
+ destruct (CRmorph_rat _ _ f q) as [_ H3].
+ apply (CRle_lt_trans R2 _ _ _ H3) in H2. clear H3.
+ apply CRmorph_increasing_inv in H2.
+ apply (CRlt_asym R1 _ _ H2). clear H2.
+ destruct (CR_Q_dense R2 _ _ H1) as [r [H2 H3]].
+ apply lt_CR_of_Q in H3.
+ destruct (CR_archimedean R1 y) as [A Amaj].
+ destruct (CR_Q_dense R1 x (CRplus R1 x (CR_of_Q R1 ((q-r) * (1#A)))))
+ as [s [H4 H5]].
+ - apply (CRle_lt_trans R1 _ (CRplus R1 x (CRzero R1))).
+ apply CRplus_0_r. apply CRplus_lt_compat_l.
+ apply (CRle_lt_trans R1 _ (CR_of_Q R1 0)).
+ apply CR_of_Q_zero. apply CR_of_Q_lt.
+ rewrite <- (Qmult_0_r (q-r)). apply Qmult_lt_l.
+ apply Qlt_minus_iff in H3. exact H3. reflexivity.
+ - apply (CRmorph_increasing _ _ f) in H5.
+ destruct (CRmorph_plus _ _ f x (CR_of_Q R1 ((q-r) * (1#A)))) as [_ H6].
+ apply (CRlt_le_trans R2 _ _ _ H5) in H6. clear H5.
+ destruct (CRmorph_rat _ _ f s) as [H5 _ ].
+ apply (CRle_lt_trans R2 _ _ _ H5) in H6. clear H5.
+ apply (CRmult_lt_compat_r R2 (CRmorph _ _ f y)) in H6.
+ apply (CRlt_le_trans R1 _ (CRmult R1 (CR_of_Q R1 s) y)).
+ apply CRmult_lt_compat_r. exact H. exact H4. clear H4.
+ apply (CRmorph_le_inv _ _ f).
+ apply (CRle_trans R2 _ (CR_of_Q R2 q)).
+ 2: destruct (CRmorph_rat _ _ f q); exact H0.
+ apply (CRle_trans R2 _ (CRmult R2 (CR_of_Q R2 s) (CRmorph R1 R2 f y))).
+ + apply (CRle_trans R2 _ (CRmorph _ _ f (CRmult R1 y (CR_of_Q R1 s)))).
+ destruct (CRmorph_proper _ _ f (CRmult R1 (CR_of_Q R1 s) y)
+ (CRmult R1 y (CR_of_Q R1 s))).
+ apply (Rmul_comm (CRisRing R1)). exact H4.
+ apply (CRle_trans R2 _ (CRmult R2 (CRmorph R1 R2 f y) (CR_of_Q R2 s))).
+ exact (proj2 (CRmorph_mult_rat _ _ f y s)).
+ destruct (Rmul_comm (CRisRing R2) (CR_of_Q R2 s) (CRmorph R1 R2 f y)).
+ exact H0.
+ + intro H5. apply (CRlt_asym R2 _ _ H5). clear H5.
+ apply (CRlt_trans R2 _ _ _ H6). clear H6.
+ apply (CRle_lt_trans
+ R2 _ (CRplus R2
+ (CRmult R2 (CRmorph _ _ f x) (CRmorph _ _ f y))
+ (CRmult R2 (CRmorph R1 R2 f (CR_of_Q R1 ((q - r) * (1 # A))))
+ (CRmorph R1 R2 f y)))).
+ apply (Rdistr_l (CRisRing R2)).
+ apply (CRle_lt_trans
+ R2 _ (CRplus R2 (CR_of_Q R2 r)
+ (CRmult R2 (CRmorph R1 R2 f (CR_of_Q R1 ((q - r) * (1 # A))))
+ (CRmorph R1 R2 f y)))).
+ apply CRplus_le_compat_r. intro H5. apply (CRlt_asym R2 _ _ H5 H2).
+ clear H2.
+ apply (CRle_lt_trans
+ R2 _ (CRplus R2 (CR_of_Q R2 r)
+ (CRmult R2 (CR_of_Q R2 ((q - r) * (1 # A)))
+ (CRmorph R1 R2 f y)))).
+ apply CRplus_le_compat_l, CRmult_le_compat_r.
+ apply (CRle_lt_trans R2 _ (CRmorph _ _ f (CRzero R1))).
+ apply CRmorph_zero. apply CRmorph_increasing. exact H.
+ destruct (CRmorph_rat _ _ f ((q - r) * (1 # A))). exact H2.
+ apply (CRlt_le_trans R2 _ (CRplus R2 (CR_of_Q R2 r)
+ (CR_of_Q R2 ((q - r))))).
+ apply CRplus_lt_compat_l.
+ * apply (CRmult_lt_reg_l R2 (CR_of_Q R2 (/((q - r) * (1 # A))))).
+ apply (CRle_lt_trans R2 _ (CR_of_Q R2 0)). apply CR_of_Q_zero.
+ apply CR_of_Q_lt, Qinv_lt_0_compat.
+ rewrite <- (Qmult_0_r (q-r)). apply Qmult_lt_l.
+ apply Qlt_minus_iff in H3. exact H3. reflexivity.
+ apply (CRle_lt_trans R2 _ (CRmorph _ _ f y)).
+ apply (CRle_trans R2 _ (CRmult R2 (CRmult R2 (CR_of_Q R2 (/ ((q - r) * (1 # A))))
+ (CR_of_Q R2 ((q - r) * (1 # A))))
+ (CRmorph R1 R2 f y))).
+ exact (proj2 (Rmul_assoc (CRisRing R2) (CR_of_Q R2 (/ ((q - r) * (1 # A))))
+ (CR_of_Q R2 ((q - r) * (1 # A)))
+ (CRmorph _ _ f y))).
+ apply (CRle_trans R2 _ (CRmult R2 (CRone R2) (CRmorph R1 R2 f y))).
+ apply CRmult_le_compat_r.
+ apply (CRle_lt_trans R2 _ (CRmorph _ _ f (CRzero R1))).
+ apply CRmorph_zero. apply CRmorph_increasing. exact H.
+ apply (CRle_trans R2 _ (CR_of_Q R2 ((/ ((q - r) * (1 # A))) * ((q - r) * (1 # A))))).
+ exact (proj1 (CR_of_Q_mult R2 (/ ((q - r) * (1 # A))) ((q - r) * (1 # A)))).
+ apply (CRle_trans R2 _ (CR_of_Q R2 1)).
+ destruct (CR_of_Q_proper R2 (/ ((q - r) * (1 # A)) * ((q - r) * (1 # A))) 1).
+ field_simplify. reflexivity. split.
+ intro H5. inversion H5. intro H5. apply Qlt_minus_iff in H3.
+ rewrite H5 in H3. inversion H3. exact H2.
+ destruct (CR_of_Q_one R2). exact H2.
+ destruct (Rmul_1_l (CRisRing R2) (CRmorph _ _ f y)).
+ intro H5. contradiction.
+ apply (CRlt_le_trans R2 _ (CR_of_Q R2 (Z.pos A # 1))).
+ apply (CRlt_le_trans R2 _ (CRmorph _ _ f (CR_of_Q R1 (Z.pos A # 1)))).
+ apply CRmorph_increasing. exact Amaj.
+ exact (proj2 (CRmorph_rat _ _ f (Z.pos A # 1))).
+ apply (CRle_trans R2 _ (CR_of_Q R2 ((/ ((q - r) * (1 # A))) * (q - r)))).
+ 2: exact (proj2 (CR_of_Q_mult R2 (/ ((q - r) * (1 # A))) (q - r))).
+ destruct (CR_of_Q_proper R2 (Z.pos A # 1) (/ ((q - r) * (1 # A)) * (q - r))).
+ field_simplify. reflexivity. split.
+ intro H5. inversion H5. intro H5. apply Qlt_minus_iff in H3.
+ rewrite H5 in H3. inversion H3. exact H2.
+ * apply (CRle_trans R2 _ (CR_of_Q R2 (r + (q-r)))).
+ exact (proj1 (CR_of_Q_plus R2 r (q-r))).
+ destruct (CR_of_Q_proper R2 (r + (q-r)) q). ring. exact H2.
+ + apply (CRle_lt_trans R2 _ (CRmorph _ _ f (CRzero R1))).
+ apply CRmorph_zero. apply CRmorph_increasing. exact H.
+Qed.
+
+Lemma CRmorph_mult : forall (R1 R2 : ConstructiveReals)
+ (f : ConstructiveRealsMorphism R1 R2)
+ (x y : CRcarrier R1),
+ orderEq _ (CRlt R2) (CRmorph _ _ f (CRmult R1 x y))
+ (CRmult R2 (CRmorph _ _ f x) (CRmorph _ _ f y)).
+Proof.
+ intros.
+ destruct (CR_archimedean R1 (CRopp R1 y)) as [p pmaj].
+ apply (CRplus_eq_reg_r R2 (CRmult R2 (CRmorph _ _ f x)
+ (CR_of_Q R2 (Z.pos p # 1)))).
+ apply (CReq_trans R2 _ (CRmorph _ _ f (CRmult R1 x (CRplus R1 y (CR_of_Q R1 (Z.pos p # 1)))))).
+ - apply (CReq_trans R2 _ (CRplus R2 (CRmorph R1 R2 f (CRmult R1 x y))
+ (CRmorph R1 R2 f (CRmult R1 x (CR_of_Q R1 (Z.pos p # 1)))))).
+ apply (Radd_ext (CRisRingExt R2)). apply CReq_refl.
+ apply CReq_sym, CRmorph_mult_int.
+ apply (CReq_trans R2 _ (CRmorph _ _ f (CRplus R1 (CRmult R1 x y)
+ (CRmult R1 x (CR_of_Q R1 (Z.pos p # 1)))))).
+ apply CReq_sym, CRmorph_plus. apply CRmorph_proper.
+ apply CReq_sym, CRmult_plus_distr_l.
+ - apply (CReq_trans R2 _ (CRmult R2 (CRmorph _ _ f x)
+ (CRmorph _ _ f (CRplus R1 y (CR_of_Q R1 (Z.pos p # 1)))))).
+ apply CRmorph_mult_pos_pos.
+ apply (CRplus_lt_compat_l R1 y) in pmaj.
+ apply (CRle_lt_trans R1 _ (CRplus R1 y (CRopp R1 y))).
+ 2: exact pmaj. apply (CRisRing R1).
+ apply (CReq_trans R2 _ (CRmult R2 (CRmorph R1 R2 f x)
+ (CRplus R2 (CRmorph R1 R2 f y) (CR_of_Q R2 (Z.pos p # 1))))).
+ apply (Rmul_ext (CRisRingExt R2)). apply CReq_refl.
+ apply (CReq_trans R2 _ (CRplus R2 (CRmorph R1 R2 f y)
+ (CRmorph _ _ f (CR_of_Q R1 (Z.pos p # 1))))).
+ apply CRmorph_plus.
+ apply (Radd_ext (CRisRingExt R2)). apply CReq_refl.
+ apply CRmorph_rat.
+ apply CRmult_plus_distr_l.
+Qed.
+
+Lemma CRmorph_appart : forall (R1 R2 : ConstructiveReals)
+ (f : ConstructiveRealsMorphism R1 R2)
+ (x y : CRcarrier R1)
+ (app : orderAppart _ (CRlt R1) x y),
+ orderAppart _ (CRlt R2) (CRmorph _ _ f x) (CRmorph _ _ f y).
+Proof.
+ intros. destruct app.
+ - left. apply CRmorph_increasing. exact c.
+ - right. apply CRmorph_increasing. exact c.
+Defined.
+
+Lemma CRmorph_appart_zero : forall (R1 R2 : ConstructiveReals)
+ (f : ConstructiveRealsMorphism R1 R2)
+ (x : CRcarrier R1)
+ (app : orderAppart _ (CRlt R1) x (CRzero R1)),
+ orderAppart _ (CRlt R2) (CRmorph _ _ f x) (CRzero R2).
+Proof.
+ intros. destruct app.
+ - left. apply (CRlt_le_trans R2 _ (CRmorph _ _ f (CRzero R1))).
+ apply CRmorph_increasing. exact c.
+ exact (proj2 (CRmorph_zero _ _ f)).
+ - right. apply (CRle_lt_trans R2 _ (CRmorph _ _ f (CRzero R1))).
+ exact (proj1 (CRmorph_zero _ _ f)).
+ apply CRmorph_increasing. exact c.
+Defined.
+
+Lemma CRmorph_inv : forall (R1 R2 : ConstructiveReals)
+ (f : ConstructiveRealsMorphism R1 R2)
+ (x : CRcarrier R1)
+ (xnz : orderAppart _ (CRlt R1) x (CRzero R1))
+ (fxnz : orderAppart _ (CRlt R2) (CRmorph _ _ f x) (CRzero R2)),
+ orderEq _ (CRlt R2) (CRmorph _ _ f (CRinv R1 x xnz))
+ (CRinv R2 (CRmorph _ _ f x) fxnz).
+Proof.
+ intros. apply (CRmult_eq_reg_r R2 (CRmorph _ _ f x)).
+ destruct fxnz. right. exact c. left. exact c.
+ apply (CReq_trans R2 _ (CRone R2)).
+ 2: apply CReq_sym, CRinv_l.
+ apply (CReq_trans R2 _ (CRmorph _ _ f (CRmult R1 (CRinv R1 x xnz) x))).
+ apply CReq_sym, CRmorph_mult.
+ apply (CReq_trans R2 _ (CRmorph _ _ f (CRone R1))).
+ apply CRmorph_proper. apply CRinv_l.
+ apply CRmorph_one.
+Qed.
+
+Definition CauchyMorph (R : ConstructiveReals)
+ : CReal -> CRcarrier R.
+Proof.
+ intros [xn xcau].
+ destruct (CR_complete R (fun n:nat => CR_of_Q R (xn n))).
+ - intros p. exists (Pos.to_nat p). intros.
+ specialize (xcau p i j H H0). apply Qlt_le_weak in xcau.
+ rewrite Qabs_Qle_condition in xcau. split.
+ + unfold CRminus.
+ apply (CRle_trans R _ (CRplus R (CR_of_Q R (xn i)) (CR_of_Q R (-xn j)))).
+ apply (CRle_trans R _ (CR_of_Q R (xn i-xn j))).
+ apply CR_of_Q_le. apply xcau. exact (proj2 (CR_of_Q_plus R _ _)).
+ apply CRplus_le_compat_l. exact (proj2 (CR_of_Q_opp R (xn j))).
+ + unfold CRminus.
+ apply (CRle_trans R _ (CRplus R (CR_of_Q R (xn i)) (CR_of_Q R (-xn j)))).
+ apply CRplus_le_compat_l. exact (proj1 (CR_of_Q_opp R (xn j))).
+ apply (CRle_trans R _ (CR_of_Q R (xn i-xn j))).
+ exact (proj1 (CR_of_Q_plus R _ _)).
+ apply CR_of_Q_le. apply xcau.
+ - exact x.
+Defined.
+
+Lemma CauchyMorph_rat : forall (R : ConstructiveReals) (q : Q),
+ orderEq _ (CRlt R) (CauchyMorph R (inject_Q q)) (CR_of_Q R q).
+Proof.
+ intros.
+ unfold CauchyMorph; simpl;
+ destruct (CRltLinear R), p, (CR_complete R (fun _ : nat => CR_of_Q R q)).
+ apply CR_cv_const in c0. apply CReq_sym. exact c0.
+Qed.
+
+Lemma CauchyMorph_increasing_Ql : forall (R : ConstructiveReals) (x : CReal) (q : Q),
+ CRealLt x (inject_Q q) -> CRlt R (CauchyMorph R x) (CR_of_Q R q).
+Proof.
+ intros.
+ unfold CauchyMorph; simpl;
+ destruct x as [xn xcau], (CRltLinear R), p, (CR_complete R (fun n : nat => CR_of_Q R (xn n))).
+ destruct (CRealQ_dense _ _ H) as [r [H0 H1]].
+ apply lt_inject_Q in H1.
+ destruct (s _ x _ (CR_of_Q_lt R _ _ H1)). 2: exact c1. exfalso.
+ clear H1 H q.
+ (* For an index high enough, xn should be both higher
+ and lower than r, which is absurd. *)
+ apply CRealLt_above in H0.
+ destruct H0 as [p pmaj]. simpl in pmaj.
+ destruct (CR_cv_above_rat R xn x r c0 c1).
+ assert (x0 <= Nat.max (Pos.to_nat p) (S x0))%nat.
+ { apply (le_trans _ (S x0)). apply le_S, le_refl. apply Nat.le_max_r. }
+ specialize (q (Nat.max (Pos.to_nat p) (S x0)) H). clear H.
+ specialize (pmaj (Pos.max p (Pos.of_nat (S x0))) (Pos.le_max_l _ _)).
+ rewrite Pos2Nat.inj_max, Nat2Pos.id in pmaj. 2: discriminate.
+ apply (Qlt_not_le _ _ q). apply Qlt_le_weak.
+ apply Qlt_minus_iff. apply (Qlt_trans _ (2#p)). reflexivity. exact pmaj.
+Qed.
+
+Lemma CauchyMorph_increasing_Qr : forall (R : ConstructiveReals) (x : CReal) (q : Q),
+ CRealLt (inject_Q q) x -> CRlt R (CR_of_Q R q) (CauchyMorph R x).
+Proof.
+ intros.
+ unfold CauchyMorph; simpl;
+ destruct x as [xn xcau], (CRltLinear R), p, (CR_complete R (fun n : nat => CR_of_Q R (xn n))).
+ destruct (CRealQ_dense _ _ H) as [r [H0 H1]].
+ apply lt_inject_Q in H0.
+ destruct (s _ x _ (CR_of_Q_lt R _ _ H0)). exact c1. exfalso.
+ clear H0 H q.
+ (* For an index high enough, xn should be both higher
+ and lower than r, which is absurd. *)
+ apply CRealLt_above in H1.
+ destruct H1 as [p pmaj]. simpl in pmaj.
+ destruct (CR_cv_below_rat R xn x r c0 c1).
+ assert (x0 <= Nat.max (Pos.to_nat p) (S x0))%nat.
+ { apply (le_trans _ (S x0)). apply le_S, le_refl. apply Nat.le_max_r. }
+ specialize (q (Nat.max (Pos.to_nat p) (S x0)) H). clear H.
+ specialize (pmaj (Pos.max p (Pos.of_nat (S x0))) (Pos.le_max_l _ _)).
+ rewrite Pos2Nat.inj_max, Nat2Pos.id in pmaj. 2: discriminate.
+ apply (Qlt_not_le _ _ q). apply Qlt_le_weak.
+ apply Qlt_minus_iff. apply (Qlt_trans _ (2#p)). reflexivity. exact pmaj.
+Qed.
+
+Lemma CauchyMorph_increasing : forall (R : ConstructiveReals) (x y : CReal),
+ CRealLt x y -> CRlt R (CauchyMorph R x) (CauchyMorph R y).
+Proof.
+ intros.
+ destruct (CRealQ_dense _ _ H) as [q [H0 H1]].
+ apply (CRlt_trans R _ (CR_of_Q R q)).
+ apply CauchyMorph_increasing_Ql. exact H0.
+ apply CauchyMorph_increasing_Qr. exact H1.
+Qed.
+
+Definition CauchyMorphism (R : ConstructiveReals) : ConstructiveRealsMorphism CRealImplem R.
+Proof.
+ apply (Build_ConstructiveRealsMorphism CRealImplem R (CauchyMorph R)).
+ exact (CauchyMorph_rat R).
+ exact (CauchyMorph_increasing R).
+Defined.
+
+Lemma RightBound : forall (R : ConstructiveReals) (x : CRcarrier R) (p q r : Q),
+ CRlt R x (CR_of_Q R q)
+ -> CRlt R x (CR_of_Q R r)
+ -> CRlt R (CR_of_Q R q) (CRplus R x (CR_of_Q R p))
+ -> CRlt R (CR_of_Q R r) (CRplus R x (CR_of_Q R p))
+ -> Qlt (Qabs (q - r)) p.
+Proof.
+ intros. apply Qabs_case.
+ - intros. apply (Qplus_lt_l _ _ r). ring_simplify.
+ apply (lt_CR_of_Q R), (CRlt_le_trans R _ _ _ H1).
+ apply (CRle_trans R _ (CRplus R (CR_of_Q R r) (CR_of_Q R p))).
+ intro abs. apply CRplus_lt_reg_r in abs.
+ exact (CRlt_asym R _ _ abs H0).
+ destruct (CR_of_Q_plus R r p). exact H4.
+ - intros. apply (Qplus_lt_l _ _ q). ring_simplify.
+ apply (lt_CR_of_Q R), (CRlt_le_trans R _ _ _ H2).
+ apply (CRle_trans R _ (CRplus R (CR_of_Q R q) (CR_of_Q R p))).
+ intro abs. apply CRplus_lt_reg_r in abs.
+ exact (CRlt_asym R _ _ abs H).
+ destruct (CR_of_Q_plus R q p). exact H4.
+Qed.
+
+Definition CauchyMorph_inv (R : ConstructiveReals)
+ : CRcarrier R -> CReal.
+Proof.
+ intro x.
+ exists (fun n:nat => let (q,_) := CR_Q_dense
+ R x _ (CRplus_pos_rat_lt R x (1 # Pos.of_nat (S n)) (eq_refl _))
+ in q).
+ intros n p q H0 H1.
+ destruct (CR_Q_dense R x (CRplus R x (CR_of_Q R (1 # Pos.of_nat (S p))))
+ (CRplus_pos_rat_lt R x (1 # Pos.of_nat (S p)) (eq_refl _)))
+ as [r [H2 H3]].
+ destruct (CR_Q_dense R x (CRplus R x (CR_of_Q R (1 # Pos.of_nat (S q))))
+ (CRplus_pos_rat_lt R x (1 # Pos.of_nat (S q)) (eq_refl _)))
+ as [s [H4 H5]].
+ apply (RightBound R x (1#n) r s). exact H2. exact H4.
+ apply (CRlt_trans R _ _ _ H3), CRplus_lt_compat_l, CR_of_Q_lt.
+ unfold Qlt. do 2 rewrite Z.mul_1_l. unfold Qden.
+ apply Pos2Z.pos_lt_pos, Pos2Nat.inj_lt. rewrite Nat2Pos.id.
+ 2: discriminate. apply le_n_S. exact H0.
+ apply (CRlt_trans R _ _ _ H5), CRplus_lt_compat_l, CR_of_Q_lt.
+ unfold Qlt. do 2 rewrite Z.mul_1_l. unfold Qden.
+ apply Pos2Z.pos_lt_pos, Pos2Nat.inj_lt. rewrite Nat2Pos.id.
+ 2: discriminate. apply le_n_S. exact H1.
+Defined.
+
+Lemma CauchyMorph_inv_rat : forall (R : ConstructiveReals) (q : Q),
+ CRealEq (CauchyMorph_inv R (CR_of_Q R q)) (inject_Q q).
+Proof.
+ split.
+ - intros [n nmaj]. unfold CauchyMorph_inv, proj1_sig, inject_Q in nmaj.
+ destruct (CR_Q_dense R (CR_of_Q R q)
+ (CRplus R (CR_of_Q R q) (CR_of_Q R (1 # Pos.of_nat (S (Pos.to_nat n)))))
+ (CRplus_pos_rat_lt R (CR_of_Q R q) (1 # Pos.of_nat (S (Pos.to_nat n)))
+ eq_refl))
+ as [r [H _]].
+ apply lt_CR_of_Q, Qlt_minus_iff in H.
+ apply (Qlt_not_le _ _ H), (Qplus_le_l _ _ (q-r)).
+ ring_simplify. apply (Qle_trans _ (2#n)). discriminate.
+ apply Qlt_le_weak. ring_simplify in nmaj. rewrite Qplus_comm. exact nmaj.
+ - intros [n nmaj]. unfold CauchyMorph_inv, proj1_sig, inject_Q in nmaj.
+ destruct (CR_Q_dense R (CR_of_Q R q)
+ (CRplus R (CR_of_Q R q) (CR_of_Q R (1 # Pos.of_nat (S (Pos.to_nat n)))))
+ (CRplus_pos_rat_lt R (CR_of_Q R q) (1 # Pos.of_nat (S (Pos.to_nat n)))
+ eq_refl))
+ as [r [_ H0]].
+ destruct (CR_of_Q_plus R q (1 # Pos.of_nat (S (Pos.to_nat n)))) as [H1 _].
+ apply (CRlt_le_trans R _ _ _ H0) in H1. clear H0.
+ apply lt_CR_of_Q, (Qplus_lt_l _ _ (-q)) in H1.
+ ring_simplify in H1. ring_simplify in nmaj.
+ apply (Qlt_trans _ _ _ nmaj) in H1. clear nmaj.
+ apply (Qlt_not_le _ _ H1). clear H1.
+ apply (Qle_trans _ (1#n)).
+ unfold Qle, Qnum, Qden. do 2 rewrite Z.mul_1_l.
+ apply Pos2Z.pos_le_pos. apply Pos2Nat.inj_le.
+ rewrite Nat2Pos.id. 2: discriminate. apply le_S, le_refl.
+ unfold Qle, Qnum, Qden. apply Z.mul_le_mono_nonneg_r.
+ 2: discriminate. apply Pos2Z.pos_is_nonneg.
+Qed.
+
+(* The easier side, because CauchyMorph_inv takes a limit from above. *)
+Lemma CauchyMorph_inv_increasing_Qr
+ : forall (R : ConstructiveReals) (x : CRcarrier R) (q : Q),
+ CRlt R (CR_of_Q R q) x -> CRealLt (inject_Q q) (CauchyMorph_inv R x).
+Proof.
+ intros.
+ destruct (CR_Q_dense R _ _ H) as [r [H2 H3]].
+ apply lt_CR_of_Q in H2.
+ destruct (Qarchimedean (/(r-q))) as [p pmaj].
+ exists (2*p)%positive. unfold CauchyMorph_inv, inject_Q, proj1_sig.
+ destruct (CR_Q_dense
+ R x (CRplus R x (CR_of_Q R (1 # Pos.of_nat (S (Pos.to_nat (2*p))))))
+ (CRplus_pos_rat_lt R x (1 # Pos.of_nat (S (Pos.to_nat (2*p)))) eq_refl))
+ as [t [H4 H5]].
+ setoid_replace (2#2*p) with (1#p). 2: reflexivity.
+ apply (Qlt_trans _ (r-q)).
+ apply (Qmult_lt_l _ _ (r-q)) in pmaj.
+ rewrite Qmult_inv_r in pmaj.
+ apply Qlt_shift_inv_r in pmaj. 2: reflexivity. exact pmaj.
+ intro abs. apply Qlt_minus_iff in H2.
+ rewrite abs in H2. inversion H2.
+ apply Qlt_minus_iff in H2. exact H2.
+ apply Qplus_lt_l, (lt_CR_of_Q R), (CRlt_trans R _ x _ H3 H4).
+Qed.
+
+Lemma CauchyMorph_inv_increasing : forall (R : ConstructiveReals) (x y : CRcarrier R),
+ CRlt R x y -> CRealLt (CauchyMorph_inv R x) (CauchyMorph_inv R y).
+Proof.
+ intros.
+ destruct (CR_Q_dense R _ _ H) as [q [H0 H1]].
+ apply (CReal_lt_trans _ (inject_Q q)).
+ - clear H1 H y.
+ destruct (CR_Q_dense R _ _ H0) as [r [H2 H3]].
+ apply lt_CR_of_Q in H3.
+ destruct (Qarchimedean (/(q-r))) as [p pmaj].
+ exists (4*p)%positive. unfold CauchyMorph_inv, inject_Q, proj1_sig.
+ destruct (CR_Q_dense
+ R x (CRplus R x (CR_of_Q R (1 # Pos.of_nat (S (Pos.to_nat (4*p))))))
+ (CRplus_pos_rat_lt R x (1 # Pos.of_nat (S (Pos.to_nat (4*p)))) eq_refl))
+ as [t [H4 H5]].
+ setoid_replace (2#4*p) with (1#2*p). 2: reflexivity.
+ assert (1 # 2 * p < (q - r) / 2) as H.
+ { apply Qlt_shift_div_l. reflexivity.
+ setoid_replace ((1#2*p)*2) with (1#p).
+ apply (Qmult_lt_l _ _ (q-r)) in pmaj.
+ rewrite Qmult_inv_r in pmaj.
+ apply Qlt_shift_inv_r in pmaj. 2: reflexivity. exact pmaj.
+ intro abs. apply Qlt_minus_iff in H3.
+ rewrite abs in H3. inversion H3.
+ apply Qlt_minus_iff in H3. exact H3.
+ rewrite Qmult_comm. reflexivity. }
+ apply (Qlt_trans _ ((q-r)/2)). exact H.
+ apply (Qplus_lt_l _ _ (t + (r-q)/2)). field_simplify.
+ setoid_replace (2*t/2) with t. 2: field.
+ apply (lt_CR_of_Q R). apply (CRlt_trans R _ _ _ H5).
+ apply (CRlt_trans
+ R _ (CRplus R (CR_of_Q R r) (CR_of_Q R (1 # Pos.of_nat (S (Pos.to_nat (4 * p))))))).
+ apply CRplus_lt_compat_r. exact H2.
+ apply (CRle_lt_trans
+ R _ (CR_of_Q R (r + (1 # Pos.of_nat (S (Pos.to_nat (4 * p))))))).
+ apply CR_of_Q_plus. apply CR_of_Q_lt.
+ apply (Qlt_le_trans _ (r + (q-r)/2)).
+ 2: field_simplify; apply Qle_refl.
+ apply Qplus_lt_r.
+ apply (Qlt_trans _ (1#2*p)). 2: exact H.
+ unfold Qlt. do 2 rewrite Z.mul_1_l. unfold Qden.
+ apply Pos2Z.pos_lt_pos.
+ rewrite Nat2Pos.inj_succ, Pos2Nat.id.
+ apply (Pos.lt_trans _ (4*p)). apply Pos2Nat.inj_lt.
+ do 2 rewrite Pos2Nat.inj_mul.
+ apply Nat.mul_lt_mono_pos_r. apply Pos2Nat.is_pos.
+ unfold Pos.to_nat. simpl. auto.
+ apply Pos.lt_succ_diag_r.
+ intro abs. pose proof (Pos2Nat.is_pos (4*p)).
+ rewrite abs in H1. inversion H1.
+ - apply CauchyMorph_inv_increasing_Qr. exact H1.
+Qed.
+
+Definition CauchyMorphismInv (R : ConstructiveReals)
+ : ConstructiveRealsMorphism R CRealImplem.
+Proof.
+ apply (Build_ConstructiveRealsMorphism R CRealImplem (CauchyMorph_inv R)).
+ - apply CauchyMorph_inv_rat.
+ - apply CauchyMorph_inv_increasing.
+Defined.
+
+Lemma CauchyMorph_surject : forall (R : ConstructiveReals) (x : CRcarrier R),
+ orderEq _ (CRlt R) (CauchyMorph R (CauchyMorph_inv R x)) x.
+Proof.
+ intros.
+ apply (Endomorph_id
+ R (CRmorph_compose _ _ _ (CauchyMorphismInv R) (CauchyMorphism R)) x).
+Qed.
+
+Lemma CauchyMorph_inject : forall (R : ConstructiveReals) (x : CReal),
+ CRealEq (CauchyMorph_inv R (CauchyMorph R x)) x.
+Proof.
+ intros.
+ apply (Endomorph_id CRealImplem (CRmorph_compose _ _ _ (CauchyMorphism R) (CauchyMorphismInv R)) x).
+Qed.
+
+(* We call this morphism slow to remind that it should only be used
+ for proofs, not for computations. *)
+Definition SlowConstructiveRealsMorphism (R1 R2 : ConstructiveReals)
+ : ConstructiveRealsMorphism R1 R2
+ := CRmorph_compose R1 CRealImplem R2
+ (CauchyMorphismInv R1) (CauchyMorphism R2).
diff --git a/theories/Reals/Raxioms.v b/theories/Reals/Raxioms.v
index f03b0ccea3..d856d1c7fe 100644
--- a/theories/Reals/Raxioms.v
+++ b/theories/Reals/Raxioms.v
@@ -21,6 +21,7 @@
Require Export ZArith_base.
Require Import ConstructiveRIneq.
+Require Import ConstructiveRealsLUB.
Require Export Rdefinitions.
Declare Scope R_scope.
Local Open Scope R_scope.
@@ -408,6 +409,10 @@ Lemma completeness :
bound E -> (exists x : R, E x) -> { m:R | is_lub E m }.
Proof.
intros. pose (fun x:ConstructiveRIneq.R => E (Rabst x)) as Er.
+ assert (forall x y : CRcarrier CR, orderEq (CRcarrier CR) (CRlt CR) x y -> Er x <-> Er y)
+ as Erproper.
+ { intros. unfold Er. replace (Rabst x) with (Rabst y). reflexivity.
+ apply Rquot1. do 2 rewrite Rquot2. split; apply H1. }
assert (exists x : ConstructiveRIneq.R, Er x) as Einhab.
{ destruct H0. exists (Rrepr x). unfold Er.
replace (Rabst (Rrepr x)) with x. exact H0.
@@ -418,7 +423,7 @@ Proof.
{ destruct H. exists (Rrepr x). intros y Ey. rewrite <- (Rquot2 y).
apply Rrepr_le. apply H. exact Ey. }
destruct (CR_sig_lub CR
- Er sig_forall_dec sig_not_dec Einhab Ebound).
+ Er Erproper sig_forall_dec sig_not_dec Einhab Ebound).
exists (Rabst x). split.
intros y Ey. apply Rrepr_le. rewrite Rquot2.
unfold ConstructiveRIneq.Rle. apply a.