Skip to content

Commit 7f2aab9

Browse files
authored
Merge pull request #43 from c-corn/v8.6
V8.6
2 parents f895f2e + 8a64c76 commit 7f2aab9

File tree

6 files changed

+100
-85
lines changed

6 files changed

+100
-85
lines changed

coq_reals/Rreals.v

Lines changed: 8 additions & 4 deletions
Original file line numberDiff line numberDiff line change
@@ -452,23 +452,27 @@ Proof.
452452
exists (Zabs_nat (up x)).
453453
unfold Zabs_nat.
454454
elim (archimed x).
455-
destruct (up x); simpl.
456-
intros; fourier.
457-
unfold nat_of_P.
455+
destruct (up x).
456+
simpl; intros; fourier.
457+
rewrite <- positive_nat_Z.
458458
intros H _.
459459
apply Rlt_le.
460460
rewrite <- R_INR_as_IR.
461-
auto.
461+
now rewrite INR_IZR_INZ.
462462
intros I _.
463463
cut (x < 0%R).
464464
intro H; clear I.
465465
rewrite <- R_INR_as_IR.
466466
cut (0 <= INR (nat_of_P p)).
467467
intro.
468+
apply Rlt_le.
468469
fourier.
469470
apply pos_INR.
470471
cut (0 <= INR (nat_of_P p)).
472+
rewrite INR_IZR_INZ.
471473
intro.
474+
change (IZR (Zneg p)) with (Ropp (IZR (Zpos p))) in I.
475+
rewrite <- positive_nat_Z in I.
472476
fourier.
473477
apply pos_INR.
474478
Qed.

coq_reals/Rreals_iso.v

Lines changed: 81 additions & 73 deletions
Original file line numberDiff line numberDiff line change
@@ -468,6 +468,85 @@ Proof.
468468
Qed.
469469
Hint Rewrite R_nring_as_IR : RtoIR.
470470

471+
Add Morphism RasIR with signature (@cs_eq _) ==> (@cs_eq _) as R_as_IR_wd.
472+
Proof.
473+
intros.
474+
rewrite H.
475+
reflexivity.
476+
Qed.
477+
478+
(** integers *)
479+
480+
Lemma R_pring_as_IR : forall x, RasIR (pring _ x) [=] pring _ x.
481+
Proof.
482+
intro x.
483+
(*rewrite pring_convert.*)
484+
stepr (nring (R := IR) (nat_of_P x)).
485+
stepl (RasIR (nring (R := RReals) (nat_of_P x))).
486+
apply R_nring_as_IR.
487+
apply R_as_IR_wd.
488+
symmetry.
489+
apply (pring_convert RReals x).
490+
symmetry.
491+
apply (pring_convert IR x).
492+
Qed.
493+
494+
Lemma R_zring_as_IR : forall x, RasIR (zring x) [=] zring x.
495+
Proof.
496+
induction x; simpl.
497+
apply R_Zero_as_IR.
498+
apply R_pring_as_IR.
499+
rewrite -> R_opp_as_IR.
500+
rewrite -> R_pring_as_IR.
501+
reflexivity.
502+
Qed.
503+
504+
Lemma INR_as_nring : forall x, INR x = nring (R:=RRing) x.
505+
Proof.
506+
induction x.
507+
reflexivity.
508+
simpl nring.
509+
rewrite <- IHx.
510+
apply S_INR.
511+
Qed.
512+
513+
Lemma IZR_as_zring : forall x, IZR x = zring (R:=RRing) x.
514+
Proof.
515+
induction x.
516+
reflexivity.
517+
rewrite <- positive_nat_Z, <- INR_IZR_INZ.
518+
rewrite INR_as_nring.
519+
(* rewrite pring_convert *)
520+
symmetry.
521+
rewrite positive_nat_Z.
522+
apply (pring_convert RRing p).
523+
change (IZR (Zneg p)) with (Ropp (IZR (Zpos p))).
524+
rewrite <- positive_nat_Z, <- INR_IZR_INZ.
525+
rewrite INR_as_nring.
526+
apply Ropp_eq_compat.
527+
symmetry.
528+
apply (pring_convert RRing p).
529+
Qed.
530+
531+
Lemma R_IZR_as_IR : forall x, RasIR (IZR x) [=] zring x.
532+
Proof.
533+
induction x.
534+
apply R_Zero_as_IR.
535+
rewrite <- positive_nat_Z, <- INR_IZR_INZ.
536+
rewrite R_INR_as_IR.
537+
rewrite -> R_nring_as_IR.
538+
auto with *.
539+
change (IZR (Zneg p)) with (Ropp (IZR (Zpos p))).
540+
rewrite <- positive_nat_Z, <- INR_IZR_INZ.
541+
rewrite -> R_opp_as_IR.
542+
rewrite R_INR_as_IR.
543+
rewrite -> R_nring_as_IR.
544+
simpl zring.
545+
auto with *.
546+
Qed.
547+
548+
Hint Rewrite R_IZR_as_IR : RtoIR.
549+
471550
(** exponents *)
472551

473552
Lemma R_pow_as_IR : forall x i,
@@ -714,13 +793,6 @@ Proof.
714793
apply: c.
715794
Qed.
716795

717-
Add Morphism RasIR with signature (@cs_eq _) ==> (@cs_eq _) as R_as_IR_wd.
718-
Proof.
719-
intros.
720-
rewrite H.
721-
reflexivity.
722-
Qed.
723-
724796
(** logarithm *)
725797

726798
Lemma R_ln_as_IR : forall x prf, RasIR (ln x) [=] Log (RasIR x) prf.
@@ -736,70 +808,6 @@ Proof.
736808
assumption.
737809
Qed.
738810

739-
(** integers *)
740-
741-
Lemma R_pring_as_IR : forall x, RasIR (pring _ x) [=] pring _ x.
742-
Proof.
743-
intro x.
744-
(*rewrite pring_convert.*)
745-
stepr (nring (R := IR) (nat_of_P x)).
746-
stepl (RasIR (nring (R := RReals) (nat_of_P x))).
747-
apply R_nring_as_IR.
748-
apply R_as_IR_wd.
749-
symmetry.
750-
apply (pring_convert RReals x).
751-
symmetry.
752-
apply (pring_convert IR x).
753-
Qed.
754-
755-
Lemma R_zring_as_IR : forall x, RasIR (zring x) [=] zring x.
756-
Proof.
757-
induction x; simpl.
758-
apply R_Zero_as_IR.
759-
apply R_pring_as_IR.
760-
rewrite -> R_opp_as_IR.
761-
rewrite -> R_pring_as_IR.
762-
reflexivity.
763-
Qed.
764-
765-
Lemma INR_as_nring : forall x, INR x = nring (R:=RRing) x.
766-
Proof.
767-
induction x.
768-
reflexivity.
769-
simpl nring.
770-
rewrite <- IHx.
771-
apply S_INR.
772-
Qed.
773-
774-
Lemma IZR_as_zring : forall x, IZR x = zring (R:=RRing) x.
775-
Proof.
776-
induction x; simpl.
777-
reflexivity.
778-
rewrite INR_as_nring.
779-
(* rewrite pring_convert *)
780-
symmetry.
781-
apply (pring_convert RRing p).
782-
rewrite INR_as_nring.
783-
apply Ropp_eq_compat.
784-
symmetry.
785-
apply (pring_convert RRing p).
786-
Qed.
787-
788-
Lemma R_IZR_as_IR : forall x, RasIR (IZR x) [=] zring x.
789-
Proof.
790-
induction x; simpl.
791-
apply R_Zero_as_IR.
792-
rewrite R_INR_as_IR.
793-
rewrite -> R_nring_as_IR.
794-
auto with *.
795-
rewrite -> R_opp_as_IR.
796-
rewrite R_INR_as_IR.
797-
rewrite -> R_nring_as_IR.
798-
auto with *.
799-
Qed.
800-
801-
Hint Rewrite R_IZR_as_IR : RtoIR.
802-
803811
(** pi *)
804812

805813
Lemma R_pi_as_IR : RasIR (PI) [=] Pi.
@@ -896,6 +904,7 @@ Proof.
896904
unfold PI_tg in prf.
897905
rewrite -> R_mult_as_IR.
898906
apply mult_wd.
907+
change 4%R with ((1 + 1) * (1 + 1))%R.
899908
rewrite -> R_mult_as_IR.
900909
rewrite -> R_plus_as_IR.
901910
rewrite -> R_One_as_IR.
@@ -943,7 +952,6 @@ Proof.
943952
intro q.
944953
destruct q.
945954
unfold Q2R.
946-
simpl.
947955
cut (Dom (f_rcpcl' IR) (RasIR (nring (R:=RRing) (nat_of_P Qden)))).
948956
intro Hy.
949957
stepr (RasIR (zring (R:=RRing) Qnum)[/]RasIR (nring (R:=RRing) (nat_of_P Qden))[//]Hy).
@@ -953,7 +961,7 @@ Proof.
953961
unfold Rdiv.
954962
replace (nring (R:=RRing) (nat_of_P Qden)) with (INR (nat_of_P Qden)).
955963
replace (zring (R:=RRing) Qnum) with (IZR Qnum).
956-
simpl; reflexivity.
964+
now rewrite <- positive_nat_Z, INR_IZR_INZ.
957965
apply IZR_as_zring.
958966
apply INR_as_nring.
959967
apply div_wd.

ode/Picard.v

Lines changed: 6 additions & 3 deletions
Original file line numberDiff line numberDiff line change
@@ -290,7 +290,8 @@ transitivity ('(abs (x - x0) * M)).
290290
intros t A.
291291
assert (A1 : mspc_ball rx x0 t) by
292292
(apply (mspc_ball_convex x0 x); [apply mspc_refl, (proj2_sig rx) | |]; trivial).
293-
apply extend_inside in A1. destruct A1 as [p A1]. rewrite A1. apply bounded.
293+
apply (extend_inside (Y:=CR)) in A1.
294+
destruct A1 as [p A1]. rewrite A1. apply bounded.
294295
+ apply CRle_Qle. change (abs (x - x0) * M ≤ ry). transitivity (`rx * M).
295296
- now apply (orders.order_preserving (.* M)), mspc_ball_Qabs_flip.
296297
- apply rx_M.
@@ -319,7 +320,7 @@ rewrite <- int_minus. transitivity ('(abs (x - x0) * (L * e))).
319320
intros x' B. assert (B1 : ball rx x0 x') by
320321
(apply (mspc_ball_convex x0 x); [apply mspc_refl | |]; solve_propholds).
321322
unfold plus, negate, ext_plus, ext_negate.
322-
apply extend_inside in B1. destruct B1 as [p B1]. rewrite !B1.
323+
apply (extend_inside (Y:=CR)) in B1. destruct B1 as [p B1]. rewrite !B1.
323324
apply mspc_ball_CRabs. unfold diag, together, Datatypes.id, Basics.compose; simpl.
324325
apply (lip_prf (λ y, v (_, y)) L), A.
325326
+ apply CRle_Qle. mc_setoid_replace (L * rx * e) with ((to_Q rx) * (L * e)) by ring.
@@ -366,7 +367,9 @@ Definition L : Q := 1.
366367
Instance : Bounded v M.
367368
Proof.
368369
intros [x [y H]]. unfold v; simpl. unfold M, ry, y0 in *.
369-
apply mspc_ball_CRabs, CRdistance_CRle in H. destruct H as [H1 H2].
370+
apply mspc_ball_CRabs in H.
371+
(* MS: why is this taking ages? *)
372+
apply <- (CRdistance_CRle 1) in H. destruct H as [H1 H2].
370373
change (1 - 1 ≤ y) in H1. change (y ≤ 1 + 1) in H2. change (abs y ≤ 2).
371374
rewrite plus_negate_r in H1. apply CRabs_AbsSmall. split; [| assumption].
372375
change (-2 ≤ y). transitivity (0 : CR); [| easy]. rewrite <- negate_0.

reals/fast/ModulusDerivative.v

Lines changed: 3 additions & 3 deletions
Original file line numberDiff line numberDiff line change
@@ -55,7 +55,7 @@ end.
5555

5656
Let properI : proper I.
5757
Proof.
58-
destruct l as [|l];destruct r as [|r]; try constructor.
58+
destruct l; destruct r; try constructor.
5959
simpl.
6060
apply inj_Q_less.
6161
assumption.
@@ -71,7 +71,7 @@ end.
7171

7272
Lemma ball_clamp : forall e a b, ball e a b -> ball e (clamp a) (clamp b).
7373
Proof.
74-
destruct l as [|l]; destruct r as [|r]; unfold clamp; intros e a b Hab; try apply uc_prf; apply Hab.
74+
destruct l; destruct r; unfold clamp; intros e a b Hab; try apply uc_prf; apply Hab.
7575
Qed.
7676

7777
Variable f f' : PartFunct IR.
@@ -91,7 +91,7 @@ Proof.
9191
assert (X:forall x, I (inj_Q _ (clamp x))).
9292
clear -I Hlr.
9393
intros x.
94-
destruct l as [|l];destruct r as [|r]; try split; unfold clamp; apply: inj_Q_leEq; simpl;
94+
destruct l; destruct r; try split; unfold clamp; apply: inj_Q_leEq; simpl;
9595
auto with *.
9696
assert (Y:=(fun a=> (Hg _ (Derivative_imp_inc _ _ _ _ Hf _ (X a)) (X a)))).
9797
do 2 rewrite -> Y.

reals/faster/ARbigD.v

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -1,5 +1,5 @@
11
Require Import
2-
Coq.Program.Program Coq.QArith.QArith Coq.ZArith.ZArith Coq.Numbers.Integer.BigZ.BigZ CoRN.model.structures.Qpossec
2+
Coq.Program.Program Coq.QArith.QArith Coq.ZArith.ZArith Bignums.BigZ.BigZ CoRN.model.structures.Qpossec
33
CoRN.metric2.MetricMorphisms CoRN.model.metric2.Qmetric CoRN.util.Qdlog CoRN.reals.faster.ARArith
44
MathClasses.theory.int_pow MathClasses.theory.nat_pow
55
MathClasses.implementations.stdlib_rationals MathClasses.implementations.stdlib_binary_integers MathClasses.implementations.fast_integers MathClasses.implementations.dyadics.

reals/faster/ARbigQ.v

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -1,5 +1,5 @@
11
Require Import
2-
Coq.Program.Program Coq.QArith.QArith Coq.ZArith.ZArith Coq.Numbers.Integer.BigZ.BigZ Coq.Numbers.Rational.BigQ.BigQ CoRN.model.structures.Qpossec
2+
Coq.Program.Program Coq.QArith.QArith Coq.ZArith.ZArith Bignums.BigZ.BigZ Bignums.BigQ.BigQ CoRN.model.structures.Qpossec
33
CoRN.reals.fast.Compress CoRN.reals.faster.ARQ
44
CoRN.metric2.MetricMorphisms CoRN.model.metric2.Qmetric CoRN.reals.faster.ARArith
55
MathClasses.implementations.stdlib_rationals MathClasses.implementations.stdlib_binary_integers MathClasses.implementations.field_of_fractions

0 commit comments

Comments
 (0)