@@ -468,6 +468,85 @@ Proof.
468
468
Qed .
469
469
Hint Rewrite R_nring_as_IR : RtoIR.
470
470
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
+
471
550
(** exponents *)
472
551
473
552
Lemma R_pow_as_IR : forall x i,
@@ -714,13 +793,6 @@ Proof.
714
793
apply: c.
715
794
Qed .
716
795
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
-
724
796
(** logarithm *)
725
797
726
798
Lemma R_ln_as_IR : forall x prf, RasIR (ln x) [=] Log (RasIR x) prf.
@@ -736,70 +808,6 @@ Proof.
736
808
assumption.
737
809
Qed .
738
810
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
-
803
811
(** pi *)
804
812
805
813
Lemma R_pi_as_IR : RasIR (PI) [=] Pi.
@@ -896,6 +904,7 @@ Proof.
896
904
unfold PI_tg in prf.
897
905
rewrite -> R_mult_as_IR.
898
906
apply mult_wd.
907
+ change 4%R with ((1 + 1) * (1 + 1))%R.
899
908
rewrite -> R_mult_as_IR.
900
909
rewrite -> R_plus_as_IR.
901
910
rewrite -> R_One_as_IR.
@@ -943,7 +952,6 @@ Proof.
943
952
intro q.
944
953
destruct q.
945
954
unfold Q2R.
946
- simpl.
947
955
cut (Dom (f_rcpcl' IR) (RasIR (nring (R:=RRing) (nat_of_P Qden)))).
948
956
intro Hy.
949
957
stepr (RasIR (zring (R:=RRing) Qnum)[/]RasIR (nring (R:=RRing) (nat_of_P Qden))[//]Hy).
@@ -953,7 +961,7 @@ Proof.
953
961
unfold Rdiv.
954
962
replace (nring (R:=RRing) (nat_of_P Qden)) with (INR (nat_of_P Qden)).
955
963
replace (zring (R:=RRing) Qnum) with (IZR Qnum).
956
- simpl; reflexivity .
964
+ now rewrite <- positive_nat_Z, INR_IZR_INZ .
957
965
apply IZR_as_zring.
958
966
apply INR_as_nring.
959
967
apply div_wd.
0 commit comments