summaryrefslogtreecommitdiff
path: root/theories/Reals/Rtrigo1.v
diff options
context:
space:
mode:
Diffstat (limited to 'theories/Reals/Rtrigo1.v')
-rw-r--r--theories/Reals/Rtrigo1.v383
1 files changed, 102 insertions, 281 deletions
diff --git a/theories/Reals/Rtrigo1.v b/theories/Reals/Rtrigo1.v
index 4d241863..bf00f736 100644
--- a/theories/Reals/Rtrigo1.v
+++ b/theories/Reals/Rtrigo1.v
@@ -1,9 +1,11 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* * The Coq Proof Assistant / The Coq Development Team *)
+(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *)
+(* <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 *)
+(* // * 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) *)
(************************************************************************)
Require Import Rbase.
@@ -182,13 +184,10 @@ destruct (pre_cos_bound _ 0 lo up) as [_ upper].
apply Rle_lt_trans with (1 := upper).
apply Rlt_le_trans with (2 := lower).
unfold cos_approx, sin_approx.
-simpl sum_f_R0; replace 7 with (IZR 7) by (simpl; field).
-replace 8 with (IZR 8) by (simpl; field).
+simpl sum_f_R0.
unfold cos_term, sin_term; simpl fact; rewrite !INR_IZR_INZ.
-simpl plus; simpl mult.
-field_simplify;
- try (repeat apply conj; apply not_eq_sym, Rlt_not_eq, (IZR_lt 0); reflexivity).
-unfold Rminus; rewrite !pow_IZR, <- !mult_IZR, <- !opp_IZR, <- ?plus_IZR.
+simpl plus; simpl mult; simpl Z_of_nat.
+field_simplify.
match goal with
|- IZR ?a / ?b < ?c / ?d =>
apply Rmult_lt_reg_r with d;[apply (IZR_lt 0); reflexivity |
@@ -198,7 +197,7 @@ match goal with
end.
unfold Rdiv; rewrite !Rmult_assoc, Rinv_l, Rmult_1_r;
[ | apply not_eq_sym, Rlt_not_eq, (IZR_lt 0); reflexivity].
-repeat (rewrite <- !plus_IZR || rewrite <- !mult_IZR).
+rewrite <- !mult_IZR.
apply IZR_lt; reflexivity.
Qed.
@@ -323,6 +322,7 @@ Lemma sin_PI : sin PI = 0.
Proof.
assert (H := sin2_cos2 PI).
rewrite cos_PI in H.
+ change (-1) with (-(1)) in H.
rewrite <- Rsqr_neg in H.
rewrite Rsqr_1 in H.
cut (Rsqr (sin PI) = 0).
@@ -533,9 +533,8 @@ Qed.
Lemma sin_PI_x : forall x:R, sin (PI - x) = sin x.
Proof.
- intro x; rewrite sin_minus; rewrite sin_PI; rewrite cos_PI; rewrite Rmult_0_l;
- unfold Rminus in |- *; rewrite Rplus_0_l; rewrite Ropp_mult_distr_l_reverse;
- rewrite Ropp_involutive; apply Rmult_1_l.
+ intro x; rewrite sin_minus; rewrite sin_PI; rewrite cos_PI.
+ ring.
Qed.
Lemma sin_period : forall (x:R) (k:nat), sin (x + 2 * INR k * PI) = sin x.
@@ -593,9 +592,9 @@ Proof.
generalize
(Rsqr_incrst_1 1 (sin x) H (Rlt_le 0 1 Rlt_0_1)
(Rlt_le 0 (sin x) (Rlt_trans 0 1 (sin x) Rlt_0_1 H)));
- rewrite Rsqr_1; intro; rewrite sin2 in H0; unfold Rminus in H0;
+ rewrite Rsqr_1; intro; rewrite sin2 in H0; unfold Rminus in H0.
generalize (Rplus_lt_compat_l (-1) 1 (1 + - Rsqr (cos x)) H0);
- repeat rewrite <- Rplus_assoc; repeat rewrite Rplus_opp_l;
+ repeat rewrite <- Rplus_assoc; change (-1) with (-(1)); rewrite Rplus_opp_l;
rewrite Rplus_0_l; intro; rewrite <- Ropp_0 in H1;
generalize (Ropp_lt_gt_contravar (-0) (- Rsqr (cos x)) H1);
repeat rewrite Ropp_involutive; intro; generalize (Rle_0_sqr (cos x));
@@ -603,6 +602,7 @@ Proof.
auto with real.
cut (sin x < -1).
intro; generalize (Ropp_lt_gt_contravar (sin x) (-1) H);
+ change (-1) with (-(1));
rewrite Ropp_involutive; clear H; intro;
generalize
(Rsqr_incrst_1 1 (- sin x) H (Rlt_le 0 1 Rlt_0_1)
@@ -610,7 +610,7 @@ Proof.
rewrite Rsqr_1; intro; rewrite <- Rsqr_neg in H0;
rewrite sin2 in H0; unfold Rminus in H0;
generalize (Rplus_lt_compat_l (-1) 1 (1 + - Rsqr (cos x)) H0);
- repeat rewrite <- Rplus_assoc; repeat rewrite Rplus_opp_l;
+ rewrite <- Rplus_assoc; change (-1) with (-(1)); rewrite Rplus_opp_l;
rewrite Rplus_0_l; intro; rewrite <- Ropp_0 in H1;
generalize (Ropp_lt_gt_contravar (-0) (- Rsqr (cos x)) H1);
repeat rewrite Ropp_involutive; intro; generalize (Rle_0_sqr (cos x));
@@ -696,41 +696,38 @@ Proof.
rewrite <- Rinv_l_sym.
do 2 rewrite Rmult_1_r; apply Rle_lt_trans with (INR (fact (2 * n + 1)) * 4).
apply Rmult_le_compat_l.
- replace 0 with (INR 0); [ idtac | reflexivity ]; apply le_INR; apply le_O_n.
- simpl in |- *; rewrite Rmult_1_r; replace 4 with (Rsqr 2);
- [ idtac | ring_Rsqr ]; replace (a * a) with (Rsqr a);
- [ idtac | reflexivity ]; apply Rsqr_incr_1.
+ apply pos_INR.
+ simpl in |- *; rewrite Rmult_1_r; change 4 with (Rsqr 2);
+ apply Rsqr_incr_1.
apply Rle_trans with (PI / 2);
[ assumption
| unfold Rdiv in |- *; apply Rmult_le_reg_l with 2;
[ prove_sup0
| rewrite <- Rmult_assoc; rewrite Rinv_r_simpl_m;
- [ replace 4 with 4; [ apply PI_4 | ring ] | discrR ] ] ].
+ [ apply PI_4 | discrR ] ] ].
left; assumption.
left; prove_sup0.
rewrite H1; replace (2 * n + 1 + 2)%nat with (S (S (2 * n + 1))).
do 2 rewrite fact_simpl; do 2 rewrite mult_INR.
repeat rewrite <- Rmult_assoc.
rewrite <- (Rmult_comm (INR (fact (2 * n + 1)))).
- rewrite Rmult_assoc.
apply Rmult_lt_compat_l.
apply lt_INR_0; apply neq_O_lt.
assert (H2 := fact_neq_0 (2 * n + 1)).
red in |- *; intro; elim H2; symmetry in |- *; assumption.
do 2 rewrite S_INR; rewrite plus_INR; rewrite mult_INR; set (x := INR n);
unfold INR in |- *.
- replace ((2 * x + 1 + 1 + 1) * (2 * x + 1 + 1)) with (4 * x * x + 10 * x + 6);
+ replace (((1 + 1) * x + 1 + 1 + 1) * ((1 + 1) * x + 1 + 1)) with (4 * x * x + 10 * x + 6);
[ idtac | ring ].
- apply Rplus_lt_reg_l with (-4); rewrite Rplus_opp_l;
- replace (-4 + (4 * x * x + 10 * x + 6)) with (4 * x * x + 10 * x + 2);
+ apply Rplus_lt_reg_l with (-(4)); rewrite Rplus_opp_l;
+ replace (-(4) + (4 * x * x + 10 * x + 6)) with (4 * x * x + 10 * x + 2);
[ idtac | ring ].
apply Rplus_le_lt_0_compat.
cut (0 <= x).
intro; apply Rplus_le_le_0_compat; repeat apply Rmult_le_pos;
assumption || left; prove_sup.
- unfold x in |- *; replace 0 with (INR 0);
- [ apply le_INR; apply le_O_n | reflexivity ].
- prove_sup0.
+ apply pos_INR.
+ now apply IZR_lt.
ring.
apply INR_fact_neq_0.
apply INR_fact_neq_0.
@@ -738,39 +735,33 @@ Proof.
Qed.
Lemma SIN : forall a:R, 0 <= a -> a <= PI -> sin_lb a <= sin a <= sin_ub a.
+Proof.
intros; unfold sin_lb, sin_ub in |- *; apply (sin_bound a 1 H H0).
Qed.
Lemma COS :
forall a:R, - PI / 2 <= a -> a <= PI / 2 -> cos_lb a <= cos a <= cos_ub a.
+Proof.
intros; unfold cos_lb, cos_ub in |- *; apply (cos_bound a 1 H H0).
Qed.
(**********)
Lemma _PI2_RLT_0 : - (PI / 2) < 0.
Proof.
- rewrite <- Ropp_0; apply Ropp_lt_contravar; apply PI2_RGT_0.
+ assert (H := PI_RGT_0).
+ fourier.
Qed.
Lemma PI4_RLT_PI2 : PI / 4 < PI / 2.
Proof.
- unfold Rdiv in |- *; apply Rmult_lt_compat_l.
- apply PI_RGT_0.
- apply Rinv_lt_contravar.
- apply Rmult_lt_0_compat; prove_sup0.
- pattern 2 at 1 in |- *; rewrite <- Rplus_0_r.
- replace 4 with (2 + 2); [ apply Rplus_lt_compat_l; prove_sup0 | ring ].
+ assert (H := PI_RGT_0).
+ fourier.
Qed.
Lemma PI2_Rlt_PI : PI / 2 < PI.
Proof.
- unfold Rdiv in |- *; pattern PI at 2 in |- *; rewrite <- Rmult_1_r.
- apply Rmult_lt_compat_l.
- apply PI_RGT_0.
- pattern 1 at 3 in |- *; rewrite <- Rinv_1; apply Rinv_lt_contravar.
- rewrite Rmult_1_l; prove_sup0.
- pattern 1 at 1 in |- *; rewrite <- Rplus_0_r; apply Rplus_lt_compat_l;
- apply Rlt_0_1.
+ assert (H := PI_RGT_0).
+ fourier.
Qed.
(***************************************************)
@@ -787,12 +778,10 @@ Proof.
rewrite H3; rewrite sin_PI2; apply Rlt_0_1.
rewrite <- sin_PI_x; generalize (Ropp_gt_lt_contravar x (PI / 2) H3);
intro H4; generalize (Rplus_lt_compat_l PI (- x) (- (PI / 2)) H4).
- replace (PI + - x) with (PI - x).
replace (PI + - (PI / 2)) with (PI / 2).
intro H5; generalize (Ropp_lt_gt_contravar x PI H0); intro H6;
change (- PI < - x) in H6; generalize (Rplus_lt_compat_l PI (- PI) (- x) H6).
rewrite Rplus_opp_r.
- replace (PI + - x) with (PI - x).
intro H7;
elim
(SIN (PI - x) (Rlt_le 0 (PI - x) H7)
@@ -800,9 +789,7 @@ Proof.
intros H8 _;
generalize (sin_lb_gt_0 (PI - x) H7 (Rlt_le (PI - x) (PI / 2) H5));
intro H9; apply (Rlt_le_trans 0 (sin_lb (PI - x)) (sin (PI - x)) H9 H8).
- reflexivity.
- pattern PI at 2 in |- *; rewrite double_var; ring.
- reflexivity.
+ field.
Qed.
Theorem cos_gt_0 : forall x:R, - (PI / 2) < x -> x < PI / 2 -> 0 < cos x.
@@ -855,16 +842,12 @@ Proof.
rewrite <- (Ropp_involutive (cos x)); apply Ropp_le_ge_contravar;
rewrite <- neg_cos; replace (x + PI) with (x - PI + 2 * INR 1 * PI).
rewrite cos_period; apply cos_ge_0.
- replace (- (PI / 2)) with (- PI + PI / 2).
+ replace (- (PI / 2)) with (- PI + PI / 2) by field.
unfold Rminus in |- *; rewrite (Rplus_comm x); apply Rplus_le_compat_l;
assumption.
- pattern PI at 1 in |- *; rewrite (double_var PI); rewrite Ropp_plus_distr;
- ring.
unfold Rminus in |- *; rewrite Rplus_comm;
- replace (PI / 2) with (- PI + 3 * (PI / 2)).
+ replace (PI / 2) with (- PI + 3 * (PI / 2)) by field.
apply Rplus_le_compat_l; assumption.
- pattern PI at 1 in |- *; rewrite (double_var PI); rewrite Ropp_plus_distr;
- ring.
unfold INR in |- *; ring.
Qed.
@@ -905,16 +888,12 @@ Proof.
apply Ropp_lt_gt_contravar; rewrite <- neg_cos;
replace (x + PI) with (x - PI + 2 * INR 1 * PI).
rewrite cos_period; apply cos_gt_0.
- replace (- (PI / 2)) with (- PI + PI / 2).
+ replace (- (PI / 2)) with (- PI + PI / 2) by field.
unfold Rminus in |- *; rewrite (Rplus_comm x); apply Rplus_lt_compat_l;
assumption.
- pattern PI at 1 in |- *; rewrite (double_var PI); rewrite Ropp_plus_distr;
- ring.
unfold Rminus in |- *; rewrite Rplus_comm;
- replace (PI / 2) with (- PI + 3 * (PI / 2)).
+ replace (PI / 2) with (- PI + 3 * (PI / 2)) by field.
apply Rplus_lt_compat_l; assumption.
- pattern PI at 1 in |- *; rewrite (double_var PI); rewrite Ropp_plus_distr;
- ring.
unfold INR in |- *; ring.
Qed.
@@ -951,7 +930,7 @@ Lemma cos_ge_0_3PI2 :
forall x:R, 3 * (PI / 2) <= x -> x <= 2 * PI -> 0 <= cos x.
Proof.
intros; rewrite <- cos_neg; rewrite <- (cos_period (- x) 1);
- unfold INR in |- *; replace (- x + 2 * 1 * PI) with (2 * PI - x).
+ unfold INR in |- *; replace (- x + 2 * 1 * PI) with (2 * PI - x) by ring.
generalize (Ropp_le_ge_contravar x (2 * PI) H0); intro H1;
generalize (Rge_le (- x) (- (2 * PI)) H1); clear H1;
intro H1; generalize (Rplus_le_compat_l (2 * PI) (- (2 * PI)) (- x) H1).
@@ -960,36 +939,30 @@ Proof.
generalize (Rge_le (- (3 * (PI / 2))) (- x) H3); clear H3;
intro H3;
generalize (Rplus_le_compat_l (2 * PI) (- x) (- (3 * (PI / 2))) H3).
- replace (2 * PI + - (3 * (PI / 2))) with (PI / 2).
+ replace (2 * PI + - (3 * (PI / 2))) with (PI / 2) by field.
intro H4;
apply
(cos_ge_0 (2 * PI - x)
(Rlt_le (- (PI / 2)) (2 * PI - x)
(Rlt_le_trans (- (PI / 2)) 0 (2 * PI - x) _PI2_RLT_0 H2)) H4).
- rewrite double; pattern PI at 2 3 in |- *; rewrite double_var; ring.
- ring.
Qed.
Lemma form1 :
forall p q:R, cos p + cos q = 2 * cos ((p - q) / 2) * cos ((p + q) / 2).
Proof.
intros p q; pattern p at 1 in |- *;
- replace p with ((p - q) / 2 + (p + q) / 2).
- rewrite <- (cos_neg q); replace (- q) with ((p - q) / 2 - (p + q) / 2).
+ replace p with ((p - q) / 2 + (p + q) / 2) by field.
+ rewrite <- (cos_neg q); replace (- q) with ((p - q) / 2 - (p + q) / 2) by field.
rewrite cos_plus; rewrite cos_minus; ring.
- pattern q at 3 in |- *; rewrite double_var; unfold Rdiv in |- *; ring.
- pattern p at 3 in |- *; rewrite double_var; unfold Rdiv in |- *; ring.
Qed.
Lemma form2 :
forall p q:R, cos p - cos q = -2 * sin ((p - q) / 2) * sin ((p + q) / 2).
Proof.
intros p q; pattern p at 1 in |- *;
- replace p with ((p - q) / 2 + (p + q) / 2).
- rewrite <- (cos_neg q); replace (- q) with ((p - q) / 2 - (p + q) / 2).
+ replace p with ((p - q) / 2 + (p + q) / 2) by field.
+ rewrite <- (cos_neg q); replace (- q) with ((p - q) / 2 - (p + q) / 2) by field.
rewrite cos_plus; rewrite cos_minus; ring.
- pattern q at 3 in |- *; rewrite double_var; unfold Rdiv in |- *; ring.
- pattern p at 3 in |- *; rewrite double_var; unfold Rdiv in |- *; ring.
Qed.
Lemma form3 :
@@ -1007,11 +980,9 @@ Lemma form4 :
forall p q:R, sin p - sin q = 2 * cos ((p + q) / 2) * sin ((p - q) / 2).
Proof.
intros p q; pattern p at 1 in |- *;
- replace p with ((p - q) / 2 + (p + q) / 2).
- pattern q at 3 in |- *; replace q with ((p + q) / 2 - (p - q) / 2).
+ replace p with ((p - q) / 2 + (p + q) / 2) by field.
+ pattern q at 3 in |- *; replace q with ((p + q) / 2 - (p - q) / 2) by field.
rewrite sin_plus; rewrite sin_minus; ring.
- pattern q at 3 in |- *; rewrite double_var; unfold Rdiv in |- *; ring.
- pattern p at 3 in |- *; rewrite double_var; unfold Rdiv in |- *; ring.
Qed.
@@ -1067,13 +1038,13 @@ Proof.
repeat rewrite (Rmult_comm (/ 2)).
clear H4; intro H4;
generalize (Rplus_le_compat (- (PI / 2)) x (- (PI / 2)) y H H1);
- replace (- (PI / 2) + - (PI / 2)) with (- PI).
+ replace (- (PI / 2) + - (PI / 2)) with (- PI) by field.
intro H5;
generalize
(Rmult_le_compat_l (/ 2) (- PI) (x + y)
(Rlt_le 0 (/ 2) (Rinv_0_lt_compat 2 Hyp)) H5).
- replace (/ 2 * (x + y)) with ((x + y) / 2).
- replace (/ 2 * - PI) with (- (PI / 2)).
+ replace (/ 2 * (x + y)) with ((x + y) / 2) by apply Rmult_comm.
+ replace (/ 2 * - PI) with (- (PI / 2)) by field.
clear H5; intro H5; elim H4; intro H40.
elim H5; intro H50.
generalize (cos_gt_0 ((x + y) / 2) H50 H40); intro H6;
@@ -1095,13 +1066,6 @@ Proof.
rewrite H40 in H3; assert (H50 := cos_PI2); unfold Rdiv in H50;
rewrite H50 in H3; rewrite Rmult_0_r in H3; rewrite Rmult_0_l in H3;
elim (Rlt_irrefl 0 H3).
- unfold Rdiv in |- *.
- rewrite <- Ropp_mult_distr_l_reverse.
- apply Rmult_comm.
- unfold Rdiv in |- *; apply Rmult_comm.
- pattern PI at 1 in |- *; rewrite double_var.
- rewrite Ropp_plus_distr.
- reflexivity.
Qed.
Lemma sin_increasing_1 :
@@ -1111,43 +1075,42 @@ Lemma sin_increasing_1 :
Proof.
intros; generalize (Rplus_lt_compat_l x x y H3); intro H4;
generalize (Rplus_le_compat (- (PI / 2)) x (- (PI / 2)) x H H);
- replace (- (PI / 2) + - (PI / 2)) with (- PI).
+ replace (- (PI / 2) + - (PI / 2)) with (- PI) by field.
assert (Hyp : 0 < 2).
prove_sup0.
intro H5; generalize (Rle_lt_trans (- PI) (x + x) (x + y) H5 H4); intro H6;
generalize
(Rmult_lt_compat_l (/ 2) (- PI) (x + y) (Rinv_0_lt_compat 2 Hyp) H6);
- replace (/ 2 * - PI) with (- (PI / 2)).
- replace (/ 2 * (x + y)) with ((x + y) / 2).
+ replace (/ 2 * - PI) with (- (PI / 2)) by field.
+ replace (/ 2 * (x + y)) with ((x + y) / 2) by apply Rmult_comm.
clear H4 H5 H6; intro H4; generalize (Rplus_lt_compat_l y x y H3); intro H5;
rewrite Rplus_comm in H5;
generalize (Rplus_le_compat y (PI / 2) y (PI / 2) H2 H2).
rewrite <- double_var.
intro H6; generalize (Rlt_le_trans (x + y) (y + y) PI H5 H6); intro H7;
generalize (Rmult_lt_compat_l (/ 2) (x + y) PI (Rinv_0_lt_compat 2 Hyp) H7);
- replace (/ 2 * PI) with (PI / 2).
- replace (/ 2 * (x + y)) with ((x + y) / 2).
+ replace (/ 2 * PI) with (PI / 2) by apply Rmult_comm.
+ replace (/ 2 * (x + y)) with ((x + y) / 2) by apply Rmult_comm.
clear H5 H6 H7; intro H5; generalize (Ropp_le_ge_contravar (- (PI / 2)) y H1);
rewrite Ropp_involutive; clear H1; intro H1;
generalize (Rge_le (PI / 2) (- y) H1); clear H1; intro H1;
generalize (Ropp_le_ge_contravar y (PI / 2) H2); clear H2;
intro H2; generalize (Rge_le (- y) (- (PI / 2)) H2);
clear H2; intro H2; generalize (Rplus_lt_compat_l (- y) x y H3);
- replace (- y + x) with (x - y).
+ replace (- y + x) with (x - y) by apply Rplus_comm.
rewrite Rplus_opp_l.
intro H6;
generalize (Rmult_lt_compat_l (/ 2) (x - y) 0 (Rinv_0_lt_compat 2 Hyp) H6);
- rewrite Rmult_0_r; replace (/ 2 * (x - y)) with ((x - y) / 2).
+ rewrite Rmult_0_r; replace (/ 2 * (x - y)) with ((x - y) / 2) by apply Rmult_comm.
clear H6; intro H6;
generalize (Rplus_le_compat (- (PI / 2)) x (- (PI / 2)) (- y) H H2);
- replace (- (PI / 2) + - (PI / 2)) with (- PI).
- replace (x + - y) with (x - y).
+ replace (- (PI / 2) + - (PI / 2)) with (- PI) by field.
intro H7;
generalize
(Rmult_le_compat_l (/ 2) (- PI) (x - y)
(Rlt_le 0 (/ 2) (Rinv_0_lt_compat 2 Hyp)) H7);
- replace (/ 2 * - PI) with (- (PI / 2)).
- replace (/ 2 * (x - y)) with ((x - y) / 2).
+ replace (/ 2 * - PI) with (- (PI / 2)) by field.
+ replace (/ 2 * (x - y)) with ((x - y) / 2) by apply Rmult_comm.
clear H7; intro H7; clear H H0 H1 H2; apply Rminus_lt; rewrite form4;
generalize (cos_gt_0 ((x + y) / 2) H4 H5); intro H8;
generalize (Rmult_lt_0_compat 2 (cos ((x + y) / 2)) Hyp H8);
@@ -1162,23 +1125,6 @@ Proof.
2 * cos ((x + y) / 2)) H10 H8); intro H11; rewrite Rmult_0_r in H11;
rewrite Rmult_comm; assumption.
apply Ropp_lt_gt_contravar; apply PI2_Rlt_PI.
- unfold Rdiv in |- *; apply Rmult_comm.
- unfold Rdiv in |- *; rewrite <- Ropp_mult_distr_l_reverse; apply Rmult_comm.
- reflexivity.
- pattern PI at 1 in |- *; rewrite double_var.
- rewrite Ropp_plus_distr.
- reflexivity.
- unfold Rdiv in |- *; apply Rmult_comm.
- unfold Rminus in |- *; apply Rplus_comm.
- unfold Rdiv in |- *; apply Rmult_comm.
- unfold Rdiv in |- *; apply Rmult_comm.
- unfold Rdiv in |- *; apply Rmult_comm.
- unfold Rdiv in |- *.
- rewrite <- Ropp_mult_distr_l_reverse.
- apply Rmult_comm.
- pattern PI at 1 in |- *; rewrite double_var.
- rewrite Ropp_plus_distr.
- reflexivity.
Qed.
Lemma sin_decreasing_0 :
@@ -1193,33 +1139,16 @@ Proof.
generalize (Rplus_le_compat_l (- PI) (PI / 2) x H0);
generalize (Rplus_le_compat_l (- PI) y (3 * (PI / 2)) H1);
generalize (Rplus_le_compat_l (- PI) (PI / 2) y H2);
- replace (- PI + x) with (x - PI).
- replace (- PI + PI / 2) with (- (PI / 2)).
- replace (- PI + y) with (y - PI).
- replace (- PI + 3 * (PI / 2)) with (PI / 2).
- replace (- (PI - x)) with (x - PI).
- replace (- (PI - y)) with (y - PI).
+ replace (- PI + x) with (x - PI) by apply Rplus_comm.
+ replace (- PI + PI / 2) with (- (PI / 2)) by field.
+ replace (- PI + y) with (y - PI) by apply Rplus_comm.
+ replace (- PI + 3 * (PI / 2)) with (PI / 2) by field.
+ replace (- (PI - x)) with (x - PI) by ring.
+ replace (- (PI - y)) with (y - PI) by ring.
intros; change (sin (y - PI) < sin (x - PI)) in H8;
- apply Rplus_lt_reg_l with (- PI); rewrite Rplus_comm;
- replace (y + - PI) with (y - PI).
- rewrite Rplus_comm; replace (x + - PI) with (x - PI).
+ apply Rplus_lt_reg_l with (- PI); rewrite Rplus_comm.
+ rewrite (Rplus_comm _ x).
apply (sin_increasing_0 (y - PI) (x - PI) H4 H5 H6 H7 H8).
- reflexivity.
- reflexivity.
- unfold Rminus in |- *; rewrite Ropp_plus_distr.
- rewrite Ropp_involutive.
- apply Rplus_comm.
- unfold Rminus in |- *; rewrite Ropp_plus_distr.
- rewrite Ropp_involutive.
- apply Rplus_comm.
- pattern PI at 2 in |- *; rewrite double_var.
- rewrite Ropp_plus_distr.
- ring.
- unfold Rminus in |- *; apply Rplus_comm.
- pattern PI at 2 in |- *; rewrite double_var.
- rewrite Ropp_plus_distr.
- ring.
- unfold Rminus in |- *; apply Rplus_comm.
Qed.
Lemma sin_decreasing_1 :
@@ -1233,24 +1162,14 @@ Proof.
generalize (Rplus_le_compat_l (- PI) y (3 * (PI / 2)) H1);
generalize (Rplus_le_compat_l (- PI) (PI / 2) y H2);
generalize (Rplus_lt_compat_l (- PI) x y H3);
- replace (- PI + PI / 2) with (- (PI / 2)).
- replace (- PI + y) with (y - PI).
- replace (- PI + 3 * (PI / 2)) with (PI / 2).
- replace (- PI + x) with (x - PI).
+ replace (- PI + PI / 2) with (- (PI / 2)) by field.
+ replace (- PI + y) with (y - PI) by apply Rplus_comm.
+ replace (- PI + 3 * (PI / 2)) with (PI / 2) by field.
+ replace (- PI + x) with (x - PI) by apply Rplus_comm.
intros; apply Ropp_lt_cancel; repeat rewrite <- sin_neg;
- replace (- (PI - x)) with (x - PI).
- replace (- (PI - y)) with (y - PI).
+ replace (- (PI - x)) with (x - PI) by ring.
+ replace (- (PI - y)) with (y - PI) by ring.
apply (sin_increasing_1 (x - PI) (y - PI) H7 H8 H5 H6 H4).
- unfold Rminus in |- *; rewrite Ropp_plus_distr.
- rewrite Ropp_involutive.
- apply Rplus_comm.
- unfold Rminus in |- *; rewrite Ropp_plus_distr.
- rewrite Ropp_involutive.
- apply Rplus_comm.
- unfold Rminus in |- *; apply Rplus_comm.
- pattern PI at 2 in |- *; rewrite double_var; ring.
- unfold Rminus in |- *; apply Rplus_comm.
- pattern PI at 2 in |- *; rewrite double_var; ring.
Qed.
Lemma cos_increasing_0 :
@@ -1260,44 +1179,22 @@ Proof.
intros x y H1 H2 H3 H4; rewrite <- (cos_neg x); rewrite <- (cos_neg y);
rewrite <- (cos_period (- x) 1); rewrite <- (cos_period (- y) 1);
unfold INR in |- *;
- replace (- x + 2 * 1 * PI) with (PI / 2 - (x - 3 * (PI / 2))).
- replace (- y + 2 * 1 * PI) with (PI / 2 - (y - 3 * (PI / 2))).
+ replace (- x + 2 * 1 * PI) with (PI / 2 - (x - 3 * (PI / 2))) by field.
+ replace (- y + 2 * 1 * PI) with (PI / 2 - (y - 3 * (PI / 2))) by field.
repeat rewrite cos_shift; intro H5;
generalize (Rplus_le_compat_l (-3 * (PI / 2)) PI x H1);
generalize (Rplus_le_compat_l (-3 * (PI / 2)) x (2 * PI) H2);
generalize (Rplus_le_compat_l (-3 * (PI / 2)) PI y H3);
generalize (Rplus_le_compat_l (-3 * (PI / 2)) y (2 * PI) H4).
- replace (-3 * (PI / 2) + y) with (y - 3 * (PI / 2)).
- replace (-3 * (PI / 2) + x) with (x - 3 * (PI / 2)).
- replace (-3 * (PI / 2) + 2 * PI) with (PI / 2).
- replace (-3 * (PI / 2) + PI) with (- (PI / 2)).
+ replace (-3 * (PI / 2) + y) with (y - 3 * (PI / 2)) by ring.
+ replace (-3 * (PI / 2) + x) with (x - 3 * (PI / 2)) by ring.
+ replace (-3 * (PI / 2) + 2 * PI) with (PI / 2) by field.
+ replace (-3 * (PI / 2) + PI) with (- (PI / 2)) by field.
clear H1 H2 H3 H4; intros H1 H2 H3 H4;
apply Rplus_lt_reg_l with (-3 * (PI / 2));
- replace (-3 * (PI / 2) + x) with (x - 3 * (PI / 2)).
- replace (-3 * (PI / 2) + y) with (y - 3 * (PI / 2)).
+ replace (-3 * (PI / 2) + x) with (x - 3 * (PI / 2)) by ring.
+ replace (-3 * (PI / 2) + y) with (y - 3 * (PI / 2)) by ring.
apply (sin_increasing_0 (x - 3 * (PI / 2)) (y - 3 * (PI / 2)) H4 H3 H2 H1 H5).
- unfold Rminus in |- *.
- rewrite Ropp_mult_distr_l_reverse.
- apply Rplus_comm.
- unfold Rminus in |- *.
- rewrite Ropp_mult_distr_l_reverse.
- apply Rplus_comm.
- pattern PI at 3 in |- *; rewrite double_var.
- ring.
- rewrite double; pattern PI at 3 4 in |- *; rewrite double_var.
- ring.
- unfold Rminus in |- *.
- rewrite Ropp_mult_distr_l_reverse.
- apply Rplus_comm.
- unfold Rminus in |- *.
- rewrite Ropp_mult_distr_l_reverse.
- apply Rplus_comm.
- rewrite Rmult_1_r.
- rewrite (double PI); pattern PI at 3 4 in |- *; rewrite double_var.
- ring.
- rewrite Rmult_1_r.
- rewrite (double PI); pattern PI at 3 4 in |- *; rewrite double_var.
- ring.
Qed.
Lemma cos_increasing_1 :
@@ -1312,31 +1209,16 @@ Proof.
generalize (Rplus_lt_compat_l (-3 * (PI / 2)) x y H5);
rewrite <- (cos_neg x); rewrite <- (cos_neg y);
rewrite <- (cos_period (- x) 1); rewrite <- (cos_period (- y) 1);
- unfold INR in |- *; replace (-3 * (PI / 2) + x) with (x - 3 * (PI / 2)).
- replace (-3 * (PI / 2) + y) with (y - 3 * (PI / 2)).
- replace (-3 * (PI / 2) + PI) with (- (PI / 2)).
- replace (-3 * (PI / 2) + 2 * PI) with (PI / 2).
+ unfold INR in |- *; replace (-3 * (PI / 2) + x) with (x - 3 * (PI / 2)) by ring.
+ replace (-3 * (PI / 2) + y) with (y - 3 * (PI / 2)) by ring.
+ replace (-3 * (PI / 2) + PI) with (- (PI / 2)) by field.
+ replace (-3 * (PI / 2) + 2 * PI) with (PI / 2) by field.
clear H1 H2 H3 H4 H5; intros H1 H2 H3 H4 H5;
- replace (- x + 2 * 1 * PI) with (PI / 2 - (x - 3 * (PI / 2))).
- replace (- y + 2 * 1 * PI) with (PI / 2 - (y - 3 * (PI / 2))).
+ replace (- x + 2 * 1 * PI) with (PI / 2 - (x - 3 * (PI / 2))) by field.
+ replace (- y + 2 * 1 * PI) with (PI / 2 - (y - 3 * (PI / 2))) by field.
repeat rewrite cos_shift;
apply
(sin_increasing_1 (x - 3 * (PI / 2)) (y - 3 * (PI / 2)) H5 H4 H3 H2 H1).
- rewrite Rmult_1_r.
- rewrite (double PI); pattern PI at 3 4 in |- *; rewrite double_var.
- ring.
- rewrite Rmult_1_r.
- rewrite (double PI); pattern PI at 3 4 in |- *; rewrite double_var.
- ring.
- rewrite (double PI); pattern PI at 3 4 in |- *; rewrite double_var.
- ring.
- pattern PI at 3 in |- *; rewrite double_var; ring.
- unfold Rminus in |- *.
- rewrite <- Ropp_mult_distr_l_reverse.
- apply Rplus_comm.
- unfold Rminus in |- *.
- rewrite <- Ropp_mult_distr_l_reverse.
- apply Rplus_comm.
Qed.
Lemma cos_decreasing_0 :
@@ -1375,31 +1257,8 @@ Lemma tan_diff :
cos x <> 0 -> cos y <> 0 -> tan x - tan y = sin (x - y) / (cos x * cos y).
Proof.
intros; unfold tan in |- *; rewrite sin_minus.
- unfold Rdiv in |- *.
- unfold Rminus in |- *.
- rewrite Rmult_plus_distr_r.
- rewrite Rinv_mult_distr.
- repeat rewrite (Rmult_comm (sin x)).
- repeat rewrite Rmult_assoc.
- rewrite (Rmult_comm (cos y)).
- repeat rewrite Rmult_assoc.
- rewrite <- Rinv_l_sym.
- rewrite Rmult_1_r.
- rewrite (Rmult_comm (sin x)).
- apply Rplus_eq_compat_l.
- rewrite <- Ropp_mult_distr_l_reverse.
- rewrite <- Ropp_mult_distr_r_reverse.
- rewrite (Rmult_comm (/ cos x)).
- repeat rewrite Rmult_assoc.
- rewrite (Rmult_comm (cos x)).
- repeat rewrite Rmult_assoc.
- rewrite <- Rinv_l_sym.
- rewrite Rmult_1_r.
- reflexivity.
- assumption.
- assumption.
- assumption.
- assumption.
+ field.
+ now split.
Qed.
Lemma tan_increasing_0 :
@@ -1436,10 +1295,9 @@ Proof.
intro H11; generalize (Rge_le (- y) (- (PI / 4)) H11);
clear H11; intro H11;
generalize (Rplus_le_compat (- (PI / 4)) x (- (PI / 4)) (- y) H H11);
- generalize (Rplus_le_compat x (PI / 4) (- y) (PI / 4) H0 H10);
- replace (x + - y) with (x - y).
- replace (PI / 4 + PI / 4) with (PI / 2).
- replace (- (PI / 4) + - (PI / 4)) with (- (PI / 2)).
+ generalize (Rplus_le_compat x (PI / 4) (- y) (PI / 4) H0 H10).
+ replace (PI / 4 + PI / 4) with (PI / 2) by field.
+ replace (- (PI / 4) + - (PI / 4)) with (- (PI / 2)) by field.
intros; case (Rtotal_order 0 (x - y)); intro H14.
generalize
(sin_gt_0 (x - y) H14 (Rle_lt_trans (x - y) (PI / 2) PI H12 PI2_Rlt_PI));
@@ -1447,28 +1305,6 @@ Proof.
elim H14; intro H15.
rewrite <- H15 in H9; rewrite sin_0 in H9; elim (Rlt_irrefl 0 H9).
apply Rminus_lt; assumption.
- pattern PI at 1 in |- *; rewrite double_var.
- unfold Rdiv in |- *.
- rewrite Rmult_plus_distr_r.
- repeat rewrite Rmult_assoc.
- rewrite <- Rinv_mult_distr.
- rewrite Ropp_plus_distr.
- replace 4 with 4.
- reflexivity.
- ring.
- discrR.
- discrR.
- pattern PI at 1 in |- *; rewrite double_var.
- unfold Rdiv in |- *.
- rewrite Rmult_plus_distr_r.
- repeat rewrite Rmult_assoc.
- rewrite <- Rinv_mult_distr.
- replace 4 with 4.
- reflexivity.
- ring.
- discrR.
- discrR.
- reflexivity.
case (Rcase_abs (sin (x - y))); intro H9.
assumption.
generalize (Rge_le (sin (x - y)) 0 H9); clear H9; intro H9;
@@ -1482,8 +1318,7 @@ Proof.
(Rlt_le 0 (/ (cos x * cos y)) H12)); intro H13;
elim
(Rlt_irrefl 0 (Rle_lt_trans 0 (sin (x - y) * / (cos x * cos y)) 0 H13 H3)).
- rewrite Rinv_mult_distr.
- reflexivity.
+ apply Rinv_mult_distr.
assumption.
assumption.
Qed.
@@ -1521,9 +1356,8 @@ Proof.
clear H10 H11; intro H8; generalize (Ropp_le_ge_contravar y (PI / 4) H2);
intro H11; generalize (Rge_le (- y) (- (PI / 4)) H11);
clear H11; intro H11;
- generalize (Rplus_le_compat (- (PI / 4)) x (- (PI / 4)) (- y) H H11);
- replace (x + - y) with (x - y).
- replace (- (PI / 4) + - (PI / 4)) with (- (PI / 2)).
+ generalize (Rplus_le_compat (- (PI / 4)) x (- (PI / 4)) (- y) H H11).
+ replace (- (PI / 4) + - (PI / 4)) with (- (PI / 2)) by field.
clear H11; intro H9; generalize (Rlt_minus x y H3); clear H3; intro H3;
clear H H0 H1 H2 H4 H5 HP1 HP2; generalize PI2_Rlt_PI;
intro H1; generalize (Ropp_lt_gt_contravar (PI / 2) PI H1);
@@ -1534,18 +1368,6 @@ Proof.
generalize
(Rmult_lt_gt_compat_neg_l (sin (x - y)) 0 (/ (cos x * cos y)) H2 H8);
rewrite Rmult_0_r; intro H4; assumption.
- pattern PI at 1 in |- *; rewrite double_var.
- unfold Rdiv in |- *.
- rewrite Rmult_plus_distr_r.
- repeat rewrite Rmult_assoc.
- rewrite <- Rinv_mult_distr.
- replace 4 with 4.
- rewrite Ropp_plus_distr.
- reflexivity.
- ring.
- discrR.
- discrR.
- reflexivity.
apply Rinv_mult_distr; assumption.
Qed.
@@ -1737,7 +1559,7 @@ Proof.
rewrite H5.
rewrite mult_INR.
simpl in |- *.
- rewrite <- (Rplus_0_l (2 * INR x2 * PI)).
+ rewrite <- (Rplus_0_l ((1 + 1) * INR x2 * PI)).
rewrite sin_period.
apply sin_0.
rewrite H5.
@@ -1747,7 +1569,7 @@ Proof.
rewrite Rmult_1_l; rewrite sin_plus.
rewrite sin_PI.
rewrite Rmult_0_r.
- rewrite <- (Rplus_0_l (2 * INR x2 * PI)).
+ rewrite <- (Rplus_0_l ((1 + 1) * INR x2 * PI)).
rewrite sin_period.
rewrite sin_0; ring.
apply le_IZR.
@@ -1769,7 +1591,7 @@ Proof.
rewrite H5.
rewrite mult_INR.
simpl in |- *.
- rewrite <- (Rplus_0_l (2 * INR x2 * PI)).
+ rewrite <- (Rplus_0_l ((1 + 1) * INR x2 * PI)).
rewrite sin_period.
rewrite sin_0; ring.
rewrite H5.
@@ -1779,7 +1601,7 @@ Proof.
rewrite Rmult_1_l; rewrite sin_plus.
rewrite sin_PI.
rewrite Rmult_0_r.
- rewrite <- (Rplus_0_l (2 * INR x2 * PI)).
+ rewrite <- (Rplus_0_l ((1 + 1) * INR x2 * PI)).
rewrite sin_period.
rewrite sin_0; ring.
apply le_IZR.
@@ -1787,8 +1609,7 @@ Proof.
rewrite Rplus_0_r.
rewrite Ropp_Ropp_IZR.
rewrite Rplus_opp_r.
- left; replace 0 with (IZR 0); [ apply IZR_lt | reflexivity ].
- assumption.
+ now apply Rlt_le, IZR_lt.
rewrite <- sin_neg.
rewrite Ropp_mult_distr_l_reverse.
rewrite Ropp_involutive.
@@ -1858,7 +1679,7 @@ Proof.
- right; left; auto.
- left.
clear Hi. subst.
- replace 0 with (IZR 0 * PI) by (simpl; ring). f_equal. f_equal.
+ replace 0 with (IZR 0 * PI) by apply Rmult_0_l. f_equal. f_equal.
apply one_IZR_lt1.
split.
+ apply Rlt_le_trans with 0;