Library Coq.Reals.Abstract.ConstructiveRealsMorphisms


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.
WARNING: this file is experimental and likely to change in future releases.

Require Import QArith.
Require Import Qabs.
Require Import ConstructiveReals.
Require Import ConstructiveLimits.
Require Import ConstructiveAbs.

Local Open Scope ConstructiveReals.

Record ConstructiveRealsMorphism {R1 R2 : ConstructiveReals} : Set :=
  {
    CRmorph : CRcarrier R1 -> CRcarrier R2;
    CRmorph_rat : forall q : Q,
        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)
      (x y : CRcarrier R1),
    CRlt R2 (CRmorph f x) (CRmorph f y)
    -> CRlt R1 x y.

Lemma CRmorph_unique : forall {R1 R2 : ConstructiveReals}
                         (f g : @ConstructiveRealsMorphism R1 R2)
                         (x : CRcarrier R1),
    CRmorph f x == CRmorph g x.

Lemma Endomorph_id
  : forall {R : ConstructiveReals} (f : @ConstructiveRealsMorphism R R)
      (x : CRcarrier R),
    CRmorph f x == x.

Lemma CRmorph_proper
  : forall {R1 R2 : ConstructiveReals}
      (f : @ConstructiveRealsMorphism R1 R2)
      (x y : CRcarrier R1),
    x == y -> CRmorph f x == CRmorph f y.

Definition CRmorph_compose {R1 R2 R3 : ConstructiveReals}
           (f : @ConstructiveRealsMorphism R1 R2)
           (g : @ConstructiveRealsMorphism R2 R3)
  : @ConstructiveRealsMorphism R1 R3.

Lemma CRmorph_le : forall {R1 R2 : ConstructiveReals}
                     (f : @ConstructiveRealsMorphism R1 R2)
                     (x y : CRcarrier R1),
    x <= y -> CRmorph f x <= CRmorph f y.

Lemma CRmorph_le_inv : forall {R1 R2 : ConstructiveReals}
                         (f : @ConstructiveRealsMorphism R1 R2)
                         (x y : CRcarrier R1),
    CRmorph f x <= CRmorph f y -> x <= y.

Lemma CRmorph_zero : forall {R1 R2 : ConstructiveReals}
                       (f : @ConstructiveRealsMorphism R1 R2),
    CRmorph f 0 == 0.

Lemma CRmorph_one : forall {R1 R2 : ConstructiveReals}
                      (f : @ConstructiveRealsMorphism R1 R2),
    CRmorph f 1 == 1.

Lemma CRmorph_opp : forall {R1 R2 : ConstructiveReals}
                      (f : @ConstructiveRealsMorphism R1 R2)
                      (x : CRcarrier R1),
    CRmorph f (- x) == - CRmorph f x.

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)).

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.

Lemma CRmorph_plus_rat : forall {R1 R2 : ConstructiveReals}
                                (f : @ConstructiveRealsMorphism R1 R2)
                                (x : CRcarrier R1) (q : Q),
    CRmorph f (CRplus R1 x (CR_of_Q R1 q))
    == CRplus R2 (CRmorph f x) (CR_of_Q R2 q).

Lemma CRmorph_plus : forall {R1 R2 : ConstructiveReals}
                       (f : @ConstructiveRealsMorphism R1 R2)
                       (x y : CRcarrier R1),
    CRmorph f (CRplus R1 x y)
    == CRplus R2 (CRmorph f x) (CRmorph f y).

Lemma CRmorph_mult_pos : forall {R1 R2 : ConstructiveReals}
                           (f : @ConstructiveRealsMorphism R1 R2)
                           (x : CRcarrier R1) (n : nat),
    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)).

Lemma NatOfZ : forall n : Z, { p : nat | n = Z.of_nat p \/ n = Z.opp (Z.of_nat p) }.

Lemma CRmorph_mult_int : forall {R1 R2 : ConstructiveReals}
                           (f : @ConstructiveRealsMorphism R1 R2)
                           (x : CRcarrier R1) (n : Z),
    CRmorph f (CRmult R1 x (CR_of_Q R1 (n # 1)))
    == CRmult R2 (CRmorph f x) (CR_of_Q R2 (n # 1)).

Lemma CRmorph_mult_inv : forall {R1 R2 : ConstructiveReals}
                           (f : @ConstructiveRealsMorphism R1 R2)
                           (x : CRcarrier R1) (p : positive),
    CRmorph f (CRmult R1 x (CR_of_Q R1 (1 # p)))
    == CRmult R2 (CRmorph f x) (CR_of_Q R2 (1 # p)).

Lemma CRmorph_mult_rat : forall {R1 R2 : ConstructiveReals}
                           (f : @ConstructiveRealsMorphism R1 R2)
                           (x : CRcarrier R1) (q : Q),
    CRmorph f (CRmult R1 x (CR_of_Q R1 q))
    == CRmult R2 (CRmorph f x) (CR_of_Q R2 q).

Lemma CRmorph_mult_pos_pos_le : forall {R1 R2 : ConstructiveReals}
                                  (f : @ConstructiveRealsMorphism R1 R2)
                                  (x y : CRcarrier R1),
    CRlt R1 0 y
    -> CRmult R2 (CRmorph f x) (CRmorph f y)
       <= CRmorph f (CRmult R1 x y).

Lemma CRmorph_mult_pos_pos : forall {R1 R2 : ConstructiveReals}
                               (f : @ConstructiveRealsMorphism R1 R2)
                               (x y : CRcarrier R1),
    CRlt R1 0 y
    -> CRmorph f (CRmult R1 x y)
       == CRmult R2 (CRmorph f x) (CRmorph f y).

Lemma CRmorph_mult : forall {R1 R2 : ConstructiveReals}
                       (f : @ConstructiveRealsMorphism R1 R2)
                       (x y : CRcarrier R1),
    CRmorph f (CRmult R1 x y)
    == CRmult R2 (CRmorph f x) (CRmorph f y).

Lemma CRmorph_appart : forall {R1 R2 : ConstructiveReals}
                         (f : @ConstructiveRealsMorphism R1 R2)
                         (x y : CRcarrier R1)
                         (app : x y),
    CRmorph f x CRmorph f y.

Lemma CRmorph_appart_zero : forall {R1 R2 : ConstructiveReals}
                              (f : @ConstructiveRealsMorphism R1 R2)
                              (x : CRcarrier R1)
                              (app : x 0),
    CRmorph f x 0.

Lemma CRmorph_inv : forall {R1 R2 : ConstructiveReals}
                      (f : @ConstructiveRealsMorphism R1 R2)
                      (x : CRcarrier R1)
                      (xnz : x 0)
                      (fxnz : CRmorph f x 0),
    CRmorph f ((/ x) xnz)
    == (/ CRmorph f x) fxnz.

Lemma CRmorph_rat_cv
  : forall {R1 R2 : ConstructiveReals}
           (qn : nat -> Q),
  CR_cauchy R1 (fun n => CR_of_Q R1 (qn n))
  -> CR_cauchy R2 (fun n => CR_of_Q R2 (qn n)).

Definition CR_Q_limit {R : ConstructiveReals} (x : CRcarrier R) (n:nat)
  : { q:Q & x < CR_of_Q R q < x + CR_of_Q R (1 # Pos.of_nat n) }.

Lemma CR_Q_limit_cv : forall {R : ConstructiveReals} (x : CRcarrier R),
    CR_cv R (fun n => CR_of_Q R (let (q,_) := CR_Q_limit x n in q)) x.

Definition SlowMorph {R1 R2 : ConstructiveReals}
  : CRcarrier R1 -> CRcarrier R2
  := fun x => let (y,_) := CR_complete R2 _ (CRmorph_rat_cv _ (Rcv_cauchy_mod _ x (CR_Q_limit_cv x)))
              in y.

Lemma CauchyMorph_rat : forall {R1 R2 : ConstructiveReals} (q : Q),
    SlowMorph (CR_of_Q R1 q) == CR_of_Q R2 q.

Lemma SlowMorph_increasing_Qr
  : forall {R1 R2 : ConstructiveReals} (x : CRcarrier R1) (q : Q),
    CR_of_Q R1 q < x -> CR_of_Q R2 q < SlowMorph x.

Lemma SlowMorph_increasing_Ql
  : forall {R1 R2 : ConstructiveReals} (x : CRcarrier R1) (q : Q),
    x < CR_of_Q R1 q -> SlowMorph x < CR_of_Q R2 q.

Lemma SlowMorph_increasing : forall {R1 R2 : ConstructiveReals} (x y : CRcarrier R1),
    x < y -> @SlowMorph R1 R2 x < SlowMorph y.

Definition SlowConstructiveRealsMorphism {R1 R2 : ConstructiveReals}
  : @ConstructiveRealsMorphism R1 R2
  := Build_ConstructiveRealsMorphism
       R1 R2 SlowMorph CauchyMorph_rat
       SlowMorph_increasing.

Lemma CRmorph_abs : forall {R1 R2 : ConstructiveReals}
                      (f : @ConstructiveRealsMorphism R1 R2)
                      (x : CRcarrier R1),
    CRabs R2 (CRmorph f x) == CRmorph f (CRabs R1 x).

Lemma CRmorph_cv : forall {R1 R2 : ConstructiveReals}
                     (f : @ConstructiveRealsMorphism R1 R2)
                     (un : nat -> CRcarrier R1)
                     (l : CRcarrier R1),
    CR_cv R1 un l
    -> CR_cv R2 (fun n => CRmorph f (un n)) (CRmorph f l).

Lemma CRmorph_cauchy_reverse : forall {R1 R2 : ConstructiveReals}
                     (f : @ConstructiveRealsMorphism R1 R2)
                     (un : nat -> CRcarrier R1),
    CR_cauchy R2 (fun n => CRmorph f (un n))
    -> CR_cauchy R1 un.