diff options
Diffstat (limited to 'theories/Reals')
64 files changed, 2500 insertions, 2275 deletions
diff --git a/theories/Reals/Alembert.v b/theories/Reals/Alembert.v index a8548eb7..e848e4df 100644 --- a/theories/Reals/Alembert.v +++ b/theories/Reals/Alembert.v @@ -1,6 +1,6 @@ (************************************************************************) (* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *) +(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *) (* \VV/ **************************************************************) (* // * This file is distributed under the terms of the *) (* * GNU Lesser General Public License Version 2.1 *) @@ -35,10 +35,8 @@ Proof. [ intro | apply Rinv_0_lt_compat; prove_sup0 ]. elim (H0 (/ 2) H1); intros. exists (sum_f_R0 An x + 2 * An (S x)). - unfold is_upper_bound; intros; unfold EUn in H3; elim H3; intros. - rewrite H4; assert (H5 := lt_eq_lt_dec x1 x). - elim H5; intros. - elim a; intro. + unfold is_upper_bound; intros; unfold EUn in H3; destruct H3 as (x1,->). + destruct (lt_eq_lt_dec x1 x) as [[| -> ]|]. replace (sum_f_R0 An x) with (sum_f_R0 An x1 + sum_f_R0 (fun i:nat => An (S x1 + i)%nat) (x - S x1)). pattern (sum_f_R0 An x1) at 1; rewrite <- Rplus_0_r; @@ -47,7 +45,7 @@ Proof. apply tech1; intros; apply H. apply Rmult_lt_0_compat; [ prove_sup0 | apply H ]. symmetry ; apply tech2; assumption. - rewrite b; pattern (sum_f_R0 An x) at 1; rewrite <- Rplus_0_r; + pattern (sum_f_R0 An x) at 1; rewrite <- Rplus_0_r; apply Rplus_le_compat_l. left; apply Rmult_lt_0_compat; [ prove_sup0 | apply H ]. replace (sum_f_R0 An x1) with @@ -68,7 +66,7 @@ Proof. pattern 2 at 3; rewrite <- Rmult_1_r; rewrite <- (Rmult_comm 2); apply Rmult_le_compat_l. left; prove_sup0. - left; apply Rplus_lt_reg_r with ((/ 2) ^ S (x1 - S x)). + left; apply Rplus_lt_reg_l with ((/ 2) ^ S (x1 - S x)). replace ((/ 2) ^ S (x1 - S x) + (1 - (/ 2) ^ S (x1 - S x))) with 1; [ idtac | ring ]. rewrite <- (Rplus_comm 1); pattern 1 at 1; rewrite <- Rplus_0_r; @@ -86,8 +84,8 @@ Proof. apply (tech6 (fun i:nat => An (S x + i)%nat) (/ 2)). left; apply Rinv_0_lt_compat; prove_sup0. intro; cut (forall n:nat, (n >= x)%nat -> An (S n) < / 2 * An n). - intro; replace (S x + S i)%nat with (S (S x + i)). - apply H6; unfold ge; apply tech8. + intro H4; replace (S x + S i)%nat with (S (S x + i)). + apply H4; unfold ge; apply tech8. apply INR_eq; rewrite S_INR; do 2 rewrite plus_INR; do 2 rewrite S_INR; ring. intros; unfold R_dist in H2; apply Rmult_lt_reg_l with (/ An n). apply Rinv_0_lt_compat; apply H. @@ -101,17 +99,17 @@ Proof. unfold Rdiv; reflexivity. left; unfold Rdiv; change (0 < An (S n) * / An n); apply Rmult_lt_0_compat; [ apply H | apply Rinv_0_lt_compat; apply H ]. - red; intro; assert (H8 := H n); rewrite H7 in H8; + intro H5; assert (H8 := H n); rewrite H5 in H8; elim (Rlt_irrefl _ H8). replace (S x + 0)%nat with (S x); [ reflexivity | ring ]. symmetry ; apply tech2; assumption. exists (sum_f_R0 An 0); unfold EUn; exists 0%nat; reflexivity. - intro X; elim X; intros. + intros (x,H1). exists x; apply Un_cv_crit_lub; [ unfold Un_growing; intro; rewrite tech5; pattern (sum_f_R0 An n) at 1; rewrite <- Rplus_0_r; apply Rplus_le_compat_l; left; apply H - | apply p ]. + | apply H1 ]. Defined. Lemma Alembert_C2 : @@ -127,14 +125,12 @@ Proof. intro; cut (forall n:nat, 0 < Wn n). intro; cut (Un_cv (fun n:nat => Rabs (Vn (S n) / Vn n)) 0). intro; cut (Un_cv (fun n:nat => Rabs (Wn (S n) / Wn n)) 0). - intro; assert (H5 := Alembert_C1 Vn H1 H3). - assert (H6 := Alembert_C1 Wn H2 H4). - elim H5; intros. - elim H6; intros. + intro; pose proof (Alembert_C1 Vn H1 H3) as (x,p). + pose proof (Alembert_C1 Wn H2 H4) as (x0,p0). exists (x - x0); unfold Un_cv; unfold Un_cv in p; unfold Un_cv in p0; intros; cut (0 < eps / 2). - intro; elim (p (eps / 2) H8); clear p; intros. - elim (p0 (eps / 2) H8); clear p0; intros. + intro H6; destruct (p (eps / 2) H6) as (x1,H8). clear p. + destruct (p0 (eps / 2) H6) as (x2,H9). clear p0. set (N := max x1 x2). exists N; intros; replace (sum_f_R0 An n) with (sum_f_R0 Vn n - sum_f_R0 Wn n). @@ -146,9 +142,9 @@ Proof. apply Rabs_triang. rewrite Rabs_Ropp; apply Rlt_le_trans with (eps / 2 + eps / 2). apply Rplus_lt_compat. - unfold R_dist in H9; apply H9; unfold ge; apply le_trans with N; + unfold R_dist in H8; apply H8; unfold ge; apply le_trans with N; [ unfold N; apply le_max_l | assumption ]. - unfold R_dist in H10; apply H10; unfold ge; apply le_trans with N; + unfold R_dist in H9; apply H9; unfold ge; apply le_trans with N; [ unfold N; apply le_max_r | assumption ]. right; symmetry ; apply double_var. symmetry ; apply tech11; intro; unfold Vn, Wn; @@ -315,7 +311,7 @@ Proof. intro; unfold Wn; unfold Rdiv; rewrite <- (Rmult_0_r (/ 2)); rewrite <- (Rmult_comm (/ 2)); apply Rmult_lt_compat_l. apply Rinv_0_lt_compat; prove_sup0. - apply Rplus_lt_reg_r with (An n); rewrite Rplus_0_r; unfold Rminus; + apply Rplus_lt_reg_l with (An n); rewrite Rplus_0_r; unfold Rminus; rewrite (Rplus_comm (An n)); rewrite Rplus_assoc; rewrite Rplus_opp_l; rewrite Rplus_0_r; apply Rle_lt_trans with (Rabs (An n)). @@ -325,7 +321,7 @@ Proof. intro; unfold Vn; unfold Rdiv; rewrite <- (Rmult_0_r (/ 2)); rewrite <- (Rmult_comm (/ 2)); apply Rmult_lt_compat_l. apply Rinv_0_lt_compat; prove_sup0. - apply Rplus_lt_reg_r with (- An n); rewrite Rplus_0_r; unfold Rminus; + apply Rplus_lt_reg_l with (- An n); rewrite Rplus_0_r; unfold Rminus; rewrite (Rplus_comm (- An n)); rewrite Rplus_assoc; rewrite Rplus_opp_r; rewrite Rplus_0_r; apply Rle_lt_trans with (Rabs (An n)). @@ -344,9 +340,8 @@ Proof. intros; set (Bn := fun i:nat => An i * x ^ i). cut (forall n:nat, Bn n <> 0). intro; cut (Un_cv (fun n:nat => Rabs (Bn (S n) / Bn n)) 0). - intro; assert (H4 := Alembert_C2 Bn H2 H3). - elim H4; intros. - exists x0; unfold Bn in p; apply tech12; assumption. + intro; destruct (Alembert_C2 Bn H2 H3) as (x0,H4). + exists x0; unfold Bn in H4; apply tech12; assumption. unfold Un_cv; intros; unfold Un_cv in H1; cut (0 < eps / Rabs x). intro; elim (H1 (eps / Rabs x) H4); intros. exists x0; intros; unfold R_dist; unfold Rminus; @@ -400,15 +395,14 @@ Theorem Alembert_C3 : Un_cv (fun n:nat => Rabs (An (S n) / An n)) 0 -> { l:R | Pser An x l }. Proof. - intros; case (total_order_T x 0); intro. - elim s; intro. + intros; destruct (total_order_T x 0) as [[Hlt|Heq]|Hgt]. cut (x <> 0). intro; apply AlembertC3_step1; assumption. - red; intro; rewrite H1 in a; elim (Rlt_irrefl _ a). + red; intro; rewrite H1 in Hlt; elim (Rlt_irrefl _ Hlt). apply AlembertC3_step2; assumption. cut (x <> 0). intro; apply AlembertC3_step1; assumption. - red; intro; rewrite H1 in r; elim (Rlt_irrefl _ r). + red; intro; rewrite H1 in Hgt; elim (Rlt_irrefl _ Hgt). Defined. Lemma Alembert_C4 : @@ -432,9 +426,7 @@ Proof. unfold is_upper_bound; intros; unfold EUn in H6. elim H6; intros. rewrite H7. - assert (H8 := lt_eq_lt_dec x2 x0). - elim H8; intros. - elim a; intro. + destruct (lt_eq_lt_dec x2 x0) as [[| -> ]|]. replace (sum_f_R0 An x0) with (sum_f_R0 An x2 + sum_f_R0 (fun i:nat => An (S x2 + i)%nat) (x0 - S x2)). pattern (sum_f_R0 An x2) at 1; rewrite <- Rplus_0_r. @@ -443,14 +435,14 @@ Proof. apply tech1. intros; apply H. apply Rmult_lt_0_compat. - apply Rinv_0_lt_compat; apply Rplus_lt_reg_r with x; rewrite Rplus_0_r; + apply Rinv_0_lt_compat; apply Rplus_lt_reg_l with x; rewrite Rplus_0_r; replace (x + (1 - x)) with 1; [ elim H3; intros; assumption | ring ]. apply H. symmetry ; apply tech2; assumption. - rewrite b; pattern (sum_f_R0 An x0) at 1; rewrite <- Rplus_0_r; + pattern (sum_f_R0 An x0) at 1; rewrite <- Rplus_0_r; apply Rplus_le_compat_l. left; apply Rmult_lt_0_compat. - apply Rinv_0_lt_compat; apply Rplus_lt_reg_r with x; rewrite Rplus_0_r; + apply Rinv_0_lt_compat; apply Rplus_lt_reg_l with x; rewrite Rplus_0_r; replace (x + (1 - x)) with 1; [ elim H3; intros; assumption | ring ]. apply H. replace (sum_f_R0 An x2) with @@ -466,7 +458,7 @@ Proof. left; apply H. rewrite tech3. unfold Rdiv; apply Rmult_le_reg_l with (1 - x). - apply Rplus_lt_reg_r with x; rewrite Rplus_0_r. + apply Rplus_lt_reg_l with x; rewrite Rplus_0_r. replace (x + (1 - x)) with 1; [ elim H3; intros; assumption | ring ]. do 2 rewrite (Rmult_comm (1 - x)). rewrite Rmult_assoc; rewrite <- Rinv_l_sym. @@ -480,11 +472,11 @@ Proof. elim Hyp; intros; assumption. elim H3; intros; assumption. apply Rminus_eq_contra. - red; intro. - elim H3; intros. + red; intro H10. + elim H3; intros H11 H12. rewrite H10 in H12; elim (Rlt_irrefl _ H12). - red; intro. - elim H3; intros. + red; intro H10. + elim H3; intros H11 H12. rewrite H10 in H12; elim (Rlt_irrefl _ H12). replace (An (S x0)) with (An (S x0 + 0)%nat). apply (tech6 (fun i:nat => An (S x0 + i)%nat) x). @@ -493,7 +485,7 @@ Proof. elim H3; intros; assumption. intro. cut (forall n:nat, (n >= x0)%nat -> An (S n) < x * An n). - intro. + intro H9. replace (S x0 + S i)%nat with (S (S x0 + i)). apply H9. unfold ge. @@ -515,18 +507,18 @@ Proof. apply Rmult_lt_0_compat. apply H. apply Rinv_0_lt_compat; apply H. - red; intro. + red; intro H10. assert (H11 := H n). rewrite H10 in H11; elim (Rlt_irrefl _ H11). replace (S x0 + 0)%nat with (S x0); [ reflexivity | ring ]. symmetry ; apply tech2; assumption. exists (sum_f_R0 An 0); unfold EUn; exists 0%nat; reflexivity. - intro X; elim X; intros. + intros (x,H1). exists x; apply Un_cv_crit_lub; [ unfold Un_growing; intro; rewrite tech5; pattern (sum_f_R0 An n) at 1; rewrite <- Rplus_0_r; apply Rplus_le_compat_l; left; apply H - | apply p ]. + | apply H1]. Qed. Lemma Alembert_C5 : @@ -586,14 +578,13 @@ Lemma Alembert_C6 : elim X; intros. exists x0. apply tech12; assumption. - case (total_order_T x 0); intro. - elim s; intro. + destruct (total_order_T x 0) as [[Hlt|Heq]|Hgt]. eapply Alembert_C5 with (k * Rabs x). split. unfold Rdiv; apply Rmult_le_pos. left; assumption. left; apply Rabs_pos_lt. - red; intro; rewrite H3 in a; elim (Rlt_irrefl _ a). + red; intro; rewrite H3 in Hlt; elim (Rlt_irrefl _ Hlt). apply Rmult_lt_reg_l with (/ k). apply Rinv_0_lt_compat; assumption. rewrite <- Rmult_assoc. @@ -604,7 +595,7 @@ Lemma Alembert_C6 : intro; apply prod_neq_R0. apply H0. apply pow_nonzero. - red; intro; rewrite H3 in a; elim (Rlt_irrefl _ a). + red; intro; rewrite H3 in Hlt; elim (Rlt_irrefl _ Hlt). unfold Un_cv; unfold Un_cv in H1. intros. cut (0 < eps / Rabs x). @@ -621,7 +612,7 @@ Lemma Alembert_C6 : rewrite Rabs_Rabsolu. apply Rmult_lt_reg_l with (/ Rabs x). apply Rinv_0_lt_compat; apply Rabs_pos_lt. - red; intro; rewrite H7 in a; elim (Rlt_irrefl _ a). + red; intro; rewrite H7 in Hlt; elim (Rlt_irrefl _ Hlt). rewrite <- Rmult_assoc. rewrite <- Rinv_l_sym. rewrite Rmult_1_l. @@ -629,7 +620,7 @@ Lemma Alembert_C6 : unfold R_dist in H5. unfold Rdiv; unfold Rdiv in H5; apply H5; assumption. apply Rabs_no_R0. - red; intro; rewrite H7 in a; elim (Rlt_irrefl _ a). + red; intro; rewrite H7 in Hlt; elim (Rlt_irrefl _ Hlt). unfold Rdiv; replace (S n) with (n + 1)%nat; [ idtac | ring ]. rewrite pow_add. simpl. @@ -641,14 +632,14 @@ Lemma Alembert_C6 : rewrite <- Rinv_r_sym. rewrite Rmult_1_r; reflexivity. apply pow_nonzero. - red; intro; rewrite H7 in a; elim (Rlt_irrefl _ a). + red; intro; rewrite H7 in Hlt; elim (Rlt_irrefl _ Hlt). apply H0. apply pow_nonzero. - red; intro; rewrite H7 in a; elim (Rlt_irrefl _ a). + red; intro; rewrite H7 in Hlt; elim (Rlt_irrefl _ Hlt). unfold Rdiv; apply Rmult_lt_0_compat. assumption. apply Rinv_0_lt_compat; apply Rabs_pos_lt. - red; intro H7; rewrite H7 in a; elim (Rlt_irrefl _ a). + red; intro H7; rewrite H7 in Hlt; elim (Rlt_irrefl _ Hlt). exists (An 0%nat). unfold Un_cv. intros. @@ -661,14 +652,14 @@ Lemma Alembert_C6 : simpl; ring. rewrite tech5. rewrite <- Hrecn. - rewrite b; simpl; ring. + rewrite Heq; simpl; ring. unfold ge; apply le_O_n. eapply Alembert_C5 with (k * Rabs x). split. unfold Rdiv; apply Rmult_le_pos. left; assumption. left; apply Rabs_pos_lt. - red; intro; rewrite H3 in r; elim (Rlt_irrefl _ r). + red; intro; rewrite H3 in Hgt; elim (Rlt_irrefl _ Hgt). apply Rmult_lt_reg_l with (/ k). apply Rinv_0_lt_compat; assumption. rewrite <- Rmult_assoc. @@ -679,7 +670,7 @@ Lemma Alembert_C6 : intro; apply prod_neq_R0. apply H0. apply pow_nonzero. - red; intro; rewrite H3 in r; elim (Rlt_irrefl _ r). + red; intro; rewrite H3 in Hgt; elim (Rlt_irrefl _ Hgt). unfold Un_cv; unfold Un_cv in H1. intros. cut (0 < eps / Rabs x). @@ -696,7 +687,7 @@ Lemma Alembert_C6 : rewrite Rabs_Rabsolu. apply Rmult_lt_reg_l with (/ Rabs x). apply Rinv_0_lt_compat; apply Rabs_pos_lt. - red; intro; rewrite H7 in r; elim (Rlt_irrefl _ r). + red; intro; rewrite H7 in Hgt; elim (Rlt_irrefl _ Hgt). rewrite <- Rmult_assoc. rewrite <- Rinv_l_sym. rewrite Rmult_1_l. @@ -704,7 +695,7 @@ Lemma Alembert_C6 : unfold R_dist in H5. unfold Rdiv; unfold Rdiv in H5; apply H5; assumption. apply Rabs_no_R0. - red; intro; rewrite H7 in r; elim (Rlt_irrefl _ r). + red; intro; rewrite H7 in Hgt; elim (Rlt_irrefl _ Hgt). unfold Rdiv; replace (S n) with (n + 1)%nat; [ idtac | ring ]. rewrite pow_add. simpl. @@ -716,12 +707,12 @@ Lemma Alembert_C6 : rewrite <- Rinv_r_sym. rewrite Rmult_1_r; reflexivity. apply pow_nonzero. - red; intro; rewrite H7 in r; elim (Rlt_irrefl _ r). + red; intro; rewrite H7 in Hgt; elim (Rlt_irrefl _ Hgt). apply H0. apply pow_nonzero. - red; intro; rewrite H7 in r; elim (Rlt_irrefl _ r). + red; intro; rewrite H7 in Hgt; elim (Rlt_irrefl _ Hgt). unfold Rdiv; apply Rmult_lt_0_compat. assumption. apply Rinv_0_lt_compat; apply Rabs_pos_lt. - red; intro H7; rewrite H7 in r; elim (Rlt_irrefl _ r). + red; intro H7; rewrite H7 in Hgt; elim (Rlt_irrefl _ Hgt). Qed. diff --git a/theories/Reals/AltSeries.v b/theories/Reals/AltSeries.v index 6d54b791..3e99c989 100644 --- a/theories/Reals/AltSeries.v +++ b/theories/Reals/AltSeries.v @@ -1,6 +1,6 @@ (************************************************************************) (* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *) +(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *) (* \VV/ **************************************************************) (* // * This file is distributed under the terms of the *) (* * GNU Lesser General Public License Version 2.1 *) @@ -156,8 +156,7 @@ Proof. intros. assert (H2 := CV_ALT_step0 _ H). assert (H3 := CV_ALT_step4 _ H H0). - assert (X := growing_cv _ H2 H3). - elim X; intros. + destruct (growing_cv _ H2 H3) as (x,p). exists x. unfold Un_cv; unfold R_dist; unfold Un_cv in H1; unfold R_dist in H1; unfold Un_cv in p; unfold R_dist in p. @@ -388,16 +387,13 @@ Proof. apply Rle_ge; apply PI_tg_pos. apply lt_le_trans with N; assumption. elim H1; intros H5 _. - assert (H6 := lt_eq_lt_dec 0 N). - elim H6; intro. - elim a; intro. + destruct (lt_eq_lt_dec 0 N) as [[| <- ]|H6]. assumption. - rewrite <- b in H4. rewrite H4 in H5. simpl in H5. cut (0 < / (2 * eps)); [ intro | apply Rinv_0_lt_compat; assumption ]. - elim (Rlt_irrefl _ (Rlt_trans _ _ _ H7 H5)). - elim (lt_n_O _ b). + elim (Rlt_irrefl _ (Rlt_trans _ _ _ H6 H5)). + elim (lt_n_O _ H6). apply le_IZR. simpl. left; apply Rlt_trans with (/ (2 * eps)). @@ -442,7 +438,7 @@ Proof. unfold Rdiv in H; apply Rlt_le_trans with (sum_f_R0 (tg_alt PI_tg) (S (2 * 0))). simpl; unfold tg_alt; simpl; rewrite Rmult_1_l; - rewrite Rmult_1_r; apply Rplus_lt_reg_r with (PI_tg 1). + rewrite Rmult_1_r; apply Rplus_lt_reg_l with (PI_tg 1). rewrite Rplus_0_r; replace (PI_tg 1 + (PI_tg 0 + -1 * PI_tg 1)) with (PI_tg 0); [ unfold PI_tg | ring ]. diff --git a/theories/Reals/ArithProp.v b/theories/Reals/ArithProp.v index cfc74fc4..c4e410ed 100644 --- a/theories/Reals/ArithProp.v +++ b/theories/Reals/ArithProp.v @@ -1,6 +1,6 @@ (************************************************************************) (* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *) +(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *) (* \VV/ **************************************************************) (* // * This file is distributed under the terms of the *) (* * GNU Lesser General Public License Version 2.1 *) @@ -105,14 +105,14 @@ Proof. exists (x - IZR k0 * y). split. ring. - unfold k0; case (Rcase_abs y); intro. + unfold k0; case (Rcase_abs y) as [Hlt|Hge]. assert (H0 := archimed (x / - y)); rewrite <- Z_R_minus; simpl; unfold Rminus. replace (- ((1 + - IZR (up (x / - y))) * y)) with ((IZR (up (x / - y)) - 1) * y); [ idtac | ring ]. split. apply Rmult_le_reg_l with (/ - y). - apply Rinv_0_lt_compat; apply Ropp_0_gt_lt_contravar; exact r. + apply Rinv_0_lt_compat; apply Ropp_0_gt_lt_contravar; exact Hlt. rewrite Rmult_0_r; rewrite (Rmult_comm (/ - y)); rewrite Rmult_plus_distr_r; rewrite <- Ropp_inv_permute; [ idtac | assumption ]. rewrite Rmult_assoc; repeat rewrite Ropp_mult_distr_r_reverse; @@ -125,14 +125,14 @@ Proof. (- (x * / y) + - (IZR (up (x * / - y)) - 1))) with 1; [ idtac | ring ]. elim H0; intros _ H1; unfold Rdiv in H1; exact H1. - rewrite (Rabs_left _ r); apply Rmult_lt_reg_l with (/ - y). - apply Rinv_0_lt_compat; apply Ropp_0_gt_lt_contravar; exact r. + rewrite (Rabs_left _ Hlt); apply Rmult_lt_reg_l with (/ - y). + apply Rinv_0_lt_compat; apply Ropp_0_gt_lt_contravar; exact Hlt. rewrite <- Rinv_l_sym. rewrite (Rmult_comm (/ - y)); rewrite Rmult_plus_distr_r; rewrite <- Ropp_inv_permute; [ idtac | assumption ]. rewrite Rmult_assoc; repeat rewrite Ropp_mult_distr_r_reverse; rewrite <- Rinv_r_sym; [ rewrite Rmult_1_r | assumption ]; - apply Rplus_lt_reg_r with (IZR (up (x / - y)) - 1). + apply Rplus_lt_reg_l with (IZR (up (x / - y)) - 1). replace (IZR (up (x / - y)) - 1 + 1) with (IZR (up (x / - y))); [ idtac | ring ]. replace (IZR (up (x / - y)) - 1 + (- (x * / y) + - (IZR (up (x / - y)) - 1))) @@ -157,22 +157,21 @@ Proof. (IZR (up (x * / y)) - x * / y + (x * / y + (1 - IZR (up (x * / y))))) with 1; [ idtac | ring ]; elim H0; intros _ H2; unfold Rdiv in H2; exact H2. - rewrite (Rabs_right _ r); apply Rmult_lt_reg_l with (/ y). + rewrite (Rabs_right _ Hge); apply Rmult_lt_reg_l with (/ y). apply Rinv_0_lt_compat; assumption. rewrite <- (Rinv_l_sym _ H); rewrite (Rmult_comm (/ y)); rewrite Rmult_plus_distr_r; rewrite Rmult_assoc; rewrite <- Rinv_r_sym; [ rewrite Rmult_1_r | assumption ]; - apply Rplus_lt_reg_r with (IZR (up (x / y)) - 1); + apply Rplus_lt_reg_l with (IZR (up (x / y)) - 1); replace (IZR (up (x / y)) - 1 + 1) with (IZR (up (x / y))); [ idtac | ring ]; replace (IZR (up (x / y)) - 1 + (x * / y + (1 - IZR (up (x / y))))) with (x * / y); [ idtac | ring ]; elim H0; unfold Rdiv; intros H2 _; exact H2. - case (total_order_T 0 y); intro. - elim s; intro. + destruct (total_order_T 0 y) as [[Hlt|Heq]|Hgt]. assumption. - elim H; symmetry ; exact b. - assert (H1 := Rge_le _ _ r); elim (Rlt_irrefl _ (Rle_lt_trans _ _ _ H1 r0)). + elim H; symmetry ; exact Heq. + apply Rge_le in Hge; elim (Rlt_irrefl _ (Rle_lt_trans _ _ _ Hge Hgt)). Qed. Lemma tech8 : forall n i:nat, (n <= S n + i)%nat. diff --git a/theories/Reals/Binomial.v b/theories/Reals/Binomial.v index 3d6121b7..d48f42fc 100644 --- a/theories/Reals/Binomial.v +++ b/theories/Reals/Binomial.v @@ -1,6 +1,6 @@ (************************************************************************) (* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *) +(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *) (* \VV/ **************************************************************) (* // * This file is distributed under the terms of the *) (* * GNU Lesser General Public License Version 2.1 *) @@ -172,13 +172,12 @@ Proof. apply sum_eq. intros; apply H1. unfold N; apply le_lt_trans with n; [ assumption | apply lt_n_Sn ]. - intros; unfold Bn, Cn. - replace (S N - S i)%nat with (N - i)%nat; reflexivity. + reflexivity. unfold An; fold N; rewrite <- minus_n_n; rewrite H0; simpl; ring. apply sum_eq. - intros; unfold An, Bn; replace (S N - S i)%nat with (N - i)%nat; - [ idtac | reflexivity ]. + intros; unfold An, Bn. + change (S N - S i)%nat with (N - i)%nat. rewrite <- pascal; [ ring | apply le_lt_trans with n; [ assumption | unfold N; apply lt_n_Sn ] ]. diff --git a/theories/Reals/Cauchy_prod.v b/theories/Reals/Cauchy_prod.v index 34567cae..28de1186 100644 --- a/theories/Reals/Cauchy_prod.v +++ b/theories/Reals/Cauchy_prod.v @@ -1,6 +1,6 @@ (************************************************************************) (* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *) +(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *) (* \VV/ **************************************************************) (* // * This file is distributed under the terms of the *) (* * GNU Lesser General Public License Version 2.1 *) diff --git a/theories/Reals/Cos_plus.v b/theories/Reals/Cos_plus.v index 71e8d024..49ba9a6e 100644 --- a/theories/Reals/Cos_plus.v +++ b/theories/Reals/Cos_plus.v @@ -1,6 +1,6 @@ (************************************************************************) (* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *) +(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *) (* \VV/ **************************************************************) (* // * This file is distributed under the terms of the *) (* * GNU Lesser General Public License Version 2.1 *) @@ -12,6 +12,7 @@ Require Import SeqSeries. Require Import Rtrigo_def. Require Import Cos_rel. Require Import Max. +Require Import Omega. Local Open Scope nat_scope. Local Open Scope R_scope. diff --git a/theories/Reals/Cos_rel.v b/theories/Reals/Cos_rel.v index 63ab24fe..f5b34de9 100644 --- a/theories/Reals/Cos_rel.v +++ b/theories/Reals/Cos_rel.v @@ -1,6 +1,6 @@ (************************************************************************) (* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *) +(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *) (* \VV/ **************************************************************) (* // * This file is distributed under the terms of the *) (* * GNU Lesser General Public License Version 2.1 *) @@ -10,6 +10,7 @@ Require Import Rbase. Require Import Rfunctions. Require Import SeqSeries. Require Import Rtrigo_def. +Require Import Omega. Local Open Scope R_scope. Definition A1 (x:R) (N:nat) : R := @@ -257,49 +258,30 @@ Qed. Lemma A1_cvg : forall x:R, Un_cv (A1 x) (cos x). intro. -assert (H := exist_cos (x * x)). -elim H; intros. -assert (p_i := p). -unfold cos_in in p. -unfold cos_n, infinite_sum in p. -unfold R_dist in p. -cut (cos x = x0). -intro. -rewrite H0. -unfold Un_cv; unfold R_dist; intros. -elim (p eps H1); intros. +unfold cos; destruct (exist_cos (Rsqr x)) as (x0,p). +unfold cos_in, cos_n, infinite_sum, R_dist in p. +unfold Un_cv, R_dist; intros. +destruct (p eps H) as (x1,H0). exists x1; intros. unfold A1. replace (sum_f_R0 (fun k:nat => (-1) ^ k / INR (fact (2 * k)) * x ^ (2 * k)) n) with (sum_f_R0 (fun i:nat => (-1) ^ i / INR (fact (2 * i)) * (x * x) ^ i) n). -apply H2; assumption. +apply H0; assumption. apply sum_eq. intros. replace ((x * x) ^ i) with (x ^ (2 * i)). reflexivity. apply pow_sqr. -unfold cos. -case (exist_cos (Rsqr x)). -unfold Rsqr; intros. -unfold cos_in in p_i. -unfold cos_in in c. -apply uniqueness_sum with (fun i:nat => cos_n i * (x * x) ^ i); assumption. Qed. Lemma C1_cvg : forall x y:R, Un_cv (C1 x y) (cos (x + y)). intros. -assert (H := exist_cos ((x + y) * (x + y))). -elim H; intros. -assert (p_i := p). -unfold cos_in in p. -unfold cos_n, infinite_sum in p. -unfold R_dist in p. -cut (cos (x + y) = x0). -intro. -rewrite H0. -unfold Un_cv; unfold R_dist; intros. -elim (p eps H1); intros. +unfold cos. +destruct (exist_cos (Rsqr (x + y))) as (x0,p). +unfold cos_in, cos_n, infinite_sum, R_dist in p. +unfold Un_cv, R_dist; intros. +destruct (p eps H) as (x1,H0). exists x1; intros. unfold C1. replace @@ -307,19 +289,12 @@ replace with (sum_f_R0 (fun i:nat => (-1) ^ i / INR (fact (2 * i)) * ((x + y) * (x + y)) ^ i) n). -apply H2; assumption. +apply H0; assumption. apply sum_eq. intros. replace (((x + y) * (x + y)) ^ i) with ((x + y) ^ (2 * i)). reflexivity. apply pow_sqr. -unfold cos. -case (exist_cos (Rsqr (x + y))). -unfold Rsqr; intros. -unfold cos_in in p_i. -unfold cos_in in c. -apply uniqueness_sum with (fun i:nat => cos_n i * ((x + y) * (x + y)) ^ i); - assumption. Qed. Lemma B1_cvg : forall x:R, Un_cv (B1 x) (sin x). @@ -338,21 +313,14 @@ simpl; ring. rewrite tech5; rewrite <- Hrecn. simpl; ring. unfold ge; apply le_O_n. -assert (H0 := exist_sin (x * x)). -elim H0; intros. -assert (p_i := p). -unfold sin_in in p. -unfold sin_n, infinite_sum in p. -unfold R_dist in p. -cut (sin x = x * x0). -intro. -rewrite H1. -unfold Un_cv; unfold R_dist; intros. +unfold sin. destruct (exist_sin (Rsqr x)) as (x0,p). +unfold sin_in, sin_n, infinite_sum, R_dist in p. +unfold Un_cv, R_dist; intros. cut (0 < eps / Rabs x); [ intro | unfold Rdiv; apply Rmult_lt_0_compat; [ assumption | apply Rinv_0_lt_compat; apply Rabs_pos_lt; assumption ] ]. -elim (p (eps / Rabs x) H3); intros. +destruct (p (eps / Rabs x) H1) as (x1,H2). exists x1; intros. unfold B1. replace @@ -370,9 +338,7 @@ replace rewrite Rabs_mult. apply Rmult_lt_reg_l with (/ Rabs x). apply Rinv_0_lt_compat; apply Rabs_pos_lt; assumption. -rewrite <- Rmult_assoc. -rewrite <- Rinv_l_sym. -rewrite Rmult_1_l; rewrite <- (Rmult_comm eps); unfold Rdiv in H4; apply H4; +rewrite <- Rmult_assoc, <- Rinv_l_sym, Rmult_1_l, <- (Rmult_comm eps). apply H2; assumption. apply Rabs_no_R0; assumption. rewrite scal_sum. @@ -382,12 +348,4 @@ rewrite pow_add. rewrite pow_sqr. simpl. ring. -unfold sin. -case (exist_sin (Rsqr x)). -unfold Rsqr; intros. -unfold sin_in in p_i. -unfold sin_in in s. -assert - (H1 := uniqueness_sum (fun i:nat => sin_n i * (x * x) ^ i) x0 x1 p_i s). -rewrite H1; reflexivity. Qed. diff --git a/theories/Reals/DiscrR.v b/theories/Reals/DiscrR.v index 3a2d51f9..75fd4c0a 100644 --- a/theories/Reals/DiscrR.v +++ b/theories/Reals/DiscrR.v @@ -1,6 +1,6 @@ (************************************************************************) (* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *) +(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *) (* \VV/ **************************************************************) (* // * This file is distributed under the terms of the *) (* * GNU Lesser General Public License Version 2.1 *) @@ -11,16 +11,19 @@ Require Import Omega. Local Open Scope R_scope. Lemma Rlt_R0_R2 : 0 < 2. +Proof. change 2 with (INR 2); apply lt_INR_0; apply lt_O_Sn. Qed. Notation Rplus_lt_pos := Rplus_lt_0_compat (only parsing). Lemma IZR_eq : forall z1 z2:Z, z1 = z2 -> IZR z1 = IZR z2. +Proof. intros; rewrite H; reflexivity. Qed. Lemma IZR_neq : forall z1 z2:Z, z1 <> z2 -> IZR z1 <> IZR z2. +Proof. intros; red; intro; elim H; apply eq_IZR; assumption. Qed. diff --git a/theories/Reals/Exp_prop.v b/theories/Reals/Exp_prop.v index 0d418bc3..be96b94e 100644 --- a/theories/Reals/Exp_prop.v +++ b/theories/Reals/Exp_prop.v @@ -1,6 +1,6 @@ (************************************************************************) (* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *) +(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *) (* \VV/ **************************************************************) (* // * This file is distributed under the terms of the *) (* * GNU Lesser General Public License Version 2.1 *) @@ -15,6 +15,7 @@ Require Import PSeries_reg. Require Import Div2. Require Import Even. Require Import Max. +Require Import Omega. Local Open Scope nat_scope. Local Open Scope R_scope. @@ -85,18 +86,17 @@ Qed. Lemma div2_not_R0 : forall N:nat, (1 < N)%nat -> (0 < div2 N)%nat. Proof. - intros; induction N as [| N HrecN]. - elim (lt_n_O _ H). - cut ((1 < N)%nat \/ N = 1%nat). - intro; elim H0; intro. - assert (H2 := even_odd_dec N). - elim H2; intro. - rewrite <- (even_div2 _ a); apply HrecN; assumption. - rewrite <- (odd_div2 _ b); apply lt_O_Sn. - rewrite H1; simpl; apply lt_O_Sn. - inversion H. - right; reflexivity. - left; apply lt_le_trans with 2%nat; [ apply lt_n_Sn | apply H1 ]. + intros; induction N as [| N HrecN]. + - elim (lt_n_O _ H). + - cut ((1 < N)%nat \/ N = 1%nat). + { intro; elim H0; intro. + + destruct (even_odd_dec N) as [Heq|Heq]. + * rewrite <- (even_div2 _ Heq); apply HrecN; assumption. + * rewrite <- (odd_div2 _ Heq); apply lt_O_Sn. + + rewrite H1; simpl; apply lt_O_Sn. } + inversion H. + right; reflexivity. + left; apply lt_le_trans with 2%nat; [ apply lt_n_Sn | apply H1 ]. Qed. Lemma Reste_E_maj : @@ -173,8 +173,7 @@ Proof. apply pow_le; apply Rabs_pos. rewrite (Rmult_comm (/ INR (fact (S n0)))); apply Rmult_le_compat_l. apply pow_le; apply Rabs_pos. - apply Rle_Rinv. - apply INR_fact_lt_0. + apply Rinv_le_contravar. apply INR_fact_lt_0. apply le_INR; apply fact_le; apply le_n_S. apply le_plus_l. @@ -254,8 +253,7 @@ Proof. do 2 rewrite <- (Rmult_comm (/ INR (fact (N - n0)))). apply Rmult_le_compat_l. left; apply Rinv_0_lt_compat; apply INR_fact_lt_0. - apply Rle_Rinv. - apply INR_fact_lt_0. + apply Rinv_le_contravar. apply INR_fact_lt_0. apply le_INR. apply fact_le. @@ -724,15 +722,14 @@ Qed. (**********) Lemma exp_pos : forall x:R, 0 < exp x. Proof. - intro; case (total_order_T 0 x); intro. - elim s; intro. - apply (exp_pos_pos _ a). - rewrite <- b; rewrite exp_0; apply Rlt_0_1. + intro; destruct (total_order_T 0 x) as [[Hlt|<-]|Hgt]. + apply (exp_pos_pos _ Hlt). + rewrite exp_0; apply Rlt_0_1. replace (exp x) with (1 / exp (- x)). unfold Rdiv; apply Rmult_lt_0_compat. apply Rlt_0_1. apply Rinv_0_lt_compat; apply exp_pos_pos. - apply (Ropp_0_gt_lt_contravar _ r). + apply (Ropp_0_gt_lt_contravar _ Hgt). cut (exp (- x) <> 0). intro; unfold Rdiv; apply Rmult_eq_reg_l with (exp (- x)). rewrite Rmult_1_l; rewrite <- Rinv_r_sym. @@ -773,10 +770,10 @@ Proof. apply (not_eq_sym H6). rewrite Rminus_0_r; apply H7. unfold SFL. - case (cv 0); intros. + case (cv 0) as (x,Hu). eapply UL_sequence. - apply u. - unfold Un_cv, SP. + apply Hu. + unfold Un_cv, SP in |- *. intros; exists 1%nat; intros. unfold R_dist; rewrite decomp_sum. rewrite (Rplus_comm (fn 0%nat 0)). @@ -793,14 +790,13 @@ Proof. unfold Rdiv; rewrite Rinv_1; rewrite Rmult_1_r; reflexivity. apply lt_le_trans with 1%nat; [ apply lt_n_Sn | apply H9 ]. unfold SFL, exp. - case (cv h); case (exist_exp h); simpl; intros. + case (cv h) as (x0,Hu); case (exist_exp h) as (x,Hexp); simpl. eapply UL_sequence. - apply u. + apply Hu. unfold Un_cv; intros. - unfold exp_in in e. - unfold infinite_sum in e. + unfold exp_in, infinite_sum in Hexp. cut (0 < eps0 * Rabs h). - intro; elim (e _ H9); intros N0 H10. + intro; elim (Hexp _ H9); intros N0 H10. exists N0; intros. unfold R_dist. apply Rmult_lt_reg_l with (Rabs h). @@ -860,8 +856,7 @@ Proof. Un_cv (fun n:nat => sum_f_R0 (fun k:nat => Rabs (r ^ k / INR (fact (S k)))) n) l }. - intro X. - elim X; intros. + intros (x,p). exists x; intros. split. apply p. diff --git a/theories/Reals/Integration.v b/theories/Reals/Integration.v index 50b57374..222d106f 100644 --- a/theories/Reals/Integration.v +++ b/theories/Reals/Integration.v @@ -1,6 +1,6 @@ (************************************************************************) (* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *) +(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *) (* \VV/ **************************************************************) (* // * This file is distributed under the terms of the *) (* * GNU Lesser General Public License Version 2.1 *) diff --git a/theories/Reals/LegacyRfield.v b/theories/Reals/LegacyRfield.v deleted file mode 100644 index cc8a8f7c..00000000 --- a/theories/Reals/LegacyRfield.v +++ /dev/null @@ -1,38 +0,0 @@ -(************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *) -(* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) -(************************************************************************) - -Require Export Raxioms. -Require Export LegacyField. -Import LegacyRing_theory. - -Section LegacyRfield. - -Open Scope R_scope. - -Lemma RLegacyTheory : Ring_Theory Rplus Rmult 1 0 Ropp (fun x y:R => false). - split. - exact Rplus_comm. - symmetry ; apply Rplus_assoc. - exact Rmult_comm. - symmetry ; apply Rmult_assoc. - intro; apply Rplus_0_l. - intro; apply Rmult_1_l. - exact Rplus_opp_r. - intros. - rewrite Rmult_comm. - rewrite (Rmult_comm n p). - rewrite (Rmult_comm m p). - apply Rmult_plus_distr_l. - intros; contradiction. -Defined. - -End LegacyRfield. - -Add Legacy Field -R Rplus Rmult 1%R 0%R Ropp (fun x y:R => false) Rinv RLegacyTheory Rinv_l - with minus := Rminus div := Rdiv. diff --git a/theories/Reals/MVT.v b/theories/Reals/MVT.v index d3970069..59976957 100644 --- a/theories/Reals/MVT.v +++ b/theories/Reals/MVT.v @@ -1,6 +1,6 @@ (************************************************************************) (* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *) +(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *) (* \VV/ **************************************************************) (* // * This file is distributed under the terms of the *) (* * GNU Lesser General Public License Version 2.1 *) @@ -151,14 +151,14 @@ Proof. cut (forall c:R, a <= c <= b -> continuity_pt id c); [ intro | intros; apply derivable_continuous_pt; apply derivable_id ]. assert (H2 := MVT f id a b X X0 H H0 H1). - elim H2; intros c H3; elim H3; intros. + destruct H2 as (c & P & H4). exists c; split. - cut (derive_pt id c (X0 c x) = derive_pt id c (derivable_pt_id c)); - [ intro | apply pr_nu ]. + cut (derive_pt id c (X0 c P) = derive_pt id c (derivable_pt_id c)); + [ intro H5 | apply pr_nu ]. rewrite H5 in H4; rewrite (derive_pt_id c) in H4; rewrite Rmult_1_r in H4; - rewrite <- H4; replace (derive_pt f c (X c x)) with (derive_pt f c (pr c)); + rewrite <- H4; replace (derive_pt f c (X c P)) with (derive_pt f c (pr c)); [ idtac | apply pr_nu ]; apply Rmult_comm. - apply x. + apply P. Qed. Theorem MVT_cor2 : @@ -173,14 +173,14 @@ Proof. intro; cut (forall c:R, a <= c <= b -> derivable_pt id c). intro X1; cut (forall c:R, a < c < b -> derivable_pt id c). intro X2; cut (forall c:R, a <= c <= b -> continuity_pt id c). - intro; elim (MVT f id a b X0 X2 H H1 H2); intros; elim H3; clear H3; intros; - exists x; split. - cut (derive_pt id x (X2 x x0) = 1). - cut (derive_pt f x (X0 x x0) = f' x). + intro; elim (MVT f id a b X0 X2 H H1 H2); intros x (P,H3). + exists x; split. + cut (derive_pt id x (X2 x P) = 1). + cut (derive_pt f x (X0 x P) = f' x). intros; rewrite H4 in H3; rewrite H5 in H3; unfold id in H3; rewrite Rmult_1_r in H3; rewrite Rmult_comm; symmetry ; assumption. - apply derive_pt_eq_0; apply H0; elim x0; intros; split; left; assumption. + apply derive_pt_eq_0; apply H0; elim P; intros; split; left; assumption. apply derive_pt_eq_0; apply derivable_pt_lim_id. assumption. intros; apply derivable_continuous_pt; apply X1; assumption. @@ -217,12 +217,12 @@ Proof. assert (H3 := MVT f id a b pr H2 H0 H); assert (H4 : forall x:R, a <= x <= b -> continuity_pt id x). intros; apply derivable_continuous; apply derivable_id. - elim (H3 H4); intros; elim H5; intros; exists x; exists x0; rewrite H1 in H6; - unfold id in H6; unfold Rminus in H6; rewrite Rplus_opp_r in H6; - rewrite Rmult_0_l in H6; apply Rmult_eq_reg_l with (b - a); - [ rewrite Rmult_0_r; apply H6 - | apply Rminus_eq_contra; red; intro; rewrite H7 in H0; - elim (Rlt_irrefl _ H0) ]. + destruct (H3 H4) as (c & P & H6). exists c; exists P; rewrite H1 in H6. + unfold id in H6; unfold Rminus in H6; rewrite Rplus_opp_r in H6. + rewrite Rmult_0_l in H6; apply Rmult_eq_reg_l with (b - a); + [ rewrite Rmult_0_r; apply H6 + | apply Rminus_eq_contra; red; intro H7; rewrite H7 in H0; + elim (Rlt_irrefl _ H0) ]. Qed. (**********) @@ -233,21 +233,18 @@ Proof. intros. unfold increasing. intros. - case (total_order_T x y); intro. - elim s; intro. + destruct (total_order_T x y) as [[H1| ->]|H1]. apply Rplus_le_reg_l with (- f x). rewrite Rplus_opp_l; rewrite Rplus_comm. - assert (H1 := MVT_cor1 f _ _ pr a). - elim H1; intros. - elim H2; intros. + pose proof (MVT_cor1 f _ _ pr H1) as (c & H3 & H4). unfold Rminus in H3. rewrite H3. apply Rmult_le_pos. apply H. apply Rplus_le_reg_l with x. rewrite Rplus_0_r; replace (x + (y + - x)) with y; [ assumption | ring ]. - rewrite b; right; reflexivity. - elim (Rlt_irrefl _ (Rle_lt_trans _ _ _ H0 r)). + right; reflexivity. + elim (Rlt_irrefl _ (Rle_lt_trans _ _ _ H0 H1)). Qed. (**********) @@ -269,7 +266,7 @@ Proof. cut ((f (x + delta / 2) - f x) / (delta / 2) <= 0). intro; cut (0 < - ((f (x + delta / 2) - f x) / (delta / 2) - l)). intro; unfold Rabs; - case (Rcase_abs ((f (x + delta / 2) - f x) / (delta / 2) - l)). + case (Rcase_abs ((f (x + delta / 2) - f x) / (delta / 2) - l)) as [Hlt|Hge]. intros; generalize (Rplus_lt_compat_r (- l) (- ((f (x + delta / 2) - f x) / (delta / 2) - l)) @@ -294,7 +291,7 @@ Proof. ring. intros. generalize - (Ropp_ge_le_contravar ((f (x + delta / 2) - f x) / (delta / 2) - l) 0 r). + (Ropp_ge_le_contravar ((f (x + delta / 2) - f x) / (delta / 2) - l) _ Hge). rewrite Ropp_0. intro. elim @@ -412,7 +409,7 @@ Proof. intros. unfold strict_increasing. intros. - apply Rplus_lt_reg_r with (- f x). + apply Rplus_lt_reg_l with (- f x). rewrite Rplus_opp_l; rewrite Rplus_comm. assert (H1 := MVT_cor1 f _ _ pr H0). elim H1; intros. @@ -421,7 +418,7 @@ Proof. rewrite H3. apply Rmult_lt_0_compat. apply H. - apply Rplus_lt_reg_r with x. + apply Rplus_lt_reg_l with x. rewrite Rplus_0_r; replace (x + (y + - x)) with y; [ assumption | ring ]. Qed. @@ -517,7 +514,7 @@ Lemma derive_increasing_interv_ax : Proof. intros. split; intros. - apply Rplus_lt_reg_r with (- f x). + apply Rplus_lt_reg_l with (- f x). rewrite Rplus_opp_l; rewrite Rplus_comm. assert (H4 := MVT_cor1 f _ _ pr H3). elim H4; intros. @@ -532,7 +529,7 @@ Proof. apply Rle_lt_trans with x; assumption. elim H2; intros. apply Rlt_le_trans with y; assumption. - apply Rplus_lt_reg_r with x. + apply Rplus_lt_reg_l with x. rewrite Rplus_0_r; replace (x + (y + - x)) with y; [ assumption | ring ]. apply Rplus_le_reg_l with (- f x). rewrite Rplus_opp_l; rewrite Rplus_comm. @@ -587,12 +584,8 @@ Theorem IAF : f b - f a <= k * (b - a). Proof. intros. - case (total_order_T a b); intro. - elim s; intro. - assert (H1 := MVT_cor1 f _ _ pr a0). - elim H1; intros. - elim H2; intros. - rewrite H3. + destruct (total_order_T a b) as [[H1| -> ]|H1]. + pose proof (MVT_cor1 f _ _ pr H1) as (c & -> & H4). do 2 rewrite <- (Rmult_comm (b - a)). apply Rmult_le_compat_l. apply Rplus_le_reg_l with a; rewrite Rplus_0_r. @@ -600,10 +593,9 @@ Proof. apply H0. elim H4; intros. split; left; assumption. - rewrite b0. unfold Rminus; do 2 rewrite Rplus_opp_r. rewrite Rmult_0_r; right; reflexivity. - elim (Rlt_irrefl _ (Rle_lt_trans _ _ _ H r)). + elim (Rlt_irrefl _ (Rle_lt_trans _ _ _ H H1)). Qed. Lemma IAF_var : @@ -648,8 +640,7 @@ Lemma null_derivative_loc : (forall (x:R) (P:a < x < b), derive_pt f x (pr x P) = 0) -> constant_D_eq f (fun x:R => a <= x <= b) (f a). Proof. - intros; unfold constant_D_eq; intros; case (total_order_T a b); intro. - elim s; intro. + intros; unfold constant_D_eq; intros; destruct (total_order_T a b) as [[Hlt|Heq]|Hgt]. assert (H2 : forall y:R, a < y < x -> derivable_pt id y). intros; apply derivable_pt_id. assert (H3 : forall y:R, a <= y <= x -> continuity_pt id y). @@ -664,24 +655,25 @@ Proof. elim H1; intros; apply Rle_trans with x; assumption. elim H1; clear H1; intros; elim H1; clear H1; intro. assert (H7 := MVT f id a x H4 H2 H1 H5 H3). - elim H7; intros; elim H8; intros; assert (H10 : a < x0 < b). - elim x1; intros; split. - assumption. - apply Rlt_le_trans with x; assumption. - assert (H11 : derive_pt f x0 (H4 x0 x1) = 0). - replace (derive_pt f x0 (H4 x0 x1)) with (derive_pt f x0 (pr x0 H10)); + destruct H7 as (c & P & H9). + assert (H10 : a < c < b). + split. + apply P. + apply Rlt_le_trans with x; [apply P|assumption]. + assert (H11 : derive_pt f c (H4 c P) = 0). + replace (derive_pt f c (H4 c P)) with (derive_pt f c (pr c H10)); [ apply H0 | apply pr_nu ]. - assert (H12 : derive_pt id x0 (H2 x0 x1) = 1). + assert (H12 : derive_pt id c (H2 c P) = 1). apply derive_pt_eq_0; apply derivable_pt_lim_id. rewrite H11 in H9; rewrite H12 in H9; rewrite Rmult_0_r in H9; rewrite Rmult_1_r in H9; apply Rminus_diag_uniq; symmetry ; assumption. rewrite H1; reflexivity. assert (H2 : x = a). - rewrite <- b0 in H1; elim H1; intros; apply Rle_antisym; assumption. + rewrite <- Heq in H1; elim H1; intros; apply Rle_antisym; assumption. rewrite H2; reflexivity. elim H1; intros; - elim (Rlt_irrefl _ (Rle_lt_trans _ _ _ (Rle_trans _ _ _ H2 H3) r)). + elim (Rlt_irrefl _ (Rle_lt_trans _ _ _ (Rle_trans _ _ _ H2 H3) Hgt)). Qed. (* Unicity of the antiderivative *) @@ -718,3 +710,32 @@ Proof. unfold constant_D_eq in H8; assert (H9 := H8 _ H2); unfold minus_fct in H9; rewrite <- H9; ring. Qed. + +(* A variant of MVT using absolute values. *) +Lemma MVT_abs : + forall (f f' : R -> R) (a b : R), + (forall c : R, Rmin a b <= c <= Rmax a b -> + derivable_pt_lim f c (f' c)) -> + exists c : R, Rabs (f b - f a) = Rabs (f' c) * Rabs (b - a) /\ + Rmin a b <= c <= Rmax a b. +Proof. +intros f f' a b. +destruct (Rle_dec a b) as [aleb | blta]. + destruct (Req_dec a b) as [ab | anb]. + unfold Rminus; intros _; exists a; split. + now rewrite <- ab, !Rplus_opp_r, Rabs_R0, Rmult_0_r. + split;[apply Rmin_l | apply Rmax_l]. + rewrite Rmax_right, Rmin_left; auto; intros derv. + destruct (MVT_cor2 f f' a b) as [c [hc intc]]; + [destruct aleb;[assumption | contradiction] | apply derv | ]. + exists c; rewrite hc, Rabs_mult;split; + [reflexivity | unfold Rle; tauto]. +assert (b < a) by (apply Rnot_le_gt; assumption). +assert (b <= a) by (apply Rlt_le; assumption). +rewrite Rmax_left, Rmin_right; try assumption; intros derv. +destruct (MVT_cor2 f f' b a) as [c [hc intc]]; + [assumption | apply derv | ]. +exists c; rewrite <- Rabs_Ropp, Ropp_minus_distr, hc, Rabs_mult. +split;[now rewrite <- (Rabs_Ropp (b - a)), Ropp_minus_distr| unfold Rle; tauto]. +Qed. + diff --git a/theories/Reals/Machin.v b/theories/Reals/Machin.v index 40a857e3..1a94f6a8 100644 --- a/theories/Reals/Machin.v +++ b/theories/Reals/Machin.v @@ -1,6 +1,6 @@ (************************************************************************) (* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *) +(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *) (* \VV/ **************************************************************) (* // * This file is distributed under the terms of the *) (* * GNU Lesser General Public License Version 2.1 *) @@ -16,6 +16,7 @@ Require Import Rseries. Require Import SeqProp. Require Import PartSum. Require Import Ratan. +Require Import Omega. Local Open Scope R_scope. @@ -27,6 +28,7 @@ Lemma atan_sub_correct : forall u v, 1 + u * v <> 0 -> -PI/2 < atan u - atan v < PI/2 -> -PI/2 < atan (atan_sub u v) < PI/2 -> atan u = atan v + atan (atan_sub u v). +Proof. intros u v pn0 uvint aint. assert (cos (atan u) <> 0). destruct (atan_bound u); apply Rgt_not_eq, cos_gt_0; auto. @@ -44,6 +46,7 @@ Qed. Lemma tech : forall x y , -1 <= x <= 1 -> -1 < y < 1 -> -PI/2 < atan x - atan y < PI/2. +Proof. assert (ut := PI_RGT_0). intros x y [xm1 x1] [ym1 y1]. assert (-(PI/4) <= atan x). @@ -67,6 +70,7 @@ Qed. (* A simple formula, reasonably efficient. *) Lemma Machin_2_3 : PI/4 = atan(/2) + atan(/3). +Proof. assert (utility : 0 < PI/2) by (apply PI2_RGT_0). rewrite <- atan_1. rewrite (atan_sub_correct 1 (/2)). @@ -77,6 +81,7 @@ apply atan_bound. Qed. Lemma Machin_4_5_239 : PI/4 = 4 * atan (/5) - atan(/239). +Proof. rewrite <- atan_1. rewrite (atan_sub_correct 1 (/5)); [ | apply Rgt_not_eq; fourier | apply tech; try split; fourier | @@ -105,6 +110,7 @@ unfold atan_sub; field. Qed. Lemma Machin_2_3_7 : PI/4 = 2 * atan(/3) + (atan (/7)). +Proof. rewrite <- atan_1. rewrite (atan_sub_correct 1 (/3)); [ | apply Rgt_not_eq; fourier | apply tech; try split; fourier | diff --git a/theories/Reals/NewtonInt.v b/theories/Reals/NewtonInt.v index 8faf3b41..832e7adc 100644 --- a/theories/Reals/NewtonInt.v +++ b/theories/Reals/NewtonInt.v @@ -1,6 +1,6 @@ (************************************************************************) (* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *) +(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *) (* \VV/ **************************************************************) (* // * This file is distributed under the terms of the *) (* * GNU Lesser General Public License Version 2.1 *) @@ -63,14 +63,16 @@ Proof. [ apply derivable_pt_lim_const | apply derivable_pt_lim_id ] | unfold id, fct_cte; rewrite H2; ring ]. right; reflexivity. -Defined. +Qed. (* $\int_a^a f = 0$ *) Lemma NewtonInt_P2 : forall (f:R -> R) (a:R), NewtonInt f a a (NewtonInt_P1 f a) = 0. Proof. intros; unfold NewtonInt; simpl; - unfold mult_fct, fct_cte, id; ring. + unfold mult_fct, fct_cte, id. + destruct NewtonInt_P1 as [g _]. + now apply Rminus_diag_eq. Qed. (* If $\int_a^b f$ exists, then $\int_b^a f$ exists too *) @@ -87,42 +89,7 @@ Lemma NewtonInt_P4 : forall (f:R -> R) (a b:R) (pr:Newton_integrable f a b), NewtonInt f a b pr = - NewtonInt f b a (NewtonInt_P3 f a b pr). Proof. - intros; unfold Newton_integrable in pr; elim pr; intros; elim p; intro. - unfold NewtonInt; - case - (NewtonInt_P3 f a b - (exist - (fun g:R -> R => antiderivative f g a b \/ antiderivative f g b a) x - p)). - intros; elim o; intro. - unfold antiderivative in H0; elim H0; intros; elim H2; intro. - unfold antiderivative in H; elim H; intros; - elim (Rlt_irrefl _ (Rle_lt_trans _ _ _ H5 H3)). - rewrite H3; ring. - assert (H1 := antiderivative_Ucte f x x0 a b H H0); elim H1; intros; - unfold antiderivative in H0; elim H0; clear H0; intros _ H0. - assert (H3 : a <= a <= b). - split; [ right; reflexivity | assumption ]. - assert (H4 : a <= b <= b). - split; [ assumption | right; reflexivity ]. - assert (H5 := H2 _ H3); assert (H6 := H2 _ H4); rewrite H5; rewrite H6; ring. - unfold NewtonInt; - case - (NewtonInt_P3 f a b - (exist - (fun g:R -> R => antiderivative f g a b \/ antiderivative f g b a) x - p)); intros; elim o; intro. - assert (H1 := antiderivative_Ucte f x x0 b a H H0); elim H1; intros; - unfold antiderivative in H0; elim H0; clear H0; intros _ H0. - assert (H3 : b <= a <= a). - split; [ assumption | right; reflexivity ]. - assert (H4 : b <= b <= a). - split; [ right; reflexivity | assumption ]. - assert (H5 := H2 _ H3); assert (H6 := H2 _ H4); rewrite H5; rewrite H6; ring. - unfold antiderivative in H0; elim H0; intros; elim H2; intro. - unfold antiderivative in H; elim H; intros; - elim (Rlt_irrefl _ (Rle_lt_trans _ _ _ H5 H3)). - rewrite H3; ring. + intros f a b (x,H). unfold NewtonInt, NewtonInt_P3; simpl; ring. Qed. (* The set of Newton integrable functions is a vectorial space *) @@ -133,7 +100,7 @@ Lemma NewtonInt_P5 : Newton_integrable (fun x:R => l * f x + g x) a b. Proof. unfold Newton_integrable; intros f g l a b X X0; - elim X; intros; elim X0; intros; + elim X; intros x p; elim X0; intros x0 p0; exists (fun y:R => l * x y + x0 y). elim p; intro. elim p0; intro. @@ -227,10 +194,8 @@ Lemma NewtonInt_P6 : l * NewtonInt f a b pr1 + NewtonInt g a b pr2. Proof. intros f g l a b pr1 pr2; unfold NewtonInt; - case (NewtonInt_P5 f g l a b pr1 pr2); intros; case pr1; - intros; case pr2; intros; case (total_order_T a b); - intro. - elim s; intro. + destruct (NewtonInt_P5 f g l a b pr1 pr2) as (x,o); destruct pr1 as (x0,o0); + destruct pr2 as (x1,o1); destruct (total_order_T a b) as [[Hlt|Heq]|Hgt]. elim o; intro. elim o0; intro. elim o1; intro. @@ -242,21 +207,21 @@ Proof. split; [ left; assumption | right; reflexivity ]. assert (H7 := H4 _ H5); assert (H8 := H4 _ H6); rewrite H7; rewrite H8; ring. unfold antiderivative in H1; elim H1; intros; - elim (Rlt_irrefl _ (Rle_lt_trans _ _ _ H3 a0)). + elim (Rlt_irrefl _ (Rle_lt_trans _ _ _ H3 Hlt)). unfold antiderivative in H0; elim H0; intros; - elim (Rlt_irrefl _ (Rle_lt_trans _ _ _ H2 a0)). + elim (Rlt_irrefl _ (Rle_lt_trans _ _ _ H2 Hlt)). unfold antiderivative in H; elim H; intros; - elim (Rlt_irrefl _ (Rle_lt_trans _ _ _ H1 a0)). - rewrite b0; ring. + elim (Rlt_irrefl _ (Rle_lt_trans _ _ _ H1 Hlt)). + rewrite Heq; ring. elim o; intro. unfold antiderivative in H; elim H; intros; - elim (Rlt_irrefl _ (Rle_lt_trans _ _ _ H1 r)). + elim (Rlt_irrefl _ (Rle_lt_trans _ _ _ H1 Hgt)). elim o0; intro. unfold antiderivative in H0; elim H0; intros; - elim (Rlt_irrefl _ (Rle_lt_trans _ _ _ H2 r)). + elim (Rlt_irrefl _ (Rle_lt_trans _ _ _ H2 Hgt)). elim o1; intro. unfold antiderivative in H1; elim H1; intros; - elim (Rlt_irrefl _ (Rle_lt_trans _ _ _ H3 r)). + elim (Rlt_irrefl _ (Rle_lt_trans _ _ _ H3 Hgt)). assert (H2 := antiderivative_P1 f g x0 x1 l b a H0 H1); assert (H3 := antiderivative_Ucte _ _ _ _ _ H H2); elim H3; intros; assert (H5 : b <= a <= a). @@ -277,14 +242,12 @@ Lemma antiderivative_P2 : | right _ => F1 x + (F0 b - F1 b) end) a c. Proof. - unfold antiderivative; intros; elim H; clear H; intros; elim H0; - clear H0; intros; split. + intros; destruct H as (H,H1), H0 as (H0,H2); split. 2: apply Rle_trans with b; assumption. - intros; elim H3; clear H3; intros; case (total_order_T x b); intro. - elim s; intro. + intros x (H3,H4); destruct (total_order_T x b) as [[Hlt|Heq]|Hgt]. assert (H5 : a <= x <= b). split; [ assumption | left; assumption ]. - assert (H6 := H _ H5); elim H6; clear H6; intros; + destruct (H _ H5) as (x0,H6). assert (H7 : derivable_pt_lim @@ -293,27 +256,26 @@ Proof. | left _ => F0 x | right _ => F1 x + (F0 b - F1 b) end) x (f x)). - unfold derivable_pt_lim; assert (H7 : derive_pt F0 x x0 = f x). - symmetry ; assumption. - assert (H8 := derive_pt_eq_1 F0 x (f x) x0 H7); unfold derivable_pt_lim in H8; - intros; elim (H8 _ H9); intros; set (D := Rmin x1 (b - x)). + unfold derivable_pt_lim. intros eps H9. + assert (H7 : derive_pt F0 x x0 = f x) by (symmetry; assumption). + destruct (derive_pt_eq_1 F0 x (f x) x0 H7 _ H9) as (x1,H10); set (D := Rmin x1 (b - x)). assert (H11 : 0 < D). - unfold D; unfold Rmin; case (Rle_dec x1 (b - x)); intro. + unfold D, Rmin; case (Rle_dec x1 (b - x)); intro. apply (cond_pos x1). apply Rlt_Rminus; assumption. - exists (mkposreal _ H11); intros; case (Rle_dec x b); intro. - case (Rle_dec (x + h) b); intro. + exists (mkposreal _ H11); intros h H12 H13. case (Rle_dec x b) as [|[]]. + case (Rle_dec (x + h) b) as [|[]]. apply H10. assumption. apply Rlt_le_trans with D; [ assumption | unfold D; apply Rmin_l ]. - elim n; left; apply Rlt_le_trans with (x + D). + left; apply Rlt_le_trans with (x + D). apply Rplus_lt_compat_l; apply Rle_lt_trans with (Rabs h). apply RRle_abs. apply H13. apply Rplus_le_reg_l with (- x); rewrite <- Rplus_assoc; rewrite Rplus_opp_l; rewrite Rplus_0_l; rewrite Rplus_comm; unfold D; apply Rmin_r. - elim n; left; assumption. + left; assumption. assert (H8 : derivable_pt @@ -348,7 +310,7 @@ Proof. unfold D; unfold Rmin; case (Rle_dec x2 x3); intro. apply (cond_pos x2). apply (cond_pos x3). - exists (mkposreal _ H16); intros; case (Rle_dec x b); intro. + exists (mkposreal _ H16); intros; case (Rle_dec x b) as [|[]]. case (Rle_dec (x + h) b); intro. apply H15. assumption. @@ -357,8 +319,8 @@ Proof. apply H14. assumption. apply Rlt_le_trans with D; [ assumption | unfold D; apply Rmin_l ]. - rewrite b0; ring. - elim n; right; assumption. + rewrite Heq; ring. + right; assumption. assert (H14 : derivable_pt @@ -388,12 +350,12 @@ Proof. unfold D; unfold Rmin; case (Rle_dec x1 (x - b)); intro. apply (cond_pos x1). apply Rlt_Rminus; assumption. - exists (mkposreal _ H11); intros; case (Rle_dec x b); intro. - elim (Rlt_irrefl _ (Rle_lt_trans _ _ _ r0 r)). - case (Rle_dec (x + h) b); intro. + exists (mkposreal _ H11); intros; destruct (Rle_dec x b) as [Hle|Hnle]. + elim (Rlt_irrefl _ (Rle_lt_trans _ _ _ Hle Hgt)). + destruct (Rle_dec (x + h) b) as [Hle'|Hnle']. cut (b < x + h). - intro; elim (Rlt_irrefl _ (Rle_lt_trans _ _ _ r0 H14)). - apply Rplus_lt_reg_r with (- h - b); replace (- h - b + b) with (- h); + intro; elim (Rlt_irrefl _ (Rle_lt_trans _ _ _ Hle' H14)). + apply Rplus_lt_reg_l with (- h - b); replace (- h - b + b) with (- h); [ idtac | ring ]; replace (- h - b + (x + h)) with (x - b); [ idtac | ring ]; apply Rle_lt_trans with (Rabs h). rewrite <- Rabs_Ropp; apply RRle_abs. @@ -425,8 +387,7 @@ Lemma antiderivative_P3 : antiderivative f F1 c a \/ antiderivative f F0 a c. Proof. intros; unfold antiderivative in H, H0; elim H; clear H; elim H0; clear H0; - intros; case (total_order_T a c); intro. - elim s; intro. + intros; destruct (total_order_T a c) as [[Hle|Heq]|Hgt]. right; unfold antiderivative; split. intros; apply H1; elim H3; intros; split; [ assumption | apply Rle_trans with c; assumption ]. @@ -448,8 +409,7 @@ Lemma antiderivative_P4 : antiderivative f F1 b c \/ antiderivative f F0 c b. Proof. intros; unfold antiderivative in H, H0; elim H; clear H; elim H0; clear H0; - intros; case (total_order_T c b); intro. - elim s; intro. + intros; destruct (total_order_T c b) as [[Hlt|Heq]|Hgt]. right; unfold antiderivative; split. intros; apply H1; elim H3; intros; split; [ apply Rle_trans with c; assumption | assumption ]. @@ -499,10 +459,8 @@ Proof. intros. elim X; intros F0 H0. elim X0; intros F1 H1. - case (total_order_T a b); intro. - elim s; intro. - case (total_order_T b c); intro. - elim s0; intro. + destruct (total_order_T a b) as [[Hlt|Heq]|Hgt]. + destruct (total_order_T b c) as [[Hlt'|Heq']|Hgt']. (* a<b & b<c *) unfold Newton_integrable; exists @@ -515,84 +473,81 @@ Proof. elim H1; intro. left; apply antiderivative_P2; assumption. unfold antiderivative in H2; elim H2; clear H2; intros _ H2. - elim (Rlt_irrefl _ (Rle_lt_trans _ _ _ H2 a1)). + elim (Rlt_irrefl _ (Rle_lt_trans _ _ _ H2 Hlt')). unfold antiderivative in H; elim H; clear H; intros _ H. - elim (Rlt_irrefl _ (Rle_lt_trans _ _ _ H a0)). + elim (Rlt_irrefl _ (Rle_lt_trans _ _ _ H Hlt)). (* a<b & b=c *) - rewrite b0 in X; apply X. + rewrite Heq' in X; apply X. (* a<b & b>c *) - case (total_order_T a c); intro. - elim s0; intro. + destruct (total_order_T a c) as [[Hlt''|Heq'']|Hgt'']. unfold Newton_integrable; exists F0. left. elim H1; intro. unfold antiderivative in H; elim H; clear H; intros _ H. - elim (Rlt_irrefl _ (Rle_lt_trans _ _ _ H r)). + elim (Rlt_irrefl _ (Rle_lt_trans _ _ _ H Hgt')). elim H0; intro. assert (H3 := antiderivative_P3 f F0 F1 a b c H2 H). elim H3; intro. unfold antiderivative in H4; elim H4; clear H4; intros _ H4. - elim (Rlt_irrefl _ (Rle_lt_trans _ _ _ H4 a1)). + elim (Rlt_irrefl _ (Rle_lt_trans _ _ _ H4 Hlt'')). assumption. unfold antiderivative in H2; elim H2; clear H2; intros _ H2. - elim (Rlt_irrefl _ (Rle_lt_trans _ _ _ H2 a0)). - rewrite b0; apply NewtonInt_P1. + elim (Rlt_irrefl _ (Rle_lt_trans _ _ _ H2 Hlt)). + rewrite Heq''; apply NewtonInt_P1. unfold Newton_integrable; exists F1. right. elim H1; intro. unfold antiderivative in H; elim H; clear H; intros _ H. - elim (Rlt_irrefl _ (Rle_lt_trans _ _ _ H r)). + elim (Rlt_irrefl _ (Rle_lt_trans _ _ _ H Hgt')). elim H0; intro. assert (H3 := antiderivative_P3 f F0 F1 a b c H2 H). elim H3; intro. assumption. unfold antiderivative in H4; elim H4; clear H4; intros _ H4. - elim (Rlt_irrefl _ (Rle_lt_trans _ _ _ H4 r0)). + elim (Rlt_irrefl _ (Rle_lt_trans _ _ _ H4 Hgt'')). unfold antiderivative in H2; elim H2; clear H2; intros _ H2. - elim (Rlt_irrefl _ (Rle_lt_trans _ _ _ H2 a0)). + elim (Rlt_irrefl _ (Rle_lt_trans _ _ _ H2 Hlt)). (* a=b *) - rewrite b0; apply X0. - case (total_order_T b c); intro. - elim s; intro. + rewrite Heq; apply X0. + destruct (total_order_T b c) as [[Hlt'|Heq']|Hgt']. (* a>b & b<c *) - case (total_order_T a c); intro. - elim s0; intro. + destruct (total_order_T a c) as [[Hlt''|Heq'']|Hgt'']. unfold Newton_integrable; exists F1. left. elim H1; intro. (*****************) elim H0; intro. unfold antiderivative in H2; elim H2; clear H2; intros _ H2. - elim (Rlt_irrefl _ (Rle_lt_trans _ _ _ H2 r)). + elim (Rlt_irrefl _ (Rle_lt_trans _ _ _ H2 Hgt)). assert (H3 := antiderivative_P4 f F0 F1 b a c H2 H). elim H3; intro. assumption. unfold antiderivative in H4; elim H4; clear H4; intros _ H4. - elim (Rlt_irrefl _ (Rle_lt_trans _ _ _ H4 a1)). + elim (Rlt_irrefl _ (Rle_lt_trans _ _ _ H4 Hlt'')). unfold antiderivative in H; elim H; clear H; intros _ H. - elim (Rlt_irrefl _ (Rle_lt_trans _ _ _ H a0)). - rewrite b0; apply NewtonInt_P1. + elim (Rlt_irrefl _ (Rle_lt_trans _ _ _ H Hlt')). + rewrite Heq''; apply NewtonInt_P1. unfold Newton_integrable; exists F0. right. elim H0; intro. unfold antiderivative in H; elim H; clear H; intros _ H. - elim (Rlt_irrefl _ (Rle_lt_trans _ _ _ H r)). + elim (Rlt_irrefl _ (Rle_lt_trans _ _ _ H Hgt)). elim H1; intro. assert (H3 := antiderivative_P4 f F0 F1 b a c H H2). elim H3; intro. unfold antiderivative in H4; elim H4; clear H4; intros _ H4. - elim (Rlt_irrefl _ (Rle_lt_trans _ _ _ H4 r0)). + elim (Rlt_irrefl _ (Rle_lt_trans _ _ _ H4 Hgt'')). assumption. unfold antiderivative in H2; elim H2; clear H2; intros _ H2. - elim (Rlt_irrefl _ (Rle_lt_trans _ _ _ H2 a0)). + elim (Rlt_irrefl _ (Rle_lt_trans _ _ _ H2 Hlt')). (* a>b & b=c *) - rewrite b0 in X; apply X. + rewrite Heq' in X; apply X. (* a>b & b>c *) assert (X1 := NewtonInt_P3 f a b X). assert (X2 := NewtonInt_P3 f b c X0). apply NewtonInt_P3. apply NewtonInt_P7 with b; assumption. -Defined. +Qed. (* Chasles' relation *) Lemma NewtonInt_P9 : @@ -602,17 +557,15 @@ Lemma NewtonInt_P9 : NewtonInt f a b pr1 + NewtonInt f b c pr2. Proof. intros; unfold NewtonInt. - case (NewtonInt_P8 f a b c pr1 pr2); intros. - case pr1; intros. - case pr2; intros. - case (total_order_T a b); intro. - elim s; intro. - case (total_order_T b c); intro. - elim s0; intro. + case (NewtonInt_P8 f a b c pr1 pr2) as (x,Hor). + case pr1 as (x0,Hor0). + case pr2 as (x1,Hor1). + destruct (total_order_T a b) as [[Hlt|Heq]|Hgt]. + destruct (total_order_T b c) as [[Hlt'|Heq']|Hgt']. (* a<b & b<c *) - elim o0; intro. - elim o1; intro. - elim o; intro. + case Hor0; intro. + case Hor1; intro. + case Hor; intro. assert (H2 := antiderivative_P2 f x0 x1 a b c H H0). assert (H3 := @@ -628,23 +581,23 @@ Proof. assert (H6 : a <= c <= c). split; [ left; apply Rlt_trans with b; assumption | right; reflexivity ]. rewrite (H4 _ H5); rewrite (H4 _ H6). - case (Rle_dec a b); intro. - case (Rle_dec c b); intro. - elim (Rlt_irrefl _ (Rle_lt_trans _ _ _ r0 a1)). + destruct (Rle_dec a b) as [Hlea|Hnlea]. + destruct (Rle_dec c b) as [Hlec|Hnlec]. + elim (Rlt_irrefl _ (Rle_lt_trans _ _ _ Hlec Hlt')). ring. - elim n; left; assumption. + elim Hnlea; left; assumption. unfold antiderivative in H1; elim H1; clear H1; intros _ H1. - elim (Rlt_irrefl _ (Rle_lt_trans _ _ _ H1 (Rlt_trans _ _ _ a0 a1))). + elim (Rlt_irrefl _ (Rle_lt_trans _ _ _ H1 (Rlt_trans _ _ _ Hlt Hlt'))). unfold antiderivative in H0; elim H0; clear H0; intros _ H0. - elim (Rlt_irrefl _ (Rle_lt_trans _ _ _ H0 a1)). + elim (Rlt_irrefl _ (Rle_lt_trans _ _ _ H0 Hlt')). unfold antiderivative in H; elim H; clear H; intros _ H. - elim (Rlt_irrefl _ (Rle_lt_trans _ _ _ H a0)). + elim (Rlt_irrefl _ (Rle_lt_trans _ _ _ H Hlt)). (* a<b & b=c *) - rewrite <- b0. + rewrite <- Heq'. unfold Rminus; rewrite Rplus_opp_r; rewrite Rplus_0_r. - rewrite <- b0 in o. - elim o0; intro. - elim o; intro. + rewrite <- Heq' in Hor. + elim Hor0; intro. + elim Hor; intro. assert (H1 := antiderivative_Ucte f x x0 a b H0 H). elim H1; intros. rewrite (H2 b). @@ -653,25 +606,25 @@ Proof. split; [ right; reflexivity | left; assumption ]. split; [ left; assumption | right; reflexivity ]. unfold antiderivative in H0; elim H0; clear H0; intros _ H0. - elim (Rlt_irrefl _ (Rle_lt_trans _ _ _ H0 a0)). + elim (Rlt_irrefl _ (Rle_lt_trans _ _ _ H0 Hlt)). unfold antiderivative in H; elim H; clear H; intros _ H. - elim (Rlt_irrefl _ (Rle_lt_trans _ _ _ H a0)). + elim (Rlt_irrefl _ (Rle_lt_trans _ _ _ H Hlt)). (* a<b & b>c *) - elim o1; intro. + elim Hor1; intro. unfold antiderivative in H; elim H; clear H; intros _ H. - elim (Rlt_irrefl _ (Rle_lt_trans _ _ _ H r)). - elim o0; intro. - elim o; intro. + elim (Rlt_irrefl _ (Rle_lt_trans _ _ _ H Hgt')). + elim Hor0; intro. + elim Hor; intro. assert (H2 := antiderivative_P2 f x x1 a c b H1 H). assert (H3 := antiderivative_Ucte _ _ _ a b H0 H2). elim H3; intros. rewrite (H4 a). rewrite (H4 b). - case (Rle_dec b c); intro. - elim (Rlt_irrefl _ (Rle_lt_trans _ _ _ r0 r)). - case (Rle_dec a c); intro. + destruct (Rle_dec b c) as [Hle|Hnle]. + elim (Rlt_irrefl _ (Rle_lt_trans _ _ _ Hle Hgt')). + destruct (Rle_dec a c) as [Hle'|Hnle']. ring. - elim n0; unfold antiderivative in H1; elim H1; intros; assumption. + elim Hnle'; unfold antiderivative in H1; elim H1; intros; assumption. split; [ left; assumption | right; reflexivity ]. split; [ right; reflexivity | left; assumption ]. assert (H2 := antiderivative_P2 _ _ _ _ _ _ H1 H0). @@ -679,19 +632,19 @@ Proof. elim H3; intros. rewrite (H4 c). rewrite (H4 b). - case (Rle_dec b a); intro. - elim (Rlt_irrefl _ (Rle_lt_trans _ _ _ r0 a0)). - case (Rle_dec c a); intro. + destruct (Rle_dec b a) as [Hle|Hnle]. + elim (Rlt_irrefl _ (Rle_lt_trans _ _ _ Hle Hlt)). + destruct (Rle_dec c a) as [Hle'|[]]. ring. - elim n0; unfold antiderivative in H1; elim H1; intros; assumption. + unfold antiderivative in H1; elim H1; intros; assumption. split; [ left; assumption | right; reflexivity ]. split; [ right; reflexivity | left; assumption ]. unfold antiderivative in H0; elim H0; clear H0; intros _ H0. - elim (Rlt_irrefl _ (Rle_lt_trans _ _ _ H0 a0)). + elim (Rlt_irrefl _ (Rle_lt_trans _ _ _ H0 Hlt)). (* a=b *) - rewrite b0 in o; rewrite b0. - elim o; intro. - elim o1; intro. + rewrite Heq in Hor |- *. + elim Hor; intro. + elim Hor1; intro. assert (H1 := antiderivative_Ucte _ _ _ b c H H0). elim H1; intros. assert (H3 : b <= c). @@ -705,7 +658,7 @@ Proof. unfold antiderivative in H, H0; elim H; elim H0; intros; apply Rle_antisym; assumption. rewrite H1; ring. - elim o1; intro. + elim Hor1; intro. assert (H1 : b = c). unfold antiderivative in H, H0; elim H; elim H0; intros; apply Rle_antisym; assumption. @@ -720,25 +673,24 @@ Proof. split; [ assumption | right; reflexivity ]. split; [ right; reflexivity | assumption ]. (* a>b & b<c *) - case (total_order_T b c); intro. - elim s; intro. - elim o0; intro. + destruct (total_order_T b c) as [[Hlt'|Heq']|Hgt']. + elim Hor0; intro. unfold antiderivative in H; elim H; clear H; intros _ H. - elim (Rlt_irrefl _ (Rle_lt_trans _ _ _ H r)). - elim o1; intro. - elim o; intro. + elim (Rlt_irrefl _ (Rle_lt_trans _ _ _ H Hgt)). + elim Hor1; intro. + elim Hor; intro. assert (H2 := antiderivative_P2 _ _ _ _ _ _ H H1). assert (H3 := antiderivative_Ucte _ _ _ b c H0 H2). elim H3; intros. rewrite (H4 b). rewrite (H4 c). - case (Rle_dec b a); intro. - case (Rle_dec c a); intro. + case (Rle_dec b a) as [|[]]. + case (Rle_dec c a) as [|]. assert (H5 : a = c). unfold antiderivative in H1; elim H1; intros; apply Rle_antisym; assumption. rewrite H5; ring. ring. - elim n; left; assumption. + left; assumption. split; [ left; assumption | right; reflexivity ]. split; [ right; reflexivity | left; assumption ]. assert (H2 := antiderivative_P2 _ _ _ _ _ _ H0 H1). @@ -746,27 +698,27 @@ Proof. elim H3; intros. rewrite (H4 a). rewrite (H4 b). - case (Rle_dec b c); intro. - case (Rle_dec a c); intro. + case (Rle_dec b c) as [|[]]. + case (Rle_dec a c) as [|]. assert (H5 : a = c). unfold antiderivative in H1; elim H1; intros; apply Rle_antisym; assumption. rewrite H5; ring. ring. - elim n; left; assumption. + left; assumption. split; [ right; reflexivity | left; assumption ]. split; [ left; assumption | right; reflexivity ]. unfold antiderivative in H0; elim H0; clear H0; intros _ H0. - elim (Rlt_irrefl _ (Rle_lt_trans _ _ _ H0 a0)). + elim (Rlt_irrefl _ (Rle_lt_trans _ _ _ H0 Hlt')). (* a>b & b=c *) - rewrite <- b0. + rewrite <- Heq'. unfold Rminus; rewrite Rplus_opp_r; rewrite Rplus_0_r. - rewrite <- b0 in o. - elim o0; intro. + rewrite <- Heq' in Hor. + elim Hor0; intro. unfold antiderivative in H; elim H; clear H; intros _ H. - elim (Rlt_irrefl _ (Rle_lt_trans _ _ _ H r)). - elim o; intro. + elim (Rlt_irrefl _ (Rle_lt_trans _ _ _ H Hgt)). + elim Hor; intro. unfold antiderivative in H0; elim H0; clear H0; intros _ H0. - elim (Rlt_irrefl _ (Rle_lt_trans _ _ _ H0 r)). + elim (Rlt_irrefl _ (Rle_lt_trans _ _ _ H0 Hgt)). assert (H1 := antiderivative_Ucte f x x0 b a H0 H). elim H1; intros. rewrite (H2 b). @@ -775,15 +727,15 @@ Proof. split; [ left; assumption | right; reflexivity ]. split; [ right; reflexivity | left; assumption ]. (* a>b & b>c *) - elim o0; intro. + elim Hor0; intro. unfold antiderivative in H; elim H; clear H; intros _ H. - elim (Rlt_irrefl _ (Rle_lt_trans _ _ _ H r)). - elim o1; intro. + elim (Rlt_irrefl _ (Rle_lt_trans _ _ _ H Hgt)). + elim Hor1; intro. unfold antiderivative in H0; elim H0; clear H0; intros _ H0. - elim (Rlt_irrefl _ (Rle_lt_trans _ _ _ H0 r0)). - elim o; intro. + elim (Rlt_irrefl _ (Rle_lt_trans _ _ _ H0 Hgt')). + elim Hor; intro. unfold antiderivative in H1; elim H1; clear H1; intros _ H1. - elim (Rlt_irrefl _ (Rle_lt_trans _ _ _ H1 (Rlt_trans _ _ _ r0 r))). + elim (Rlt_irrefl _ (Rle_lt_trans _ _ _ H1 (Rlt_trans _ _ _ Hgt' Hgt))). assert (H2 := antiderivative_P2 _ _ _ _ _ _ H0 H). assert (H3 := antiderivative_Ucte _ _ _ c a H1 H2). elim H3; intros. @@ -791,11 +743,11 @@ Proof. unfold antiderivative in H1; elim H1; intros; assumption. rewrite (H4 c). rewrite (H4 a). - case (Rle_dec a b); intro. - elim (Rlt_irrefl _ (Rle_lt_trans _ _ _ r1 r)). - case (Rle_dec c b); intro. + destruct (Rle_dec a b) as [Hle|Hnle]. + elim (Rlt_irrefl _ (Rle_lt_trans _ _ _ Hle Hgt)). + destruct (Rle_dec c b) as [|[]]. ring. - elim n0; left; assumption. + left; assumption. split; [ assumption | right; reflexivity ]. split; [ right; reflexivity | assumption ]. Qed. diff --git a/theories/Reals/PSeries_reg.v b/theories/Reals/PSeries_reg.v index 199c2014..30a26f77 100644 --- a/theories/Reals/PSeries_reg.v +++ b/theories/Reals/PSeries_reg.v @@ -1,6 +1,6 @@ (************************************************************************) (* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *) +(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *) (* \VV/ **************************************************************) (* // * This file is distributed under the terms of the *) (* * GNU Lesser General Public License Version 2.1 *) @@ -10,12 +10,116 @@ Require Import Rbase. Require Import Rfunctions. Require Import SeqSeries. Require Import Ranalysis1. +Require Import MVT. Require Import Max. Require Import Even. +Require Import Fourier. Local Open Scope R_scope. +(* Boule is French for Ball *) + Definition Boule (x:R) (r:posreal) (y:R) : Prop := Rabs (y - x) < r. +(* General properties of balls. *) + +Lemma Boule_convex : forall c d x y z, + Boule c d x -> Boule c d y -> x <= z <= y -> Boule c d z. +intros c d x y z bx b_y intz. +unfold Boule in bx, b_y; apply Rabs_def2 in bx; +apply Rabs_def2 in b_y; apply Rabs_def1; + [apply Rle_lt_trans with (y - c);[apply Rplus_le_compat_r|]| + apply Rlt_le_trans with (x - c);[|apply Rplus_le_compat_r]];tauto. +Qed. + +Definition boule_of_interval x y (h : x < y) : + {c :R & {r : posreal | c - r = x /\ c + r = y}}. +exists ((x + y)/2). +assert (radius : 0 < (y - x)/2). + unfold Rdiv; apply Rmult_lt_0_compat. + apply Rlt_Rminus; assumption. + now apply Rinv_0_lt_compat, Rlt_0_2. + exists (mkposreal _ radius). + simpl; split; unfold Rdiv; field. +Qed. + +Definition boule_in_interval x y z (h : x < z < y) : + {c : R & {r | Boule c r z /\ x < c - r /\ c + r < y}}. +Proof. +assert (cmp : x * /2 + z * /2 < z * /2 + y * /2). +destruct h as [h1 h2]. +rewrite Rplus_comm; apply Rplus_lt_compat_l, Rmult_lt_compat_r. + apply Rinv_0_lt_compat, Rlt_0_2. +apply Rlt_trans with z; assumption. +destruct (boule_of_interval _ _ cmp) as [c [r [P1 P2]]]. +assert (0 < /2) by (apply Rinv_0_lt_compat, Rlt_0_2). +exists c, r; split. + destruct h; unfold Boule; simpl; apply Rabs_def1. + apply Rplus_lt_reg_l with c; rewrite P2; + replace (c + (z - c)) with (z * / 2 + z * / 2) by field. + apply Rplus_lt_compat_l, Rmult_lt_compat_r;assumption. + apply Rplus_lt_reg_l with c; change (c + - r) with (c - r); + rewrite P1; + replace (c + (z - c)) with (z * / 2 + z * / 2) by field. + apply Rplus_lt_compat_r, Rmult_lt_compat_r;assumption. +destruct h; split. + replace x with (x * / 2 + x * / 2) by field; rewrite P1. + apply Rplus_lt_compat_l, Rmult_lt_compat_r;assumption. +replace y with (y * / 2 + y * /2) by field; rewrite P2. +apply Rplus_lt_compat_r, Rmult_lt_compat_r;assumption. +Qed. + +Lemma Ball_in_inter : forall c1 c2 r1 r2 x, + Boule c1 r1 x -> Boule c2 r2 x -> + {r3 : posreal | forall y, Boule x r3 y -> Boule c1 r1 y /\ Boule c2 r2 y}. +intros c1 c2 [r1 r1p] [r2 r2p] x; unfold Boule; simpl; intros in1 in2. +assert (Rmax (c1 - r1)(c2 - r2) < x). + apply Rmax_lub_lt;[revert in1 | revert in2]; intros h; + apply Rabs_def2 in h; destruct h as [_ u]; + apply (fun h => Rplus_lt_reg_r _ _ _ (Rle_lt_trans _ _ _ h u)), Req_le; ring. +assert (x < Rmin (c1 + r1) (c2 + r2)). + apply Rmin_glb_lt;[revert in1 | revert in2]; intros h; + apply Rabs_def2 in h; destruct h as [u _]; + apply (fun h => Rplus_lt_reg_r _ _ _ (Rlt_le_trans _ _ _ u h)), Req_le; ring. +assert (t: 0 < Rmin (x - Rmax (c1 - r1) (c2 - r2)) + (Rmin (c1 + r1) (c2 + r2) - x)). + apply Rmin_glb_lt; apply Rlt_Rminus; assumption. +exists (mkposreal _ t). +apply Rabs_def2 in in1; destruct in1. +apply Rabs_def2 in in2; destruct in2. +assert (c1 - r1 <= Rmax (c1 - r1) (c2 - r2)) by apply Rmax_l. +assert (c2 - r2 <= Rmax (c1 - r1) (c2 - r2)) by apply Rmax_r. +assert (Rmin (c1 + r1) (c2 + r2) <= c1 + r1) by apply Rmin_l. +assert (Rmin (c1 + r1) (c2 + r2) <= c2 + r2) by apply Rmin_r. +assert (Rmin (x - Rmax (c1 - r1) (c2 - r2)) + (Rmin (c1 + r1) (c2 + r2) - x) <= x - Rmax (c1 - r1) (c2 - r2)) + by apply Rmin_l. +assert (Rmin (x - Rmax (c1 - r1) (c2 - r2)) + (Rmin (c1 + r1) (c2 + r2) - x) <= Rmin (c1 + r1) (c2 + r2) - x) + by apply Rmin_r. +simpl. +intros y h; apply Rabs_def2 in h; destruct h as [h h']. +apply Rmin_Rgt in h; destruct h as [cmp1 cmp2]. +apply Rplus_lt_reg_r in cmp2; apply Rmin_Rgt in cmp2. +rewrite Ropp_Rmin, Ropp_minus_distr in h'. +apply Rmax_Rlt in h'; destruct h' as [cmp3 cmp4]; +apply Rplus_lt_reg_r in cmp3; apply Rmax_Rlt in cmp3; +split; apply Rabs_def1. +apply (fun h => Rplus_lt_reg_l _ _ _ (Rle_lt_trans _ _ _ h (proj1 cmp2))), Req_le; + ring. +apply (fun h => Rplus_lt_reg_l _ _ _ (Rlt_le_trans _ _ _ (proj1 cmp3) h)), Req_le; + ring. +apply (fun h => Rplus_lt_reg_l _ _ _ (Rle_lt_trans _ _ _ h (proj2 cmp2))), Req_le; + ring. +apply (fun h => Rplus_lt_reg_l _ _ _ (Rlt_le_trans _ _ _ (proj2 cmp3) h)), Req_le; + ring. +Qed. + +Lemma Boule_center : forall x r, Boule x r x. +Proof. +intros x [r rpos]; unfold Boule, Rminus; simpl; rewrite Rplus_opp_r. +rewrite Rabs_pos_eq;[assumption | apply Rle_refl]. +Qed. + (** Uniform convergence *) Definition CVU (fn:nat -> R -> R) (f:R -> R) (x:R) (r:posreal) : Prop := @@ -153,7 +257,7 @@ Proof. unfold Boule; replace (y + h - x) with (h + (y - x)); [ idtac | ring ]; apply Rle_lt_trans with (Rabs h + Rabs (y - x)). apply Rabs_triang. - apply Rplus_lt_reg_r with (- Rabs (x - y)). + apply Rplus_lt_reg_l with (- Rabs (x - y)). rewrite <- (Rabs_Ropp (y - x)); rewrite Ropp_minus_distr'. replace (- Rabs (x - y) + r) with (r - Rabs (x - y)). replace (- Rabs (x - y) + (Rabs h + Rabs (x - y))) with (Rabs h). @@ -161,7 +265,7 @@ Proof. ring. ring. unfold Boule in H1; rewrite <- (Rabs_Ropp (x - y)); rewrite Ropp_minus_distr'; - apply Rplus_lt_reg_r with (Rabs (y - x)). + apply Rplus_lt_reg_l with (Rabs (y - x)). rewrite Rplus_0_r; replace (Rabs (y - x) + (r - Rabs (y - x))) with (pos r); [ apply H1 | ring ]. Qed. @@ -258,3 +362,242 @@ Proof. rewrite <- Rplus_0_r; apply Rplus_lt_compat_l; apply Rlt_0_1. apply Rplus_le_lt_0_compat; [ apply Rabs_pos | apply Rlt_0_1 ]. Qed. + +(* Uniform convergence implies pointwise simple convergence *) +Lemma CVU_cv : forall f g c d, CVU f g c d -> + forall x, Boule c d x -> Un_cv (fun n => f n x) (g x). +intros f g c d cvu x bx eps ep; destruct (cvu eps ep) as [N Pn]. + exists N; intros n nN; rewrite R_dist_sym; apply Pn; assumption. +Qed. + +(* convergence is preserved through extensional equality *) +Lemma CVU_ext_lim : + forall f g1 g2 c d, CVU f g1 c d -> (forall x, Boule c d x -> g1 x = g2 x) -> + CVU f g2 c d. +intros f g1 g2 c d cvu q eps ep; destruct (cvu _ ep) as [N Pn]. +exists N; intros; rewrite <- q; auto. +Qed. + +(* When a sequence of derivable functions converge pointwise towards + a function g, with the derivatives converging uniformly towards + a function g', then the function g' is the derivative of g. *) + +Lemma CVU_derivable : + forall f f' g g' c d, + CVU f' g' c d -> + (forall x, Boule c d x -> Un_cv (fun n => f n x) (g x)) -> + (forall n x, Boule c d x -> derivable_pt_lim (f n) x (f' n x)) -> + forall x, Boule c d x -> derivable_pt_lim g x (g' x). +intros f f' g g' c d cvu cvp dff' x bx. +set (rho_ := + fun n y => + if Req_EM_T y x then + f' n x + else ((f n y - f n x)/ (y - x))). +set (rho := fun y => + if Req_EM_T y x then + g' x + else (g y - g x)/(y - x)). +assert (ctrho : forall n z, Boule c d z -> continuity_pt (rho_ n) z). + intros n z bz. + destruct (Req_EM_T x z) as [xz | xnz]. + rewrite <- xz. + intros eps' ep'. + destruct (dff' n x bx eps' ep') as [alp Pa]. + exists (pos alp);split;[apply cond_pos | ]. + intros z'; unfold rho_, D_x, dist, R_met; simpl; intros [[_ xnz'] dxz']. + destruct (Req_EM_T z' x) as [abs | _]. + case xnz'; symmetry; exact abs. + destruct (Req_EM_T x x) as [_ | abs];[ | case abs; reflexivity]. + pattern z' at 1; replace z' with (x + (z' - x)) by ring. + apply Pa;[intros h; case xnz'; + replace z' with (z' - x + x) by ring; rewrite h, Rplus_0_l; + reflexivity | exact dxz']. + destruct (Ball_in_inter c c d d z bz bz) as [delta Pd]. + assert (dz : 0 < Rmin delta (Rabs (z - x))). + now apply Rmin_glb_lt;[apply cond_pos | apply Rabs_pos_lt; intros zx0; case xnz; + replace z with (z - x + x) by ring; rewrite zx0, Rplus_0_l]. + assert (t' : forall y : R, + R_dist y z < Rmin delta (Rabs (z - x)) -> + (fun z : R => (f n z - f n x) / (z - x)) y = rho_ n y). + intros y dyz; unfold rho_; destruct (Req_EM_T y x) as [xy | xny]. + rewrite xy in dyz. + destruct (Rle_dec delta (Rabs (z - x))). + rewrite Rmin_left, R_dist_sym in dyz; unfold R_dist in dyz; fourier. + rewrite Rmin_right, R_dist_sym in dyz; unfold R_dist in dyz; + [case (Rlt_irrefl _ dyz) |apply Rlt_le, Rnot_le_gt; assumption]. + reflexivity. + apply (continuity_pt_locally_ext (fun z => (f n z - f n x)/(z - x)) + (rho_ n) _ z dz t'); clear t'. + apply continuity_pt_div. + apply continuity_pt_minus. + apply derivable_continuous_pt; eapply exist; apply dff'; assumption. + apply continuity_pt_const; intro; intro; reflexivity. + apply continuity_pt_minus; + [apply derivable_continuous_pt; exists 1; apply derivable_pt_lim_id + | apply continuity_pt_const; intro; reflexivity]. + intros zx0; case xnz; replace z with (z - x + x) by ring. + rewrite zx0, Rplus_0_l; reflexivity. +assert (CVU rho_ rho c d ). + intros eps ep. + assert (ep8 : 0 < eps/8). + fourier. + destruct (cvu _ ep8) as [N Pn1]. + assert (cauchy1 : forall n p, (N <= n)%nat -> (N <= p)%nat -> + forall z, Boule c d z -> Rabs (f' n z - f' p z) < eps/4). + intros n p nN pN z bz; replace (eps/4) with (eps/8 + eps/8) by field. + rewrite <- Rabs_Ropp. + replace (-(f' n z - f' p z)) with (g' z - f' n z - (g' z - f' p z)) by ring. + apply Rle_lt_trans with (1 := Rabs_triang _ _); rewrite Rabs_Ropp. + apply Rplus_lt_compat; apply Pn1; assumption. + assert (step_2 : forall n p, (N <= n)%nat -> (N <= p)%nat -> + forall y, Boule c d y -> x <> y -> + Rabs ((f n y - f n x)/(y - x) - (f p y - f p x)/(y - x)) < eps/4). + intros n p nN pN y b_y xny. + assert (mm0 : (Rmin x y = x /\ Rmax x y = y) \/ + (Rmin x y = y /\ Rmax x y = x)). + destruct (Rle_dec x y) as [H | H]. + rewrite Rmin_left, Rmax_right. + left; split; reflexivity. + assumption. + assumption. + rewrite Rmin_right, Rmax_left. + right; split; reflexivity. + apply Rlt_le, Rnot_le_gt; assumption. + apply Rlt_le, Rnot_le_gt; assumption. + assert (mm : Rmin x y < Rmax x y). + destruct mm0 as [[q1 q2] | [q1 q2]]; generalize (Rminmax x y); rewrite q1, q2. + intros h; destruct h;[ assumption| contradiction]. + intros h; destruct h as [h | h];[assumption | rewrite h in xny; case xny; reflexivity]. + assert (dm : forall z, Rmin x y <= z <= Rmax x y -> + derivable_pt_lim (fun x => f n x - f p x) z (f' n z - f' p z)). + intros z intz; apply derivable_pt_lim_minus. + apply dff'; apply Boule_convex with (Rmin x y) (Rmax x y); + destruct mm0 as [[q1 q2] | [q1 q2]]; revert intz; rewrite ?q1, ?q2; intros; + try assumption. + apply dff'; apply Boule_convex with (Rmin x y) (Rmax x y); + destruct mm0 as [[q1 q2] | [q1 q2]]; revert intz; rewrite ?q1, ?q2; intros; + try assumption. + + replace ((f n y - f n x) / (y - x) - (f p y - f p x) / (y - x)) + with (((f n y - f p y) - (f n x - f p x))/(y - x)) by + (field; intros yx0; case xny; replace y with (y - x + x) by ring; + rewrite yx0, Rplus_0_l; reflexivity). + destruct (MVT_cor2 (fun x => f n x - f p x) (fun x => f' n x - f' p x) + (Rmin x y) (Rmax x y) mm dm) as [z [Pz inz]]. + destruct mm0 as [[q1 q2] | [q1 q2]]. + replace ((f n y - f p y - (f n x - f p x))/(y - x)) with + ((f n (Rmax x y) - f p (Rmax x y) - (f n (Rmin x y) - f p (Rmin x y)))/ + (Rmax x y - Rmin x y)) by (rewrite q1, q2; reflexivity). + unfold Rdiv; rewrite Pz, Rmult_assoc, Rinv_r, Rmult_1_r. + apply cauchy1; auto. + apply Boule_convex with (Rmin x y) (Rmax x y); + revert inz; rewrite ?q1, ?q2; intros; + try assumption. + split; apply Rlt_le; tauto. + rewrite q1, q2; apply Rminus_eq_contra, not_eq_sym; assumption. + replace ((f n y - f p y - (f n x - f p x))/(y - x)) with + ((f n (Rmax x y) - f p (Rmax x y) - (f n (Rmin x y) - f p (Rmin x y)))/ + (Rmax x y - Rmin x y)). + unfold Rdiv; rewrite Pz, Rmult_assoc, Rinv_r, Rmult_1_r. + apply cauchy1; auto. + apply Boule_convex with (Rmin x y) (Rmax x y); + revert inz; rewrite ?q1, ?q2; intros; + try assumption; split; apply Rlt_le; tauto. + rewrite q1, q2; apply Rminus_eq_contra; assumption. + rewrite q1, q2; field; split; + apply Rminus_eq_contra;[apply not_eq_sym |]; assumption. + assert (unif_ac : + forall n p, (N <= n)%nat -> (N <= p)%nat -> + forall y, Boule c d y -> + Rabs (rho_ n y - rho_ p y) <= eps/2). + intros n p nN pN y b_y. + destruct (Req_dec x y) as [xy | xny]. + destruct (Ball_in_inter c c d d x bx bx) as [delta Pdelta]. + destruct (ctrho n y b_y _ ep8) as [d' [dp Pd]]. + destruct (ctrho p y b_y _ ep8) as [d2 [dp2 Pd2]]. + assert (mmpos : 0 < (Rmin (Rmin d' d2) delta)/2). + apply Rmult_lt_0_compat; repeat apply Rmin_glb_lt; try assumption. + apply cond_pos. + apply Rinv_0_lt_compat, Rlt_0_2. + apply Rle_trans with (1 := R_dist_tri _ _ (rho_ n (y + Rmin (Rmin d' d2) delta/2))). + replace (eps/2) with (eps/8 + (eps/4 + eps/8)) by field. + apply Rplus_le_compat. + rewrite R_dist_sym; apply Rlt_le, Pd;split;[split;[exact I | ] | ]. + apply Rminus_not_eq_right; rewrite Rplus_comm; unfold Rminus; + rewrite Rplus_assoc, Rplus_opp_r, Rplus_0_r; apply Rgt_not_eq; assumption. + simpl; unfold R_dist. + unfold Rminus; rewrite (Rplus_comm y), Rplus_assoc, Rplus_opp_r, Rplus_0_r. + rewrite Rabs_pos_eq;[ |apply Rlt_le; assumption ]. + apply Rlt_le_trans with (Rmin (Rmin d' d2) delta);[fourier | ]. + apply Rle_trans with (Rmin d' d2); apply Rmin_l. + apply Rle_trans with (1 := R_dist_tri _ _ (rho_ p (y + Rmin (Rmin d' d2) delta/2))). + apply Rplus_le_compat. + apply Rlt_le. + replace (rho_ n (y + Rmin (Rmin d' d2) delta / 2)) with + ((f n (y + Rmin (Rmin d' d2) delta / 2) - f n x)/ + ((y + Rmin (Rmin d' d2) delta / 2) - x)). + replace (rho_ p (y + Rmin (Rmin d' d2) delta / 2)) with + ((f p (y + Rmin (Rmin d' d2) delta / 2) - f p x)/ + ((y + Rmin (Rmin d' d2) delta / 2) - x)). + apply step_2; auto; try fourier. + assert (0 < pos delta) by (apply cond_pos). + apply Boule_convex with y (y + delta/2). + assumption. + destruct (Pdelta (y + delta/2)); auto. + rewrite xy; unfold Boule; rewrite Rabs_pos_eq; try fourier; auto. + split; try fourier. + apply Rplus_le_compat_l, Rmult_le_compat_r;[ | apply Rmin_r]. + now apply Rlt_le, Rinv_0_lt_compat, Rlt_0_2. + apply Rminus_not_eq_right; rewrite xy; apply Rgt_not_eq; fourier. + unfold rho_. + destruct (Req_EM_T (y + Rmin (Rmin d' d2) delta/2) x) as [ymx | ymnx]. + case (RIneq.Rle_not_lt _ _ (Req_le _ _ ymx)); fourier. + reflexivity. + unfold rho_. + destruct (Req_EM_T (y + Rmin (Rmin d' d2) delta / 2) x) as [ymx | ymnx]. + case (RIneq.Rle_not_lt _ _ (Req_le _ _ ymx)); fourier. + reflexivity. + apply Rlt_le, Pd2; split;[split;[exact I | apply Rlt_not_eq; fourier] | ]. + simpl; unfold R_dist. + unfold Rminus; rewrite (Rplus_comm y), Rplus_assoc, Rplus_opp_r, Rplus_0_r. + rewrite Rabs_pos_eq;[ | fourier]. + apply Rlt_le_trans with (Rmin (Rmin d' d2) delta); [fourier |]. + apply Rle_trans with (Rmin d' d2). + solve[apply Rmin_l]. + solve[apply Rmin_r]. + apply Rlt_le, Rlt_le_trans with (eps/4);[ | fourier]. + unfold rho_; destruct (Req_EM_T y x); solve[auto]. + assert (unif_ac' : forall p, (N <= p)%nat -> + forall y, Boule c d y -> Rabs (rho y - rho_ p y) < eps). + assert (cvrho : forall y, Boule c d y -> Un_cv (fun n => rho_ n y) (rho y)). + intros y b_y; unfold rho_, rho; destruct (Req_EM_T y x). + intros eps' ep'; destruct (cvu eps' ep') as [N2 Pn2]. + exists N2; intros n nN2; rewrite R_dist_sym; apply Pn2; assumption. + apply CV_mult. + apply CV_minus. + apply cvp; assumption. + apply cvp; assumption. + intros eps' ep'; simpl; exists 0%nat; intros; rewrite R_dist_eq; assumption. + intros p pN y b_y. + replace eps with (eps/2 + eps/2) by field. + assert (ep2 : 0 < eps/2) by fourier. + destruct (cvrho y b_y _ ep2) as [N2 Pn2]. + apply Rle_lt_trans with (1 := R_dist_tri _ _ (rho_ (max N N2) y)). + apply Rplus_lt_le_compat. + solve[rewrite R_dist_sym; apply Pn2, Max.le_max_r]. + apply unif_ac; auto; solve [apply Max.le_max_l]. + exists N; intros; apply unif_ac'; solve[auto]. +intros eps ep. +destruct (CVU_continuity _ _ _ _ H ctrho x bx eps ep) as [delta [dp Pd]]. +exists (mkposreal _ dp); intros h hn0 dh. +replace ((g (x + h) - g x) / h) with (rho (x + h)). + replace (g' x) with (rho x). + apply Pd; unfold D_x, no_cond;split;[split;[solve[auto] | ] | ]. + intros xxh; case hn0; replace h with (x + h - x) by ring; rewrite <- xxh; ring. + simpl; unfold R_dist; replace (x + h - x) with h by ring; exact dh. + unfold rho; destruct (Req_EM_T x x) as [ _ | abs];[ | case abs]; reflexivity. +unfold rho; destruct (Req_EM_T (x + h) x) as [abs | _];[ | ]. + case hn0; replace h with (x + h - x) by ring; rewrite abs; ring. +replace (x + h - x) with h by ring; reflexivity. +Qed. diff --git a/theories/Reals/PartSum.v b/theories/Reals/PartSum.v index 364d72cb..b710c75c 100644 --- a/theories/Reals/PartSum.v +++ b/theories/Reals/PartSum.v @@ -1,6 +1,6 @@ (************************************************************************) (* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *) +(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *) (* \VV/ **************************************************************) (* // * This file is distributed under the terms of the *) (* * GNU Lesser General Public License Version 2.1 *) @@ -180,12 +180,9 @@ Proof. replace (S (S (pred N))) with (S N). rewrite (HrecN H1); ring. rewrite H2; simpl; reflexivity. - assert (H2 := O_or_S N). - elim H2; intros. - elim a; intros. - rewrite <- p. + destruct (O_or_S N) as [(m,<-)|<-]. simpl; reflexivity. - rewrite <- b in H1; elim (lt_irrefl _ H1). + elim (lt_irrefl _ H1). rewrite H1; simpl; reflexivity. inversion H. right; reflexivity. @@ -395,9 +392,7 @@ Proof. (sum_f_R0 (fun i:nat => Rabs (An i)) m)). assumption. apply H1; assumption. - assert (H4 := lt_eq_lt_dec n m). - elim H4; intro. - elim a; intro. + destruct (lt_eq_lt_dec n m) as [[ | -> ]|]. rewrite (tech2 An n m); [ idtac | assumption ]. rewrite (tech2 (fun i:nat => Rabs (An i)) n m); [ idtac | assumption ]. unfold R_dist. @@ -418,7 +413,6 @@ Proof. apply Rle_ge. apply cond_pos_sum. intro; apply Rabs_pos. - rewrite b. unfold R_dist. unfold Rminus; do 2 rewrite Rplus_opp_r. rewrite Rabs_R0; right; reflexivity. @@ -451,8 +445,7 @@ Lemma cv_cauchy_1 : { l:R | Un_cv (fun N:nat => sum_f_R0 An N) l } -> Cauchy_crit_series An. Proof. - intros An X. - elim X; intros. + intros An (x,p). unfold Un_cv in p. unfold Cauchy_crit_series; unfold Cauchy_crit. intros. @@ -508,12 +501,11 @@ Lemma sum_incr : Un_cv (fun n:nat => sum_f_R0 An n) l -> (forall n:nat, 0 <= An n) -> sum_f_R0 An N <= l. Proof. - intros; case (total_order_T (sum_f_R0 An N) l); intro. - elim s; intro. - left; apply a. - right; apply b. + intros; destruct (total_order_T (sum_f_R0 An N) l) as [[Hlt|Heq]|Hgt]. + left; apply Hlt. + right; apply Heq. cut (Un_growing (fun n:nat => sum_f_R0 An n)). - intro; set (l1 := sum_f_R0 An N) in r. + intro; set (l1 := sum_f_R0 An N) in Hgt. unfold Un_cv in H; cut (0 < l1 - l). intro; elim (H _ H2); intros. set (N0 := max x N); cut (N0 >= x)%nat. @@ -522,21 +514,21 @@ Proof. intro; unfold R_dist in H5; rewrite Rabs_right in H5. cut (sum_f_R0 An N0 < l1). intro; elim (Rlt_irrefl _ (Rlt_le_trans _ _ _ H7 H6)). - apply Rplus_lt_reg_r with (- l). + apply Rplus_lt_reg_l with (- l). do 2 rewrite (Rplus_comm (- l)). apply H5. apply Rle_ge; apply Rplus_le_reg_l with l. rewrite Rplus_0_r; replace (l + (sum_f_R0 An N0 - l)) with (sum_f_R0 An N0); [ idtac | ring ]; apply Rle_trans with l1. - left; apply r. + left; apply Hgt. apply H6. unfold l1; apply Rge_le; apply (growing_prop (fun k:nat => sum_f_R0 An k)). apply H1. unfold ge, N0; apply le_max_r. unfold ge, N0; apply le_max_l. - apply Rplus_lt_reg_r with l; rewrite Rplus_0_r; - replace (l + (l1 - l)) with l1; [ apply r | ring ]. + apply Rplus_lt_reg_l with l; rewrite Rplus_0_r; + replace (l + (l1 - l)) with l1; [ apply Hgt | ring ]. unfold Un_growing; intro; simpl; pattern (sum_f_R0 An n) at 1; rewrite <- Rplus_0_r; apply Rplus_le_compat_l; apply H0. @@ -549,10 +541,9 @@ Lemma sum_cv_maj : Un_cv (fun n:nat => sum_f_R0 An n) l2 -> (forall n:nat, Rabs (fn n x) <= An n) -> Rabs l1 <= l2. Proof. - intros; case (total_order_T (Rabs l1) l2); intro. - elim s; intro. - left; apply a. - right; apply b. + intros; destruct (total_order_T (Rabs l1) l2) as [[Hlt|Heq]|Hgt]. + left; apply Hlt. + right; apply Heq. cut (forall n0:nat, Rabs (SP fn n0 x) <= sum_f_R0 An n0). intro; cut (0 < (Rabs l1 - l2) / 2). intro; unfold Un_cv in H, H0. @@ -568,17 +559,17 @@ Proof. intro; assert (H11 := H2 N). elim (Rlt_irrefl _ (Rle_lt_trans _ _ _ H11 H10)). apply Rlt_trans with ((Rabs l1 + l2) / 2); assumption. - case (Rcase_abs (Rabs l1 - Rabs (SP fn N x))); intro. + destruct (Rcase_abs (Rabs l1 - Rabs (SP fn N x))) as [Hlt|Hge]. apply Rlt_trans with (Rabs l1). apply Rmult_lt_reg_l with 2. prove_sup0. unfold Rdiv; rewrite (Rmult_comm 2); rewrite Rmult_assoc; rewrite <- Rinv_l_sym. - rewrite Rmult_1_r; rewrite double; apply Rplus_lt_compat_l; apply r. + rewrite Rmult_1_r; rewrite double; apply Rplus_lt_compat_l; apply Hgt. discrR. - apply (Rminus_lt _ _ r0). - rewrite (Rabs_right _ r0) in H7. - apply Rplus_lt_reg_r with ((Rabs l1 - l2) / 2 - Rabs (SP fn N x)). + apply (Rminus_lt _ _ Hlt). + rewrite (Rabs_right _ Hge) in H7. + apply Rplus_lt_reg_l with ((Rabs l1 - l2) / 2 - Rabs (SP fn N x)). replace ((Rabs l1 - l2) / 2 - Rabs (SP fn N x) + (Rabs l1 + l2) / 2) with (Rabs l1 - Rabs (SP fn N x)). unfold Rminus; rewrite Rplus_assoc; rewrite Rplus_opp_l; @@ -586,18 +577,18 @@ Proof. unfold Rdiv; rewrite Rmult_plus_distr_r; rewrite <- (Rmult_comm (/ 2)); rewrite Rmult_minus_distr_l; repeat rewrite (Rmult_comm (/ 2)); pattern (Rabs l1) at 1; - rewrite double_var; unfold Rdiv; ring. - case (Rcase_abs (sum_f_R0 An N - l2)); intro. + rewrite double_var; unfold Rdiv in |- *; ring. + destruct (Rcase_abs (sum_f_R0 An N - l2)) as [Hlt|Hge]. apply Rlt_trans with l2. - apply (Rminus_lt _ _ r0). + apply (Rminus_lt _ _ Hlt). apply Rmult_lt_reg_l with 2. prove_sup0. rewrite (double l2); unfold Rdiv; rewrite (Rmult_comm 2); rewrite Rmult_assoc; rewrite <- Rinv_l_sym. rewrite Rmult_1_r; rewrite (Rplus_comm (Rabs l1)); apply Rplus_lt_compat_l; - apply r. + apply Hgt. discrR. - rewrite (Rabs_right _ r0) in H6; apply Rplus_lt_reg_r with (- l2). + rewrite (Rabs_right _ Hge) in H6; apply Rplus_lt_reg_l with (- l2). replace (- l2 + (Rabs l1 + l2) / 2) with ((Rabs l1 - l2) / 2). rewrite Rplus_comm; apply H6. unfold Rdiv; rewrite <- (Rmult_comm (/ 2)); @@ -610,9 +601,9 @@ Proof. apply H4; unfold ge, N; apply le_max_l. apply H5; unfold ge, N; apply le_max_r. unfold Rdiv; apply Rmult_lt_0_compat. - apply Rplus_lt_reg_r with l2. + apply Rplus_lt_reg_l with l2. rewrite Rplus_0_r; replace (l2 + (Rabs l1 - l2)) with (Rabs l1); - [ apply r | ring ]. + [ apply Hgt | ring ]. apply Rinv_0_lt_compat; prove_sup0. intros; induction n0 as [| n0 Hrecn0]. unfold SP; simpl; apply H1. diff --git a/theories/Reals/RIneq.v b/theories/Reals/RIneq.v index b881250f..8dca0197 100644 --- a/theories/Reals/RIneq.v +++ b/theories/Reals/RIneq.v @@ -1,7 +1,7 @@ (* -*- coding: utf-8 -*- *) (************************************************************************) (* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *) +(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *) (* \VV/ **************************************************************) (* // * This file is distributed under the terms of the *) (* * GNU Lesser General Public License Version 2.1 *) @@ -43,7 +43,7 @@ Hint Immediate Rge_refl: rorders. Lemma Rlt_irrefl : forall r, ~ r < r. Proof. - generalize Rlt_asym. intuition eauto. + intros r H; eapply Rlt_asym; eauto. Qed. Hint Resolve Rlt_irrefl: real. @@ -64,7 +64,9 @@ Qed. (**********) Lemma Rlt_dichotomy_converse : forall r1 r2, r1 < r2 \/ r1 > r2 -> r1 <> r2. Proof. - generalize Rlt_not_eq Rgt_not_eq. intuition eauto. + intuition. + - apply Rlt_not_eq in H1. eauto. + - apply Rgt_not_eq in H1. eauto. Qed. Hint Resolve Rlt_dichotomy_converse: real. @@ -74,7 +76,7 @@ Hint Resolve Rlt_dichotomy_converse: real. Lemma Req_dec : forall r1 r2, r1 = r2 \/ r1 <> r2. Proof. intros; generalize (total_order_T r1 r2) Rlt_dichotomy_converse; - intuition eauto 3. + unfold not; intuition eauto 3. Qed. Hint Resolve Req_dec: real. @@ -175,7 +177,7 @@ Proof. eauto using Rnot_gt_ge with rorders. Qed. Lemma Rlt_not_le : forall r1 r2, r2 < r1 -> ~ r1 <= r2. Proof. generalize Rlt_asym Rlt_dichotomy_converse; unfold Rle. - intuition eauto 3. + unfold not; intuition eauto 3. Qed. Hint Immediate Rlt_not_le: real. @@ -407,11 +409,20 @@ Proof. rewrite Rplus_assoc; rewrite H; ring. Qed. -Hint Resolve (f_equal (A:=R)): real. +Definition f_equal_R := (f_equal (A:=R)). + +Hint Resolve f_equal_R : real. Lemma Rplus_eq_compat_l : forall r r1 r2, r1 = r2 -> r + r1 = r + r2. Proof. - auto with real. + intros r r1 r2. + apply f_equal. +Qed. + +Lemma Rplus_eq_compat_r : forall r r1 r2, r1 = r2 -> r1 + r = r2 + r. +Proof. + intros r r1 r2. + apply (f_equal (fun v => v + r)). Qed. (*i Old i*)Hint Resolve Rplus_eq_compat_l: v62. @@ -427,6 +438,13 @@ Proof. Qed. Hint Resolve Rplus_eq_reg_l: real. +Lemma Rplus_eq_reg_r : forall r r1 r2, r1 + r = r2 + r -> r1 = r2. +Proof. + intros r r1 r2 H. + apply Rplus_eq_reg_l with r. + now rewrite 2!(Rplus_comm r). +Qed. + (**********) Lemma Rplus_0_r_uniq : forall r r1, r + r1 = r -> r1 = 0. Proof. @@ -664,6 +682,11 @@ Hint Resolve Ropp_plus_distr: real. (** ** Opposite and multiplication *) (*********************************************************) +Lemma Ropp_mult_distr_l : forall r1 r2, - (r1 * r2) = - r1 * r2. +Proof. + intros; ring. +Qed. + Lemma Ropp_mult_distr_l_reverse : forall r1 r2, - r1 * r2 = - (r1 * r2). Proof. intros; ring. @@ -677,13 +700,18 @@ Proof. Qed. Hint Resolve Rmult_opp_opp: real. +Lemma Ropp_mult_distr_r : forall r1 r2, - (r1 * r2) = r1 * - r2. +Proof. + intros; ring. +Qed. + Lemma Ropp_mult_distr_r_reverse : forall r1 r2, r1 * - r2 = - (r1 * r2). Proof. intros; ring. Qed. (*********************************************************) -(** ** Substraction *) +(** ** Subtraction *) (*********************************************************) Lemma Rminus_0_r : forall r, r - 0 = r. @@ -794,7 +822,7 @@ Hint Resolve Rinv_involutive: real. Lemma Rinv_mult_distr : forall r1 r2, r1 <> 0 -> r2 <> 0 -> / (r1 * r2) = / r1 * / r2. Proof. - intros; field; auto. + intros; field; auto. Qed. (*********) @@ -969,7 +997,7 @@ Qed. (** *** Cancellation *) -Lemma Rplus_lt_reg_r : forall r r1 r2, r + r1 < r + r2 -> r1 < r2. +Lemma Rplus_lt_reg_l : forall r r1 r2, r + r1 < r + r2 -> r1 < r2. Proof. intros; cut (- r + r + r1 < - r + r + r2). rewrite Rplus_opp_l. @@ -979,10 +1007,17 @@ Proof. apply (Rplus_lt_compat_l (- r) (r + r1) (r + r2) H). Qed. +Lemma Rplus_lt_reg_r : forall r r1 r2, r1 + r < r2 + r -> r1 < r2. +Proof. + intros. + apply (Rplus_lt_reg_l r). + now rewrite 2!(Rplus_comm r). +Qed. + Lemma Rplus_le_reg_l : forall r r1 r2, r + r1 <= r + r2 -> r1 <= r2. Proof. unfold Rle; intros; elim H; intro. - left; apply (Rplus_lt_reg_r r r1 r2 H0). + left; apply (Rplus_lt_reg_l r r1 r2 H0). right; apply (Rplus_eq_reg_l r r1 r2 H0). Qed. @@ -995,7 +1030,7 @@ Qed. Lemma Rplus_gt_reg_l : forall r r1 r2, r + r1 > r + r2 -> r1 > r2. Proof. - unfold Rgt; intros; apply (Rplus_lt_reg_r r r2 r1 H). + unfold Rgt; intros; apply (Rplus_lt_reg_l r r2 r1 H). Qed. Lemma Rplus_ge_reg_l : forall r r1 r2, r + r1 >= r + r2 -> r1 >= r2. @@ -1047,12 +1082,10 @@ Qed. Lemma Ropp_gt_lt_contravar : forall r1 r2, r1 > r2 -> - r1 < - r2. Proof. unfold Rgt; intros. - apply (Rplus_lt_reg_r (r2 + r1)). - replace (r2 + r1 + - r1) with r2. - replace (r2 + r1 + - r2) with r1. - trivial. - ring. - ring. + apply (Rplus_lt_reg_l (r2 + r1)). + replace (r2 + r1 + - r1) with r2 by ring. + replace (r2 + r1 + - r2) with r1 by ring. + exact H. Qed. Hint Resolve Ropp_gt_lt_contravar. @@ -1324,19 +1357,22 @@ Qed. Lemma Rlt_minus : forall r1 r2, r1 < r2 -> r1 - r2 < 0. Proof. - intros; apply (Rplus_lt_reg_r r2). - replace (r2 + (r1 - r2)) with r1. - replace (r2 + 0) with r2; auto with real. - ring. + intros; apply (Rplus_lt_reg_l r2). + replace (r2 + (r1 - r2)) with r1 by ring. + now rewrite Rplus_0_r. Qed. Hint Resolve Rlt_minus: real. Lemma Rgt_minus : forall r1 r2, r1 > r2 -> r1 - r2 > 0. Proof. - intros; apply (Rplus_lt_reg_r r2). - replace (r2 + (r1 - r2)) with r1. - replace (r2 + 0) with r2; auto with real. - ring. + intros; apply (Rplus_lt_reg_l r2). + replace (r2 + (r1 - r2)) with r1 by ring. + now rewrite Rplus_0_r. +Qed. + +Lemma Rlt_Rminus : forall a b:R, a < b -> 0 < b - a. +Proof. + intros a b; apply Rgt_minus. Qed. (**********) @@ -1368,6 +1404,9 @@ Proof. ring. Qed. +Lemma Rminus_gt_0_lt : forall a b, 0 < b - a -> a < b. +Proof. intro; intro; apply Rminus_gt. Qed. + (**********) Lemma Rminus_le : forall r1 r2, r1 - r2 <= 0 -> r1 <= r2. Proof. @@ -1625,7 +1664,7 @@ Proof. apply (Rlt_irrefl 0); auto. do 2 rewrite S_INR in H1; cut (INR n1 < INR n0). intro H2; generalize (H0 n0 H2); intro; auto with arith. - apply (Rplus_lt_reg_r 1 (INR n1) (INR n0)). + apply (Rplus_lt_reg_l 1 (INR n1) (INR n0)). rewrite Rplus_comm; rewrite (Rplus_comm 1 (INR n0)); trivial. Qed. Hint Resolve INR_lt: real. @@ -1931,18 +1970,26 @@ Proof. apply (Rmult_le_compat_l x 0 y H H0). Qed. +Lemma Rinv_le_contravar : + forall x y, 0 < x -> x <= y -> / y <= / x. +Proof. + intros x y H1 [H2|H2]. + apply Rlt_le. + apply Rinv_lt_contravar with (2 := H2). + apply Rmult_lt_0_compat with (1 := H1). + now apply Rlt_trans with x. + rewrite H2. + apply Rle_refl. +Qed. + Lemma Rle_Rinv : forall x y:R, 0 < x -> 0 < y -> x <= y -> / y <= / x. Proof. - intros; apply Rmult_le_reg_l with x. - apply H. - rewrite <- Rinv_r_sym. - apply Rmult_le_reg_l with y. - apply H0. - rewrite Rmult_1_r; rewrite Rmult_comm; rewrite Rmult_assoc; - rewrite <- Rinv_l_sym. - rewrite Rmult_1_r; apply H1. - red; intro; rewrite H2 in H0; elim (Rlt_irrefl _ H0). - red; intro; rewrite H2 in H; elim (Rlt_irrefl _ H). + intros x y H _. + apply Rinv_le_contravar with (1 := H). +Qed. + +Lemma Ropp_div : forall x y, -x/y = - (x / y). +intros x y; unfold Rdiv; ring. Qed. Lemma double : forall r1, 2 * r1 = r1 + r1. @@ -2018,6 +2065,29 @@ Proof. intros; elim (completeness E H H0); intros; split with x; assumption. Qed. +Lemma Rdiv_lt_0_compat : forall a b, 0 < a -> 0 < b -> 0 < a/b. +Proof. +intros; apply Rmult_lt_0_compat;[|apply Rinv_0_lt_compat]; assumption. +Qed. + +Lemma Rdiv_plus_distr : forall a b c, (a + b)/c = a/c + b/c. +intros a b c; apply Rmult_plus_distr_r. +Qed. + +Lemma Rdiv_minus_distr : forall a b c, (a - b)/c = a/c - b/c. +intros a b c; unfold Rminus, Rdiv; rewrite Rmult_plus_distr_r; ring. +Qed. + +(* A test for equality function. *) +Lemma Req_EM_T : forall r1 r2:R, {r1 = r2} + {r1 <> r2}. +Proof. + intros; destruct (total_order_T r1 r2) as [[H|]|H]. + - right; red; intros ->; elim (Rlt_irrefl r2 H). + - left; assumption. + - right; red; intros ->; elim (Rlt_irrefl r2 H). +Qed. + + (*********************************************************) (** * Definitions of new types *) (*********************************************************) @@ -2035,6 +2105,7 @@ Record negreal : Type := mknegreal {neg :> R; cond_neg : neg < 0}. Record nonzeroreal : Type := mknonzeroreal {nonzero :> R; cond_nonzero : nonzero <> 0}. + (** Compatibility *) Notation prod_neq_R0 := Rmult_integral_contrapositive_currified (only parsing). diff --git a/theories/Reals/RList.v b/theories/Reals/RList.v index ad3002b4..abf8a99d 100644 --- a/theories/Reals/RList.v +++ b/theories/Reals/RList.v @@ -1,6 +1,6 @@ (************************************************************************) (* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *) +(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *) (* \VV/ **************************************************************) (* // * This file is distributed under the terms of the *) (* * GNU Lesser General Public License Version 2.1 *) @@ -181,13 +181,13 @@ Proof. elim Hrecl; intros; assert (H3 := H1 H0); elim H3; intros; elim H4; intros; exists (S x0); split; [ simpl; apply lt_n_S; assumption | simpl; assumption ]. - elim H; intros; elim H0; intros; elim (zerop x0); intro. - rewrite a in H2; simpl in H2; left; assumption. - right; elim Hrecl; intros; apply H4; assert (H5 : S (pred x0) = x0). + elim H; intros; elim H0; intros; destruct (zerop x0) as [->|]. + simpl in H2; left; assumption. + right; elim Hrecl; intros H4 H5; apply H5; assert (H6 : S (pred x0) = x0). symmetry ; apply S_pred with 0%nat; assumption. exists (pred x0); split; - [ simpl in H1; apply lt_S_n; rewrite H5; assumption - | rewrite <- H5 in H2; simpl in H2; assumption ]. + [ simpl in H1; apply lt_S_n; rewrite H6; assumption + | rewrite <- H6 in H2; simpl in H2; assumption ]. Qed. Lemma Rlist_P1 : @@ -208,11 +208,11 @@ Proof. assert (H3 := Hrecl H2); elim H1; intros; elim H3; intros; exists (cons x x0); intros; elim H5; clear H5; intros; split. simpl; rewrite H5; reflexivity. - intros; elim (zerop i); intro. - rewrite a; simpl; assumption. - assert (H8 : i = S (pred i)). + intros; destruct (zerop i) as [->|]. + simpl; assumption. + assert (H9 : i = S (pred i)). apply S_pred with 0%nat; assumption. - rewrite H8; simpl; apply H6; simpl in H7; apply lt_S_n; rewrite <- H8; + rewrite H9; simpl; apply H6; simpl in H7; apply lt_S_n; rewrite <- H9; assumption. Qed. diff --git a/theories/Reals/ROrderedType.v b/theories/Reals/ROrderedType.v index 1e92edd6..0531bd0a 100644 --- a/theories/Reals/ROrderedType.v +++ b/theories/Reals/ROrderedType.v @@ -1,6 +1,6 @@ (************************************************************************) (* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *) +(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *) (* \VV/ **************************************************************) (* // * This file is distributed under the terms of the *) (* * GNU Lesser General Public License Version 2.1 *) @@ -15,7 +15,7 @@ Local Open Scope R_scope. Lemma Req_dec : forall r1 r2:R, {r1 = r2} + {r1 <> r2}. Proof. intros; generalize (total_order_T r1 r2) Rlt_dichotomy_converse; - intuition eauto 3. + intuition eauto. Qed. Definition Reqb r1 r2 := if Req_dec r1 r2 then true else false. diff --git a/theories/Reals/R_Ifp.v b/theories/Reals/R_Ifp.v index 4f4293f3..57ee1d9a 100644 --- a/theories/Reals/R_Ifp.v +++ b/theories/Reals/R_Ifp.v @@ -1,6 +1,6 @@ (************************************************************************) (* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *) +(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *) (* \VV/ **************************************************************) (* // * This file is distributed under the terms of the *) (* * GNU Lesser General Public License Version 2.1 *) diff --git a/theories/Reals/R_sqr.v b/theories/Reals/R_sqr.v index 5900a147..f1e2d6fa 100644 --- a/theories/Reals/R_sqr.v +++ b/theories/Reals/R_sqr.v @@ -1,6 +1,6 @@ (************************************************************************) (* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *) +(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *) (* \VV/ **************************************************************) (* // * This file is distributed under the terms of the *) (* * GNU Lesser General Public License Version 2.1 *) @@ -97,7 +97,7 @@ Qed. Lemma Rsqr_incr_0 : forall x y:R, Rsqr x <= Rsqr y -> 0 <= x -> 0 <= y -> x <= y. Proof. - intros; case (Rle_dec x y); intro; + intros; destruct (Rle_dec x y) as [Hle|Hnle]; [ assumption | cut (y < x); [ intro; unfold Rsqr in H; @@ -109,7 +109,7 @@ Qed. Lemma Rsqr_incr_0_var : forall x y:R, Rsqr x <= Rsqr y -> 0 <= y -> x <= y. Proof. - intros; case (Rle_dec x y); intro; + intros; destruct (Rle_dec x y) as [Hle|Hnle]; [ assumption | cut (y < x); [ intro; unfold Rsqr in H; @@ -146,8 +146,8 @@ Qed. Lemma Rsqr_neg_pos_le_0 : forall x y:R, Rsqr x <= Rsqr y -> 0 <= y -> - y <= x. Proof. - intros; case (Rcase_abs x); intro. - generalize (Ropp_lt_gt_contravar x 0 r); rewrite Ropp_0; intro; + intros; destruct (Rcase_abs x) as [Hlt|Hle]. + generalize (Ropp_lt_gt_contravar x 0 Hlt); rewrite Ropp_0; intro; generalize (Rlt_le 0 (- x) H1); intro; rewrite (Rsqr_neg x) in H; generalize (Rsqr_incr_0 (- x) y H H2 H0); intro; rewrite <- (Ropp_involutive x); apply Ropp_ge_le_contravar; @@ -160,25 +160,23 @@ Qed. Lemma Rsqr_neg_pos_le_1 : forall x y:R, - y <= x -> x <= y -> 0 <= y -> Rsqr x <= Rsqr y. Proof. - intros; case (Rcase_abs x); intro. - generalize (Ropp_lt_gt_contravar x 0 r); rewrite Ropp_0; intro; - generalize (Rlt_le 0 (- x) H2); intro; - generalize (Ropp_le_ge_contravar (- y) x H); rewrite Ropp_involutive; - intro; generalize (Rge_le y (- x) H4); intro; rewrite (Rsqr_neg x); - apply Rsqr_incr_1; assumption. - generalize (Rge_le x 0 r); intro; apply Rsqr_incr_1; assumption. + intros x y H H0 H1; destruct (Rcase_abs x) as [Hlt|Hle]. + apply Ropp_lt_gt_contravar, Rlt_le in Hlt; rewrite Ropp_0 in Hlt; + apply Ropp_le_ge_contravar, Rge_le in H; rewrite Ropp_involutive in H; + rewrite (Rsqr_neg x); apply Rsqr_incr_1; assumption. + apply Rge_le in Hle; apply Rsqr_incr_1; assumption. Qed. Lemma neg_pos_Rsqr_le : forall x y:R, - y <= x -> x <= y -> Rsqr x <= Rsqr y. Proof. - intros; case (Rcase_abs x); intro. - generalize (Ropp_lt_gt_contravar x 0 r); rewrite Ropp_0; intro; - generalize (Ropp_le_ge_contravar (- y) x H); rewrite Ropp_involutive; - intro; generalize (Rge_le y (- x) H2); intro; generalize (Rlt_le 0 (- x) H1); - intro; generalize (Rle_trans 0 (- x) y H4 H3); intro; - rewrite (Rsqr_neg x); apply Rsqr_incr_1; assumption. - generalize (Rge_le x 0 r); intro; generalize (Rle_trans 0 x y H1 H0); intro; - apply Rsqr_incr_1; assumption. + intros x y H H0; destruct (Rcase_abs x) as [Hlt|Hle]. + apply Ropp_lt_gt_contravar, Rlt_le in Hlt; rewrite Ropp_0 in Hlt; + apply Ropp_le_ge_contravar, Rge_le in H; rewrite Ropp_involutive in H. + assert (0 <= y) by (apply Rle_trans with (-x); assumption). + rewrite (Rsqr_neg x); apply Rsqr_incr_1; assumption. + apply Rge_le in Hle; + assert (0 <= y) by (apply Rle_trans with x; assumption). + apply Rsqr_incr_1; assumption. Qed. Lemma Rsqr_abs : forall x:R, Rsqr x = Rsqr (Rabs x). @@ -220,22 +218,22 @@ Qed. Lemma Rsqr_eq_abs_0 : forall x y:R, Rsqr x = Rsqr y -> Rabs x = Rabs y. Proof. - intros; unfold Rabs; case (Rcase_abs x); case (Rcase_abs y); intros. - rewrite (Rsqr_neg x) in H; rewrite (Rsqr_neg y) in H; - generalize (Ropp_lt_gt_contravar y 0 r); - generalize (Ropp_lt_gt_contravar x 0 r0); rewrite Ropp_0; + intros; unfold Rabs; case (Rcase_abs x) as [Hltx|Hgex]; + case (Rcase_abs y) as [Hlty|Hgey]. + rewrite (Rsqr_neg x), (Rsqr_neg y) in H; + generalize (Ropp_lt_gt_contravar y 0 Hlty); + generalize (Ropp_lt_gt_contravar x 0 Hltx); rewrite Ropp_0; intros; generalize (Rlt_le 0 (- x) H0); generalize (Rlt_le 0 (- y) H1); intros; apply Rsqr_inj; assumption. - rewrite (Rsqr_neg x) in H; generalize (Rge_le y 0 r); intro; - generalize (Ropp_lt_gt_contravar x 0 r0); rewrite Ropp_0; + rewrite (Rsqr_neg x) in H; generalize (Rge_le y 0 Hgey); intro; + generalize (Ropp_lt_gt_contravar x 0 Hltx); rewrite Ropp_0; intro; generalize (Rlt_le 0 (- x) H1); intro; apply Rsqr_inj; assumption. - rewrite (Rsqr_neg y) in H; generalize (Rge_le x 0 r0); intro; - generalize (Ropp_lt_gt_contravar y 0 r); rewrite Ropp_0; + rewrite (Rsqr_neg y) in H; generalize (Rge_le x 0 Hgex); intro; + generalize (Ropp_lt_gt_contravar y 0 Hlty); rewrite Ropp_0; intro; generalize (Rlt_le 0 (- y) H1); intro; apply Rsqr_inj; assumption. - generalize (Rge_le x 0 r0); generalize (Rge_le y 0 r); intros; apply Rsqr_inj; - assumption. + apply Rsqr_inj; auto using Rge_le. Qed. Lemma Rsqr_eq_asb_1 : forall x y:R, Rabs x = Rabs y -> Rsqr x = Rsqr y. diff --git a/theories/Reals/R_sqrt.v b/theories/Reals/R_sqrt.v index 38a38400..20319a2b 100644 --- a/theories/Reals/R_sqrt.v +++ b/theories/Reals/R_sqrt.v @@ -1,6 +1,6 @@ (************************************************************************) (* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *) +(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *) (* \VV/ **************************************************************) (* // * This file is distributed under the terms of the *) (* * GNU Lesser General Public License Version 2.1 *) @@ -37,8 +37,8 @@ Lemma sqrt_sqrt : forall x:R, 0 <= x -> sqrt x * sqrt x = x. Proof. intros. unfold sqrt. - case (Rcase_abs x); intro. - elim (Rlt_irrefl _ (Rlt_le_trans _ _ _ r H)). + case (Rcase_abs x) as [Hlt|Hge]. + elim (Rlt_irrefl _ (Rlt_le_trans _ _ _ Hlt H)). rewrite Rsqrt_Rsqrt; reflexivity. Qed. @@ -94,6 +94,10 @@ Proof. intros; unfold Rsqr; apply sqrt_square; assumption. Qed. +Lemma sqrt_pow2 : forall x, 0 <= x -> sqrt (x ^ 2) = x. +intros; simpl; rewrite Rmult_1_r, sqrt_square; auto. +Qed. + Lemma sqrt_Rsqr_abs : forall x:R, sqrt (Rsqr x) = Rabs x. Proof. intro x; rewrite Rsqr_abs; apply sqrt_Rsqr; apply Rabs_pos. @@ -517,3 +521,4 @@ Proof. reflexivity. reflexivity. Qed. + diff --git a/theories/Reals/Ranalysis.v b/theories/Reals/Ranalysis.v index d656817e..3cda675a 100644 --- a/theories/Reals/Ranalysis.v +++ b/theories/Reals/Ranalysis.v @@ -1,6 +1,6 @@ (************************************************************************) (* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *) +(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *) (* \VV/ **************************************************************) (* // * This file is distributed under the terms of the *) (* * GNU Lesser General Public License Version 2.1 *) diff --git a/theories/Reals/Ranalysis1.v b/theories/Reals/Ranalysis1.v index 2f39c00b..875eebbb 100644 --- a/theories/Reals/Ranalysis1.v +++ b/theories/Reals/Ranalysis1.v @@ -1,6 +1,6 @@ (************************************************************************) (* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *) +(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *) (* \VV/ **************************************************************) (* // * This file is distributed under the terms of the *) (* * GNU Lesser General Public License Version 2.1 *) @@ -77,6 +77,23 @@ Definition continuity f : Prop := forall x:R, continuity_pt f x. Arguments continuity_pt f%F x0%R. Arguments continuity f%F. +Lemma continuity_pt_locally_ext : + forall f g a x, 0 < a -> (forall y, R_dist y x < a -> f y = g y) -> + continuity_pt f x -> continuity_pt g x. +intros f g a x a0 q cf eps ep. +destruct (cf eps ep) as [a' [a'p Pa']]. +exists (Rmin a a'); split. + unfold Rmin; destruct (Rle_dec a a'). + assumption. + assumption. +intros y cy; rewrite <- !q. + apply Pa'. + split;[| apply Rlt_le_trans with (Rmin a a');[ | apply Rmin_r]];tauto. + rewrite R_dist_eq; assumption. +apply Rlt_le_trans with (Rmin a a');[ | apply Rmin_l]; tauto. +Qed. + + (**********) Lemma continuity_pt_plus : forall f1 f2 (x0:R), @@ -477,6 +494,47 @@ Proof. auto with real. Qed. +(* Extensionally equal functions have the same derivative. *) + +Lemma derivable_pt_lim_ext : forall f g x l, (forall z, f z = g z) -> + derivable_pt_lim f x l -> derivable_pt_lim g x l. +intros f g x l fg df e ep; destruct (df e ep) as [d pd]; exists d; intros h; +rewrite <- !fg; apply pd. +Qed. + +(* extensionally equal functions have the same derivative, locally. *) + +Lemma derivable_pt_lim_locally_ext : forall f g x a b l, + a < x < b -> + (forall z, a < z < b -> f z = g z) -> + derivable_pt_lim f x l -> derivable_pt_lim g x l. +intros f g x a b l axb fg df e ep. +destruct (df e ep) as [d pd]. +assert (d'h : 0 < Rmin d (Rmin (b - x) (x - a))). + apply Rmin_pos;[apply cond_pos | apply Rmin_pos; apply Rlt_Rminus; tauto]. +exists (mkposreal _ d'h); simpl; intros h hn0 cmp. +rewrite <- !fg;[ |assumption | ]. + apply pd;[assumption |]. + apply Rlt_le_trans with (1 := cmp), Rmin_l. +assert (-h < x - a). + apply Rle_lt_trans with (1 := Rle_abs _). + rewrite Rabs_Ropp; apply Rlt_le_trans with (1 := cmp). + rewrite Rmin_assoc; apply Rmin_r. +assert (h < b - x). + apply Rle_lt_trans with (1 := Rle_abs _). + apply Rlt_le_trans with (1 := cmp). + rewrite Rmin_comm, <- Rmin_assoc; apply Rmin_l. +split. + apply (Rplus_lt_reg_l (- h)). + replace ((-h) + (x + h)) with x by ring. + apply (Rplus_lt_reg_r (- a)). + replace (((-h) + a) + - a) with (-h) by ring. + assumption. +apply (Rplus_lt_reg_r (- x)). +replace (x + h + - x) with h by ring. +assumption. +Qed. + (***********************************) (** * derivability -> continuity *) @@ -639,6 +697,24 @@ Proof. unfold mult_real_fct, mult_fct, fct_cte; reflexivity. Qed. +Lemma derivable_pt_lim_div_scal : + forall f x l a, derivable_pt_lim f x l -> + derivable_pt_lim (fun y => f y / a) x (l / a). +intros f x l a df; + apply (derivable_pt_lim_ext (fun y => /a * f y)). + intros z; rewrite Rmult_comm; reflexivity. +unfold Rdiv; rewrite Rmult_comm; apply derivable_pt_lim_scal; assumption. +Qed. + +Lemma derivable_pt_lim_scal_right : + forall f x l a, derivable_pt_lim f x l -> + derivable_pt_lim (fun y => f y * a) x (l * a). +intros f x l a df; + apply (derivable_pt_lim_ext (fun y => a * f y)). + intros z; rewrite Rmult_comm; reflexivity. +unfold Rdiv; rewrite Rmult_comm; apply derivable_pt_lim_scal; assumption. +Qed. + Lemma derivable_pt_lim_id : forall x:R, derivable_pt_lim id x 1. Proof. intro; unfold derivable_pt_lim. @@ -1066,15 +1142,8 @@ Lemma pr_nu : forall f (x:R) (pr1 pr2:derivable_pt f x), derive_pt f x pr1 = derive_pt f x pr2. Proof. - intros. - unfold derivable_pt in pr1. - unfold derivable_pt in pr2. - elim pr1; intros. - elim pr2; intros. - unfold derivable_pt_abs in p. - unfold derivable_pt_abs in p0. - simpl. - apply (uniqueness_limite f x x0 x1 p p0). + intros f x (x0,H0) (x1,H1). + apply (uniqueness_limite f x x0 x1 H0 H1). Qed. @@ -1123,7 +1192,7 @@ Proof. case (Rcase_abs ((f (c + Rmin (delta / 2) ((b + - c) / 2)) + - f c) / - Rmin (delta / 2) ((b + - c) / 2) + - l)); intro. + Rmin (delta / 2) ((b + - c) / 2) + - l)) as [Hlt|Hge]. replace (- ((f (c + Rmin (delta / 2) ((b + - c) / 2)) + - f c) / @@ -1165,7 +1234,7 @@ Proof. (H20 := Rge_le ((f (c + Rmin (delta / 2) ((b + - c) / 2)) + - f c) / - Rmin (delta / 2) ((b + - c) / 2) + - l) 0 r). + Rmin (delta / 2) ((b + - c) / 2) + - l) 0 Hge). elim (Rlt_irrefl _ (Rle_lt_trans _ _ _ H20 H18)). assumption. rewrite <- Ropp_0; @@ -1242,17 +1311,16 @@ Proof. (mkposreal ((b - c) / 2) H8)). unfold Rdiv; apply Rmult_lt_0_compat; [ apply (cond_pos delta) | apply Rinv_0_lt_compat; prove_sup0 ]. - unfold Rabs; case (Rcase_abs (Rmin (delta / 2) ((b - c) / 2))). - intro. + unfold Rabs; case (Rcase_abs (Rmin (delta / 2) ((b - c) / 2))) as [Hlt|Hge]. cut (0 < delta / 2). intro. generalize (Rmin_stable_in_posreal (mkposreal (delta / 2) H10) (mkposreal ((b - c) / 2) H8)); simpl; intro; - elim (Rlt_irrefl 0 (Rlt_trans 0 (Rmin (delta / 2) ((b - c) / 2)) 0 H11 r)). + elim (Rlt_irrefl 0 (Rlt_trans 0 (Rmin (delta / 2) ((b - c) / 2)) 0 H11 Hlt)). unfold Rdiv; apply Rmult_lt_0_compat; [ apply (cond_pos delta) | apply Rinv_0_lt_compat; prove_sup0 ]. - intro; apply Rle_lt_trans with (delta / 2). + apply Rle_lt_trans with (delta / 2). apply Rmin_l. unfold Rdiv; apply Rmult_lt_reg_l with 2. prove_sup0. @@ -1311,13 +1379,12 @@ Proof. case (Rcase_abs ((f (c + Rmax (- (delta / 2)) ((a + - c) / 2)) + - f c) / - Rmax (- (delta / 2)) ((a + - c) / 2) + - l)). - intro; - elim + Rmax (- (delta / 2)) ((a + - c) / 2) + - l)) as [Hlt|Hge]. + elim (Rlt_irrefl 0 (Rlt_trans 0 ((f (c + Rmax (- (delta / 2)) ((a + - c) / 2)) + - f c) / - Rmax (- (delta / 2)) ((a + - c) / 2) + - l) 0 H19 r)). + Rmax (- (delta / 2)) ((a + - c) / 2) + - l) 0 H19 Hlt)). intros; generalize (Rplus_lt_compat_r l @@ -1380,8 +1447,8 @@ Proof. apply Rplus_lt_compat_l; assumption. field; discrR. assumption. - unfold Rabs; case (Rcase_abs (Rmax (- (delta / 2)) ((a - c) / 2))). - intro; generalize (RmaxLess1 (- (delta / 2)) ((a - c) / 2)); intro; + unfold Rabs; case (Rcase_abs (Rmax (- (delta / 2)) ((a - c) / 2))) as [Hlt|Hge]. + generalize (RmaxLess1 (- (delta / 2)) ((a - c) / 2)); intro; generalize (Ropp_le_ge_contravar (- (delta / 2)) (Rmax (- (delta / 2)) ((a - c) / 2)) H12); rewrite Ropp_involutive; intro; @@ -1402,7 +1469,7 @@ Proof. generalize (Rmax_stable_in_negreal (mknegreal (- (delta / 2)) H13) (mknegreal ((a - c) / 2) H12)); simpl; - intro; generalize (Rge_le (Rmax (- (delta / 2)) ((a - c) / 2)) 0 r); + intro; generalize (Rge_le (Rmax (- (delta / 2)) ((a - c) / 2)) 0 Hge); intro; elim (Rlt_irrefl 0 @@ -1494,11 +1561,10 @@ Proof. cut (0 <= (f (x + delta / 2) - f x) / (delta / 2)). intro; cut (0 <= (f (x + delta / 2) - f x) / (delta / 2) - l). intro; unfold Rabs; - case (Rcase_abs ((f (x + delta / 2) - f x) / (delta / 2) - l)). - intro; - elim + case (Rcase_abs ((f (x + delta / 2) - f x) / (delta / 2) - l)) as [Hlt|Hge]. + elim (Rlt_irrefl 0 - (Rle_lt_trans 0 ((f (x + delta / 2) - f x) / (delta / 2) - l) 0 H12 r)). + (Rle_lt_trans 0 ((f (x + delta / 2) - f x) / (delta / 2) - l) 0 H12 Hlt)). intros; generalize (Rplus_lt_compat_r l ((f (x + delta / 2) - f x) / (delta / 2) - l) @@ -1555,7 +1621,7 @@ Proof. [ apply (cond_pos delta) | apply Rinv_0_lt_compat; prove_sup0 ]. unfold Rdiv; rewrite <- Ropp_mult_distr_l_reverse; apply Rmult_lt_0_compat. - apply Rplus_lt_reg_r with l. + apply Rplus_lt_reg_l with l. unfold Rminus; rewrite Rplus_opp_r; rewrite Rplus_0_r; assumption. apply Rinv_0_lt_compat; prove_sup0. Qed. diff --git a/theories/Reals/Ranalysis2.v b/theories/Reals/Ranalysis2.v index b070cdaa..eb646913 100644 --- a/theories/Reals/Ranalysis2.v +++ b/theories/Reals/Ranalysis2.v @@ -1,6 +1,6 @@ (************************************************************************) (* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *) +(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *) (* \VV/ **************************************************************) (* // * This file is distributed under the terms of the *) (* * GNU Lesser General Public License Version 2.1 *) @@ -9,6 +9,7 @@ Require Import Rbase. Require Import Rfunctions. Require Import Ranalysis1. +Require Import Omega. Local Open Scope R_scope. (**********) @@ -432,17 +433,17 @@ Proof. unfold IZR; unfold INR, Pos.to_nat; simpl; intro; elim (Rlt_irrefl 1 (Rlt_trans _ _ _ H13 H12)). apply IZR_lt; omega. - unfold Rabs; case (Rcase_abs (/ 2)); intro. + unfold Rabs; case (Rcase_abs (/ 2)) as [Hlt|Hge]. assert (Hyp : 0 < 2). prove_sup0. - assert (H11 := Rmult_lt_compat_l 2 _ _ Hyp r); rewrite Rmult_0_r in H11; + assert (H11 := Rmult_lt_compat_l 2 _ _ Hyp Hlt); rewrite Rmult_0_r in H11; rewrite <- Rinv_r_sym in H11; [ idtac | discrR ]. elim (Rlt_irrefl 0 (Rlt_trans _ _ _ Rlt_0_1 H11)). reflexivity. apply (Rabs_pos_lt _ H0). ring. assert (H6 := Req_dec x0 (x0 + h)); elim H6; intro. - intro; rewrite <- H7; unfold dist, R_met; unfold R_dist; + intro; rewrite <- H7. unfold R_met, dist; unfold R_dist; unfold Rminus; rewrite Rplus_opp_r; rewrite Rabs_R0; apply Rabs_pos_lt. unfold Rdiv; apply prod_neq_R0; diff --git a/theories/Reals/Ranalysis3.v b/theories/Reals/Ranalysis3.v index 614f12bd..407f6410 100644 --- a/theories/Reals/Ranalysis3.v +++ b/theories/Reals/Ranalysis3.v @@ -1,6 +1,6 @@ (************************************************************************) (* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *) +(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *) (* \VV/ **************************************************************) (* // * This file is distributed under the terms of the *) (* * GNU Lesser General Public License Version 2.1 *) @@ -731,7 +731,7 @@ Proof. rewrite <- (Rabs_Ropp (f2 a - f2 x)) in H6. rewrite Ropp_minus_distr in H6. assert (H7 := Rle_lt_trans _ _ _ (Rabs_triang_inv _ _) H6). - apply Rplus_lt_reg_r with (- Rabs (f2 a) + Rabs (f2 x) / 2). + apply Rplus_lt_reg_l with (- Rabs (f2 a) + Rabs (f2 x) / 2). rewrite Rplus_assoc. rewrite <- double_var. do 2 rewrite (Rplus_comm (- Rabs (f2 a))). diff --git a/theories/Reals/Ranalysis4.v b/theories/Reals/Ranalysis4.v index 2fa17e20..ae2013f0 100644 --- a/theories/Reals/Ranalysis4.v +++ b/theories/Reals/Ranalysis4.v @@ -1,6 +1,6 @@ (************************************************************************) (* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *) +(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *) (* \VV/ **************************************************************) (* // * This file is distributed under the terms of the *) (* * GNU Lesser General Public License Version 2.1 *) @@ -13,6 +13,7 @@ Require Import Rtrigo1. Require Import Ranalysis1. Require Import Ranalysis3. Require Import Exp_prop. +Require Import MVT. Local Open Scope R_scope. (**********) @@ -26,7 +27,7 @@ Proof. apply derivable_pt_const. assumption. assumption. - unfold div_fct, inv_fct, fct_cte; intro X0; elim X0; intros; + unfold div_fct, inv_fct, fct_cte; intros (x0,p); unfold derivable_pt; exists x0; unfold derivable_pt_abs; unfold derivable_pt_lim; unfold derivable_pt_abs in p; unfold derivable_pt_lim in p; @@ -41,11 +42,7 @@ Lemma pr_nu_var : forall (f g:R -> R) (x:R) (pr1:derivable_pt f x) (pr2:derivable_pt g x), f = g -> derive_pt f x pr1 = derive_pt g x pr2. Proof. - unfold derivable_pt, derive_pt; intros. - elim pr1; intros. - elim pr2; intros. - simpl. - rewrite H in p. + unfold derivable_pt, derive_pt; intros f g x (x0,p0) (x1,p1) ->. apply uniqueness_limite with g x; assumption. Qed. @@ -54,14 +51,11 @@ Lemma pr_nu_var2 : forall (f g:R -> R) (x:R) (pr1:derivable_pt f x) (pr2:derivable_pt g x), (forall h:R, f h = g h) -> derive_pt f x pr1 = derive_pt g x pr2. Proof. - unfold derivable_pt, derive_pt; intros. - elim pr1; intros. - elim pr2; intros. - simpl. - assert (H0 := uniqueness_step2 _ _ _ p). - assert (H1 := uniqueness_step2 _ _ _ p0). + unfold derivable_pt, derive_pt; intros f g x (x0,p0) (x1,p1) H. + assert (H0 := uniqueness_step2 _ _ _ p0). + assert (H1 := uniqueness_step2 _ _ _ p1). cut (limit1_in (fun h:R => (f (x + h) - f x) / h) (fun h:R => h <> 0) x1 0). - intro; assert (H3 := uniqueness_step1 _ _ _ _ H0 H2). + intro H2; assert (H3 := uniqueness_step1 _ _ _ _ H0 H2). assumption. unfold limit1_in; unfold limit_in; unfold dist; simpl; unfold R_dist; unfold limit1_in in H1; @@ -117,14 +111,14 @@ Proof. rewrite Rplus_opp_r; rewrite Rabs_R0; apply H0. apply H1. apply Rle_ge. - case (Rcase_abs h); intro. - rewrite (Rabs_left h r) in H2. - left; rewrite Rplus_comm; apply Rplus_lt_reg_r with (- h); rewrite Rplus_0_r; + destruct (Rcase_abs h) as [Hlt|Hgt]. + rewrite (Rabs_left h Hlt) in H2. + left; rewrite Rplus_comm; apply Rplus_lt_reg_l with (- h); rewrite Rplus_0_r; rewrite <- Rplus_assoc; rewrite Rplus_opp_l; rewrite Rplus_0_l; apply H2. apply Rplus_le_le_0_compat. left; apply H. - apply Rge_le; apply r. + apply Rge_le; apply Hgt. left; apply H. Qed. @@ -145,13 +139,13 @@ Proof. rewrite <- Rinv_r_sym. rewrite Ropp_involutive; rewrite Rplus_opp_l; rewrite Rabs_R0; apply H0. apply H2. - case (Rcase_abs h); intro. + destruct (Rcase_abs h) as [Hlt|Hgt]. apply Ropp_lt_cancel. rewrite Ropp_0; rewrite Ropp_plus_distr; apply Rplus_lt_0_compat. apply H1. - apply Ropp_0_gt_lt_contravar; apply r. - rewrite (Rabs_right h r) in H3. - apply Rplus_lt_reg_r with (- x); rewrite Rplus_0_r; rewrite <- Rplus_assoc; + apply Ropp_0_gt_lt_contravar; apply Hlt. + rewrite (Rabs_right h Hgt) in H3. + apply Rplus_lt_reg_l with (- x); rewrite Rplus_0_r; rewrite <- Rplus_assoc; rewrite Rplus_opp_l; rewrite Rplus_0_l; apply H3. apply H. apply Ropp_0_gt_lt_contravar; apply H. @@ -161,13 +155,12 @@ Qed. Lemma Rderivable_pt_abs : forall x:R, x <> 0 -> derivable_pt Rabs x. Proof. intros. - case (total_order_T x 0); intro. - elim s; intro. + destruct (total_order_T x 0) as [[Hlt|Heq]|Hgt]. unfold derivable_pt; exists (-1). - apply (Rabs_derive_2 x a). - elim H; exact b. + apply (Rabs_derive_2 x Hlt). + elim H; exact Heq. unfold derivable_pt; exists 1. - apply (Rabs_derive_1 x r). + apply (Rabs_derive_1 x Hgt). Qed. (** Rabsolu is continuous for all x *) @@ -406,3 +399,14 @@ Proof. intro; apply derive_pt_eq_0. apply derivable_pt_lim_sinh. Qed. + +Lemma sinh_lt : forall x y, x < y -> sinh x < sinh y. +intros x y xy; destruct (MVT_cor2 sinh cosh x y xy) as [c [Pc _]]. + intros; apply derivable_pt_lim_sinh. +apply Rplus_lt_reg_l with (Ropp (sinh x)); rewrite Rplus_opp_l, Rplus_comm. +unfold Rminus at 1 in Pc; rewrite Pc; apply Rmult_lt_0_compat;[ | ]. + unfold cosh; apply Rmult_lt_0_compat;[|apply Rinv_0_lt_compat, Rlt_0_2]. + now apply Rplus_lt_0_compat; apply exp_pos. +now apply Rlt_Rminus; assumption. +Qed. + diff --git a/theories/Reals/Ranalysis5.v b/theories/Reals/Ranalysis5.v index 5c3b03fa..27615c59 100644 --- a/theories/Reals/Ranalysis5.v +++ b/theories/Reals/Ranalysis5.v @@ -1,6 +1,6 @@ (************************************************************************) (* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *) +(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *) (* \VV/ **************************************************************) (* // * This file is distributed under the terms of the *) (* * GNU Lesser General Public License Version 2.1 *) @@ -14,6 +14,7 @@ Require Import Fourier. Require Import RiemannInt. Require Import SeqProp. Require Import Max. +Require Import Omega. Local Open Scope R_scope. (** * Preliminaries lemmas *) @@ -164,8 +165,8 @@ assert (forall x l, lb < x < ub -> (derivable_pt_abs f x l <-> derivable_pt_abs apply Rlt_le_trans with (r2:=Rmin delta (Rmin (ub - a) (a - lb))) ; [| apply Rmin_r] ; assumption. unfold derivable_pt in Prf. unfold derivable_pt in Prg. - elim Prf; intros. - elim Prg; intros. + elim Prf; intros x0 p. + elim Prg; intros x1 p0. assert (Temp := p); rewrite H in Temp. unfold derivable_pt_abs in p. unfold derivable_pt_abs in p0. @@ -294,8 +295,8 @@ intros. (* f x y f_cont_interv x_lt_y fx_neg fy_pos.*) generalize (dicho_lb_cv x y (fun z:R => cond_positivity (f z)) H3). generalize (dicho_up_cv x y (fun z:R => cond_positivity (f z)) H3). intros X X0. - elim X; intros. - elim X0; intros. + elim X; intros x0 p. + elim X0; intros x1 p0. assert (H4 := cv_dicho _ _ _ _ _ H3 p0 p). rewrite H4 in p0. exists x0. @@ -337,14 +338,14 @@ intros. (* f x y f_cont_interv x_lt_y fx_neg fy_pos.*) left; assumption. intro. unfold cond_positivity in |- *. - case (Rle_dec 0 z); intro. + destruct (Rle_dec 0 z) as [|Hnotle]. split. intro; assumption. intro; reflexivity. split. intro feqt;discriminate feqt. intro. - elim n0; assumption. + elim Hnotle; assumption. unfold Vn in |- *. cut (forall z:R, cond_positivity z = false <-> z < 0). intros. @@ -358,10 +359,10 @@ intros. (* f x y f_cont_interv x_lt_y fx_neg fy_pos.*) assumption. intro. unfold cond_positivity in |- *. - case (Rle_dec 0 z); intro. + destruct (Rle_dec 0 z) as [Hle|]. split. intro feqt; discriminate feqt. - intro; elim (Rlt_irrefl _ (Rle_lt_trans _ _ _ r H7)). + intro; elim (Rlt_irrefl _ (Rle_lt_trans _ _ _ Hle H7)). split. intro; auto with real. intro; reflexivity. @@ -370,10 +371,9 @@ intros. (* f x y f_cont_interv x_lt_y fx_neg fy_pos.*) assert (Temp : x <= x0 <= y). apply IVT_interv_prelim1 with (D:=(fun z : R => cond_positivity (f z))) ; assumption. assert (H7 := continuity_seq f Wn x0 (H x0 Temp) H5). - case (total_order_T 0 (f x0)); intro. - elim s; intro. + destruct (total_order_T 0 (f x0)) as [[Hlt|<-]|Hgt]. left; assumption. - rewrite <- b; right; reflexivity. + right; reflexivity. unfold Un_cv in H7; unfold R_dist in H7. cut (0 < - f x0). intro. @@ -383,7 +383,7 @@ intros. (* f x y f_cont_interv x_lt_y fx_neg fy_pos.*) rewrite Rabs_right in H11. pattern (- f x0) at 1 in H11; rewrite <- Rplus_0_r in H11. unfold Rminus in H11; rewrite (Rplus_comm (f (Wn x2))) in H11. - assert (H12 := Rplus_lt_reg_r _ _ _ H11). + assert (H12 := Rplus_lt_reg_l _ _ _ H11). assert (H13 := H6 x2). elim (Rlt_irrefl _ (Rle_lt_trans _ _ _ H13 H12)). apply Rle_ge; left; unfold Rminus in |- *; apply Rplus_le_lt_0_compat. @@ -396,29 +396,28 @@ intros. (* f x y f_cont_interv x_lt_y fx_neg fy_pos.*) assert (Temp : x <= x0 <= y). apply IVT_interv_prelim1 with (D:=(fun z : R => cond_positivity (f z))) ; assumption. assert (H7 := continuity_seq f Vn x0 (H x0 Temp) H5). - case (total_order_T 0 (f x0)); intro. - elim s; intro. + destruct (total_order_T 0 (f x0)) as [[Hlt|Heq]|]. unfold Un_cv in H7; unfold R_dist in H7. - elim (H7 (f x0) a); intros. - cut (x2 >= x2)%nat; [ intro | unfold ge in |- *; apply le_n ]. + elim (H7 (f x0) Hlt); intros. + cut (x2 >= x2)%nat; [ intro | unfold ge; apply le_n ]. assert (H10 := H8 x2 H9). rewrite Rabs_left in H10. pattern (f x0) at 2 in H10; rewrite <- Rplus_0_r in H10. rewrite Ropp_minus_distr' in H10. unfold Rminus in H10. - assert (H11 := Rplus_lt_reg_r _ _ _ H10). + assert (H11 := Rplus_lt_reg_l _ _ _ H10). assert (H12 := H6 x2). cut (0 < f (Vn x2)). intro. elim (Rlt_irrefl _ (Rlt_le_trans _ _ _ H13 H12)). rewrite <- (Ropp_involutive (f (Vn x2))). apply Ropp_0_gt_lt_contravar; assumption. - apply Rplus_lt_reg_r with (f x0 - f (Vn x2)). + apply Rplus_lt_reg_l with (f x0 - f (Vn x2)). rewrite Rplus_0_r; replace (f x0 - f (Vn x2) + (f (Vn x2) - f x0)) with 0; [ unfold Rminus in |- *; apply Rplus_lt_le_0_compat | ring ]. assumption. apply Ropp_0_ge_le_contravar; apply Rle_ge; apply H6. - right; rewrite <- b; reflexivity. + right; rewrite <- Heq; reflexivity. left; assumption. unfold Vn in |- *; assumption. Qed. @@ -695,7 +694,7 @@ intros f g lb ub x Prf g_cont_pur lb_lt_ub x_encad Prg_incr f_eq_g df_neq. exists deltatemp ; exact Htemp. elim (Hf_deriv eps eps_pos). intros deltatemp Htemp. - red in Hlinv ; red in Hlinv ; simpl dist in Hlinv ; unfold R_dist in Hlinv. + red in Hlinv ; red in Hlinv ; unfold dist in Hlinv ; unfold R_dist in Hlinv. assert (Hlinv' := Hlinv (fun h => (f (y+h) - f y)/h) (fun h => h <>0) l 0). unfold limit1_in, limit_in, dist in Hlinv' ; simpl in Hlinv'. unfold R_dist in Hlinv'. assert (Premisse : (forall eps : R, @@ -1038,62 +1037,6 @@ Definition mkposreal_lb_ub (x lb ub:R) (lb_lt_x:lb<x) (x_lt_ub:x<ub) : posreal. Defined. (* end hide *) -Definition boule_of_interval x y (h : x < y) : - {c :R & {r : posreal | c - r = x /\ c + r = y}}. -exists ((x + y)/2). -assert (radius : 0 < (y - x)/2). - unfold Rdiv; apply Rmult_lt_0_compat; fourier. - exists (mkposreal _ radius). - simpl; split; unfold Rdiv; field. -Qed. - -Definition boule_in_interval x y z (h : x < z < y) : - {c : R & {r | Boule c r z /\ x < c - r /\ c + r < y}}. -Proof. -assert (cmp : x * /2 + z * /2 < z * /2 + y * /2). -destruct h as [h1 h2]; fourier. -destruct (boule_of_interval _ _ cmp) as [c [r [P1 P2]]]. -exists c, r; split. - destruct h; unfold Boule; simpl; apply Rabs_def1; fourier. -destruct h; split; fourier. -Qed. - -Lemma Ball_in_inter : forall c1 c2 r1 r2 x, - Boule c1 r1 x -> Boule c2 r2 x -> - {r3 : posreal | forall y, Boule x r3 y -> Boule c1 r1 y /\ Boule c2 r2 y}. -intros c1 c2 [r1 r1p] [r2 r2p] x; unfold Boule; simpl; intros in1 in2. -assert (Rmax (c1 - r1)(c2 - r2) < x). - apply Rmax_lub_lt;[revert in1 | revert in2]; intros h; - apply Rabs_def2 in h; destruct h; fourier. -assert (x < Rmin (c1 + r1) (c2 + r2)). - apply Rmin_glb_lt;[revert in1 | revert in2]; intros h; - apply Rabs_def2 in h; destruct h; fourier. -assert (t: 0 < Rmin (x - Rmax (c1 - r1) (c2 - r2)) - (Rmin (c1 + r1) (c2 + r2) - x)). - apply Rmin_glb_lt; fourier. -exists (mkposreal _ t). -apply Rabs_def2 in in1; destruct in1. -apply Rabs_def2 in in2; destruct in2. -assert (c1 - r1 <= Rmax (c1 - r1) (c2 - r2)) by apply Rmax_l. -assert (c2 - r2 <= Rmax (c1 - r1) (c2 - r2)) by apply Rmax_r. -assert (Rmin (c1 + r1) (c2 + r2) <= c1 + r1) by apply Rmin_l. -assert (Rmin (c1 + r1) (c2 + r2) <= c2 + r2) by apply Rmin_r. -assert (Rmin (x - Rmax (c1 - r1) (c2 - r2)) - (Rmin (c1 + r1) (c2 + r2) - x) <= x - Rmax (c1 - r1) (c2 - r2)) - by apply Rmin_l. -assert (Rmin (x - Rmax (c1 - r1) (c2 - r2)) - (Rmin (c1 + r1) (c2 + r2) - x) <= Rmin (c1 + r1) (c2 + r2) - x) - by apply Rmin_r. -simpl. -intros y h; apply Rabs_def2 in h; destruct h;split; apply Rabs_def1; fourier. -Qed. - -Lemma Boule_center : forall x r, Boule x r x. -Proof. -intros x [r rpos]; unfold Boule, Rminus; simpl; rewrite Rplus_opp_r. -rewrite Rabs_pos_eq;[assumption | apply Rle_refl]. -Qed. - Lemma derivable_pt_lim_CVU : forall (fn fn':nat -> R -> R) (f g:R->R) (x:R) c r, Boule c r x -> (forall y n, Boule c r y -> derivable_pt_lim (fn n) y (fn' n y)) -> diff --git a/theories/Reals/Ranalysis_reg.v b/theories/Reals/Ranalysis_reg.v index ea3899fc..4cf90886 100644 --- a/theories/Reals/Ranalysis_reg.v +++ b/theories/Reals/Ranalysis_reg.v @@ -1,6 +1,6 @@ (************************************************************************) (* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *) +(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *) (* \VV/ **************************************************************) (* // * This file is distributed under the terms of the *) (* * GNU Lesser General Public License Version 2.1 *) @@ -28,7 +28,10 @@ Require Export Ranalysis4. Require Export Rpower. Local Open Scope R_scope. -Axiom AppVar : R. +Definition AppVar : R. +Proof. +exact R0. +Qed. (**********) Ltac intro_hyp_glob trm := diff --git a/theories/Reals/Ratan.v b/theories/Reals/Ratan.v index 096c75fe..68718db0 100644 --- a/theories/Reals/Ratan.v +++ b/theories/Reals/Ratan.v @@ -1,6 +1,6 @@ (************************************************************************) (* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *) +(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *) (* \VV/ **************************************************************) (* // * This file is distributed under the terms of the *) (* * GNU Lesser General Public License Version 2.1 *) @@ -18,6 +18,7 @@ Require Import SeqProp. Require Import Ranalysis5. Require Import SeqSeries. Require Import PartSum. +Require Import Omega. Local Open Scope R_scope. @@ -449,9 +450,9 @@ fourier. Qed. Definition frame_tan y : {x | 0 < x < PI/2 /\ Rabs y < tan x}. -destruct (total_order_T (Rabs y) 1). - assert (yle1 : Rabs y <= 1) by (destruct s; fourier). - clear s; exists 1; split;[split; [exact Rlt_0_1 | exact PI2_1] | ]. +destruct (total_order_T (Rabs y) 1) as [Hs|Hgt]. + assert (yle1 : Rabs y <= 1) by (destruct Hs; fourier). + clear Hs; exists 1; split;[split; [exact Rlt_0_1 | exact PI2_1] | ]. apply Rle_lt_trans with (1 := yle1); exact tan_1_gt_1. assert (0 < / (Rabs y + 1)). apply Rinv_0_lt_compat; fourier. @@ -529,7 +530,7 @@ split. assumption. replace (/(Rabs y + 1)) with (2 * u). fourier. - unfold u; field; apply Rgt_not_eq; clear -r; fourier. + unfold u; field; apply Rgt_not_eq; clear -Hgt; fourier. solve[discrR]. apply Rgt_not_eq; assumption. unfold tan. @@ -735,6 +736,16 @@ replace (Rsqr x) with (x ^ 2) by (unfold Rsqr; ring). reflexivity. Qed. +Lemma derivable_pt_lim_atan : + forall x, derivable_pt_lim atan x (/(1 + x^2)). +Proof. +intros x. +apply derive_pt_eq_1 with (derivable_pt_atan x). +replace (x ^ 2) with (x * x) by ring. +rewrite <- (Rmult_1_l (Rinv _)). +apply derive_pt_atan. +Qed. + (** * Definition of the arctangent function as the sum of the arctan power series *) (* Proof taken from Guillaume Melquiond's interval package for Coq *) @@ -818,13 +829,11 @@ intros x Hx eps Heps. apply Rle_lt_trans with (/ INR (2 * N + 1))%R. unfold Rdiv. rewrite Rmult_1_l. - apply Rle_Rinv. + apply Rinv_le_contravar. apply lt_INR_0. omega. - replace 0 with (INR 0) by intuition. - apply lt_INR. + apply le_INR. omega. - intuition. rewrite <- (Rinv_involutive eps). apply Rinv_lt_contravar. apply Rmult_lt_0_compat. diff --git a/theories/Reals/Raxioms.v b/theories/Reals/Raxioms.v index cf6fdbfd..f545d3a0 100644 --- a/theories/Reals/Raxioms.v +++ b/theories/Reals/Raxioms.v @@ -1,6 +1,6 @@ (************************************************************************) (* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *) +(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *) (* \VV/ **************************************************************) (* // * This file is distributed under the terms of the *) (* * GNU Lesser General Public License Version 2.1 *) diff --git a/theories/Reals/Rbase.v b/theories/Reals/Rbase.v index 5541a0f9..7a879f45 100644 --- a/theories/Reals/Rbase.v +++ b/theories/Reals/Rbase.v @@ -1,6 +1,6 @@ (************************************************************************) (* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *) +(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *) (* \VV/ **************************************************************) (* // * This file is distributed under the terms of the *) (* * GNU Lesser General Public License Version 2.1 *) diff --git a/theories/Reals/Rbasic_fun.v b/theories/Reals/Rbasic_fun.v index 225186a6..bb30c0ef 100644 --- a/theories/Reals/Rbasic_fun.v +++ b/theories/Reals/Rbasic_fun.v @@ -1,6 +1,6 @@ (************************************************************************) (* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *) +(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *) (* \VV/ **************************************************************) (* // * This file is distributed under the terms of the *) (* * GNU Lesser General Public License Version 2.1 *) @@ -45,12 +45,12 @@ Qed. (*********) Lemma Rmin_Rgt_l : forall r1 r2 r, Rmin r1 r2 > r -> r1 > r /\ r2 > r. Proof. - intros r1 r2 r; unfold Rmin; case (Rle_dec r1 r2); intros. + intros r1 r2 r; unfold Rmin; case (Rle_dec r1 r2) as [Hle|Hnle]; intros. split. assumption. - unfold Rgt; unfold Rgt in H; exact (Rlt_le_trans r r1 r2 H r0). + unfold Rgt; exact (Rlt_le_trans r r1 r2 H Hle). split. - generalize (Rnot_le_lt r1 r2 n); intro; exact (Rgt_trans r1 r2 r H0 H). + generalize (Rnot_le_lt r1 r2 Hnle); intro; exact (Rgt_trans r1 r2 r H0 H). assumption. Qed. @@ -168,10 +168,10 @@ Lemma Rmax_Rle : forall r1 r2 r, r <= Rmax r1 r2 <-> r <= r1 \/ r <= r2. Proof. intros; split. unfold Rmax; case (Rle_dec r1 r2); intros; auto. - intro; unfold Rmax; case (Rle_dec r1 r2); elim H; clear H; intros; + intro; unfold Rmax; case (Rle_dec r1 r2) as [|Hnle]; elim H; clear H; intros; auto. apply (Rle_trans r r1 r2); auto. - generalize (Rnot_le_lt r1 r2 n); clear n; intro; unfold Rgt in H0; + generalize (Rnot_le_lt r1 r2 Hnle); clear Hnle; intro; unfold Rgt in H0; apply (Rlt_le r r1 (Rle_lt_trans r r2 r1 H H0)). Qed. @@ -262,6 +262,16 @@ Proof. intros; unfold Rmax; case (Rle_dec x y); intro; assumption. Qed. +Lemma Rmax_Rlt : forall x y z, + Rmax x y < z <-> x < z /\ y < z. +Proof. +intros x y z; split. + unfold Rmax; case (Rle_dec x y). + intros xy yz; split;[apply Rle_lt_trans with y|]; assumption. + intros xz xy; split;[|apply Rlt_trans with x;[apply Rnot_le_gt|]];assumption. + intros [h h']; apply Rmax_lub_lt; assumption. +Qed. + (*********) Lemma Rmax_neg : forall x y:R, x < 0 -> y < 0 -> Rmax x y < 0. Proof. @@ -276,9 +286,9 @@ Qed. (*********) Lemma Rcase_abs : forall r, {r < 0} + {r >= 0}. Proof. - intro; generalize (Rle_dec 0 r); intro X; elim X; intro; clear X. - right; apply (Rle_ge 0 r a). - left; fold (0 > r); apply (Rnot_le_lt 0 r b). + intro; generalize (Rle_dec 0 r); intro X; elim X; intro H; clear X. + right; apply (Rle_ge 0 r H). + left; fold (0 > r); apply (Rnot_le_lt 0 r H). Qed. (*********) @@ -320,9 +330,9 @@ Qed. (*********) Lemma Rabs_right : forall r, r >= 0 -> Rabs r = r. Proof. - intros; unfold Rabs; case (Rcase_abs r); intro. + intros; unfold Rabs; case (Rcase_abs r) as [Hlt|Hge]. absurd (r >= 0). - exact (Rlt_not_ge r 0 r0). + exact (Rlt_not_ge r 0 Hlt). assumption. trivial. Qed. @@ -337,9 +347,9 @@ Qed. (*********) Lemma Rabs_pos : forall x:R, 0 <= Rabs x. Proof. - intros; unfold Rabs; case (Rcase_abs x); intro. - generalize (Ropp_lt_gt_contravar x 0 r); intro; unfold Rgt in H; - rewrite Ropp_0 in H; unfold Rle; left; assumption. + intros; unfold Rabs; case (Rcase_abs x) as [Hlt|Hge]. + generalize (Ropp_lt_gt_contravar x 0 Hlt); intro; unfold Rgt in H; + rewrite Ropp_0 in H; left; assumption. apply Rge_le; assumption. Qed. @@ -350,11 +360,18 @@ Qed. Definition RRle_abs := Rle_abs. +Lemma Rabs_le : forall a b, -b <= a <= b -> Rabs a <= b. +Proof. +intros a b; unfold Rabs; case Rcase_abs. + intros _ [it _]; apply Ropp_le_cancel; rewrite Ropp_involutive; exact it. +intros _ [_ it]; exact it. +Qed. + (*********) Lemma Rabs_pos_eq : forall x:R, 0 <= x -> Rabs x = x. Proof. - intros; unfold Rabs; case (Rcase_abs x); intro; - [ generalize (Rgt_not_le 0 x r); intro; exfalso; auto | trivial ]. + intros; unfold Rabs; case (Rcase_abs x) as [Hlt|Hge]; + [ generalize (Rgt_not_le 0 x Hlt); intro; exfalso; auto | trivial ]. Qed. (*********) @@ -366,100 +383,70 @@ Qed. (*********) Lemma Rabs_pos_lt : forall x:R, x <> 0 -> 0 < Rabs x. Proof. - intros; generalize (Rabs_pos x); intro; unfold Rle in H0; elim H0; intro; - auto. - exfalso; clear H0; elim H; clear H; generalize H1; unfold Rabs; - case (Rcase_abs x); intros; auto. - clear r H1; generalize (Rplus_eq_compat_l x 0 (- x) H0); - rewrite (let (H1, H2) := Rplus_ne x in H1); rewrite (Rplus_opp_r x); - trivial. + intros; destruct (Rabs_pos x) as [|Heq]; auto. + apply Rabs_no_R0 in H; symmetry in Heq; contradiction. Qed. (*********) Lemma Rabs_minus_sym : forall x y:R, Rabs (x - y) = Rabs (y - x). Proof. - intros; unfold Rabs; case (Rcase_abs (x - y)); - case (Rcase_abs (y - x)); intros. - generalize (Rminus_lt y x r); generalize (Rminus_lt x y r0); intros; - generalize (Rlt_asym x y H); intro; exfalso; - auto. + intros; unfold Rabs; case (Rcase_abs (x - y)) as [Hlt|Hge]; + case (Rcase_abs (y - x)) as [Hlt'|Hge']. + apply Rminus_lt, Rlt_asym in Hlt; apply Rminus_lt in Hlt'; contradiction. rewrite (Ropp_minus_distr x y); trivial. rewrite (Ropp_minus_distr y x); trivial. - unfold Rge in r, r0; elim r; elim r0; intros; clear r r0. - generalize (Ropp_lt_gt_0_contravar (x - y) H); rewrite (Ropp_minus_distr x y); - intro; unfold Rgt in H0; generalize (Rlt_asym 0 (y - x) H0); - intro; exfalso; auto. - rewrite (Rminus_diag_uniq x y H); trivial. - rewrite (Rminus_diag_uniq y x H0); trivial. - rewrite (Rminus_diag_uniq y x H0); trivial. + destruct Hge; destruct Hge'. + apply Ropp_lt_gt_0_contravar in H; rewrite (Ropp_minus_distr x y) in H; + apply Rlt_asym in H0; contradiction. + apply Rminus_diag_uniq in H0 as ->; trivial. + apply Rminus_diag_uniq in H as ->; trivial. + apply Rminus_diag_uniq in H0 as ->; trivial. Qed. (*********) Lemma Rabs_mult : forall x y:R, Rabs (x * y) = Rabs x * Rabs y. Proof. - intros; unfold Rabs; case (Rcase_abs (x * y)); case (Rcase_abs x); - case (Rcase_abs y); intros; auto. - generalize (Rmult_lt_gt_compat_neg_l y x 0 r r0); intro; - rewrite (Rmult_0_r y) in H; generalize (Rlt_asym (x * y) 0 r1); - intro; unfold Rgt in H; exfalso; rewrite (Rmult_comm y x) in H; - auto. + intros; unfold Rabs; case (Rcase_abs (x * y)) as [Hlt|Hge]; + case (Rcase_abs x) as [Hltx|Hgex]; + case (Rcase_abs y) as [Hlty|Hgey]; auto. + apply Rmult_lt_gt_compat_neg_l with (r:=x), Rlt_asym in Hlty; trivial. + rewrite Rmult_0_r in Hlty; contradiction. rewrite (Ropp_mult_distr_l_reverse x y); trivial. rewrite (Rmult_comm x (- y)); rewrite (Ropp_mult_distr_l_reverse y x); rewrite (Rmult_comm x y); trivial. - unfold Rge in r, r0; elim r; elim r0; clear r r0; intros; unfold Rgt in H, H0. - generalize (Rmult_lt_compat_l x 0 y H H0); intro; rewrite (Rmult_0_r x) in H1; - generalize (Rlt_asym (x * y) 0 r1); intro; exfalso; - auto. - rewrite H in r1; rewrite (Rmult_0_l y) in r1; generalize (Rlt_irrefl 0); - intro; exfalso; auto. - rewrite H0 in r1; rewrite (Rmult_0_r x) in r1; generalize (Rlt_irrefl 0); - intro; exfalso; auto. - rewrite H0 in r1; rewrite (Rmult_0_r x) in r1; generalize (Rlt_irrefl 0); - intro; exfalso; auto. + destruct Hgex as [| ->], Hgey as [| ->]. + apply Rmult_lt_compat_l with (r:=x), Rlt_asym in H0; trivial. + rewrite Rmult_0_r in H0; contradiction. + rewrite Rmult_0_r in Hlt; contradiction (Rlt_irrefl 0). + rewrite Rmult_0_l in Hlt; contradiction (Rlt_irrefl 0). + rewrite Rmult_0_l in Hlt; contradiction (Rlt_irrefl 0). rewrite (Rmult_opp_opp x y); trivial. - unfold Rge in r, r1; elim r; elim r1; clear r r1; intros; unfold Rgt in H0, H. - generalize (Rmult_lt_compat_l y x 0 H0 r0); intro; - rewrite (Rmult_0_r y) in H1; rewrite (Rmult_comm y x) in H1; - generalize (Rlt_asym (x * y) 0 H1); intro; exfalso; - auto. - generalize (Rlt_dichotomy_converse x 0 (or_introl (x > 0) r0)); - generalize (Rlt_dichotomy_converse y 0 (or_intror (y < 0) H0)); - intros; generalize (Rmult_integral x y H); intro; - elim H3; intro; exfalso; auto. - rewrite H0 in H; rewrite (Rmult_0_r x) in H; unfold Rgt in H; - generalize (Rlt_irrefl 0); intro; exfalso; - auto. - rewrite H0; rewrite (Rmult_0_r x); rewrite (Rmult_0_r (- x)); trivial. - unfold Rge in r0, r1; elim r0; elim r1; clear r0 r1; intros; - unfold Rgt in H0, H. - generalize (Rmult_lt_compat_l x y 0 H0 r); intro; rewrite (Rmult_0_r x) in H1; - generalize (Rlt_asym (x * y) 0 H1); intro; exfalso; - auto. - generalize (Rlt_dichotomy_converse y 0 (or_introl (y > 0) r)); - generalize (Rlt_dichotomy_converse 0 x (or_introl (0 > x) H0)); - intros; generalize (Rmult_integral x y H); intro; - elim H3; intro; exfalso; auto. - rewrite H0 in H; rewrite (Rmult_0_l y) in H; unfold Rgt in H; - generalize (Rlt_irrefl 0); intro; exfalso; - auto. - rewrite H0; rewrite (Rmult_0_l y); rewrite (Rmult_0_l (- y)); trivial. + destruct Hge. destruct Hgey. + apply Rmult_lt_compat_r with (r:=y), Rlt_asym in Hltx; trivial. + rewrite Rmult_0_l in Hltx; contradiction. + rewrite H0, Rmult_0_r in H; contradiction (Rlt_irrefl 0). + rewrite <- Ropp_mult_distr_l, H, Ropp_0; trivial. + destruct Hge. destruct Hgex. + apply Rmult_lt_compat_l with (r:=x), Rlt_asym in Hlty; trivial. + rewrite Rmult_0_r in Hlty; contradiction. + rewrite H0, 2!Rmult_0_l; trivial. + rewrite <- Ropp_mult_distr_r, H, Ropp_0; trivial. Qed. (*********) Lemma Rabs_Rinv : forall r, r <> 0 -> Rabs (/ r) = / Rabs r. Proof. - intro; unfold Rabs; case (Rcase_abs r); case (Rcase_abs (/ r)); auto; + intro; unfold Rabs; case (Rcase_abs r) as [Hlt|Hge]; + case (Rcase_abs (/ r)) as [Hlt'|Hge']; auto; intros. apply Ropp_inv_permute; auto. - generalize (Rinv_lt_0_compat r r1); intro; unfold Rge in r0; elim r0; intros. - unfold Rgt in H1; generalize (Rlt_asym 0 (/ r) H1); intro; exfalso; - auto. - generalize (Rlt_dichotomy_converse (/ r) 0 (or_introl (/ r > 0) H0)); intro; - exfalso; auto. - unfold Rge in r1; elim r1; clear r1; intro. - unfold Rgt in H0; generalize (Rlt_asym 0 (/ r) (Rinv_0_lt_compat r H0)); - intro; exfalso; auto. - exfalso; auto. + rewrite <- Ropp_inv_permute; trivial. + destruct Hge' as [| ->]. + apply Rinv_lt_0_compat, Rlt_asym in Hlt; contradiction. + rewrite Ropp_0; trivial. + destruct Hge as [| ->]. + apply Rinv_0_lt_compat, Rlt_asym in H0; contradiction. + contradiction (refl_equal 0). Qed. Lemma Rabs_Ropp : forall x:R, Rabs (- x) = Rabs x. @@ -483,13 +470,14 @@ Qed. (*********) Lemma Rabs_triang : forall a b:R, Rabs (a + b) <= Rabs a + Rabs b. Proof. - intros a b; unfold Rabs; case (Rcase_abs (a + b)); case (Rcase_abs a); - case (Rcase_abs b); intros. + intros a b; unfold Rabs; case (Rcase_abs (a + b)) as [Hlt|Hge]; + case (Rcase_abs a) as [Hlta|Hgea]; + case (Rcase_abs b) as [Hltb|Hgeb]. apply (Req_le (- (a + b)) (- a + - b)); rewrite (Ropp_plus_distr a b); reflexivity. (**) rewrite (Ropp_plus_distr a b); apply (Rplus_le_compat_l (- a) (- b) b); - unfold Rle; unfold Rge in r; elim r; intro. + unfold Rle; elim Hgeb; intro. left; unfold Rgt in H; generalize (Rplus_lt_compat_l (- b) 0 b H); intro; elim (Rplus_ne (- b)); intros v w; rewrite v in H0; clear v w; rewrite (Rplus_opp_l b) in H0; apply (Rlt_trans (- b) 0 b H0 H). @@ -497,24 +485,24 @@ Proof. (**) rewrite (Ropp_plus_distr a b); rewrite (Rplus_comm (- a) (- b)); rewrite (Rplus_comm a (- b)); apply (Rplus_le_compat_l (- b) (- a) a); - unfold Rle; unfold Rge in r0; elim r0; intro. + unfold Rle; elim Hgea; intro. left; unfold Rgt in H; generalize (Rplus_lt_compat_l (- a) 0 a H); intro; elim (Rplus_ne (- a)); intros v w; rewrite v in H0; clear v w; rewrite (Rplus_opp_l a) in H0; apply (Rlt_trans (- a) 0 a H0 H). right; rewrite H; apply Ropp_0. (**) - exfalso; generalize (Rplus_ge_compat_l a b 0 r); intro; + exfalso; generalize (Rplus_ge_compat_l a b 0 Hgeb); intro; elim (Rplus_ne a); intros v w; rewrite v in H; clear v w; - generalize (Rge_trans (a + b) a 0 H r0); intro; clear H; + generalize (Rge_trans (a + b) a 0 H Hgea); intro; clear H; unfold Rge in H0; elim H0; intro; clear H0. - unfold Rgt in H; generalize (Rlt_asym (a + b) 0 r1); intro; auto. + unfold Rgt in H; generalize (Rlt_asym (a + b) 0 Hlt); intro; auto. absurd (a + b = 0); auto. apply (Rlt_dichotomy_converse (a + b) 0); left; assumption. (**) - exfalso; generalize (Rplus_lt_compat_l a b 0 r); intro; + exfalso; generalize (Rplus_lt_compat_l a b 0 Hltb); intro; elim (Rplus_ne a); intros v w; rewrite v in H; clear v w; - generalize (Rlt_trans (a + b) a 0 H r0); intro; clear H; - unfold Rge in r1; elim r1; clear r1; intro. + generalize (Rlt_trans (a + b) a 0 H Hlta); intro; clear H; + destruct Hge. unfold Rgt in H; generalize (Rlt_trans (a + b) 0 (a + b) H0 H); intro; apply (Rlt_irrefl (a + b)); assumption. rewrite H in H0; apply (Rlt_irrefl 0); assumption. @@ -522,16 +510,16 @@ Proof. rewrite (Rplus_comm a b); rewrite (Rplus_comm (- a) b); apply (Rplus_le_compat_l b a (- a)); apply (Rminus_le a (- a)); unfold Rminus; rewrite (Ropp_involutive a); - generalize (Rplus_lt_compat_l a a 0 r0); clear r r1; + generalize (Rplus_lt_compat_l a a 0 Hlta); clear Hge Hgeb; intro; elim (Rplus_ne a); intros v w; rewrite v in H; - clear v w; generalize (Rlt_trans (a + a) a 0 H r0); + clear v w; generalize (Rlt_trans (a + a) a 0 H Hlta); intro; apply (Rlt_le (a + a) 0 H0). (**) apply (Rplus_le_compat_l a b (- b)); apply (Rminus_le b (- b)); unfold Rminus; rewrite (Ropp_involutive b); - generalize (Rplus_lt_compat_l b b 0 r); clear r0 r1; + generalize (Rplus_lt_compat_l b b 0 Hltb); clear Hge Hgea; intro; elim (Rplus_ne b); intros v w; rewrite v in H; - clear v w; generalize (Rlt_trans (b + b) b 0 H r); + clear v w; generalize (Rlt_trans (b + b) b 0 H Hltb); intro; apply (Rlt_le (b + b) 0 H0). (**) unfold Rle; right; reflexivity. @@ -585,15 +573,15 @@ Qed. (*********) Lemma Rabs_def2 : forall x a:R, Rabs x < a -> x < a /\ - a < x. Proof. - unfold Rabs; intro x; case (Rcase_abs x); intros. - generalize (Ropp_gt_lt_0_contravar x r); unfold Rgt; intro; + unfold Rabs; intro x; case (Rcase_abs x) as [Hlt|Hge]; intros. + generalize (Ropp_gt_lt_0_contravar x Hlt); unfold Rgt; intro; generalize (Rlt_trans 0 (- x) a H0 H); intro; split. - apply (Rlt_trans x 0 a r H1). + apply (Rlt_trans x 0 a Hlt H1). generalize (Ropp_lt_gt_contravar (- x) a H); rewrite (Ropp_involutive x); unfold Rgt; trivial. - fold (a > x) in H; generalize (Rgt_ge_trans a x 0 H r); intro; + fold (a > x) in H; generalize (Rgt_ge_trans a x 0 H Hge); intro; generalize (Ropp_lt_gt_0_contravar a H0); intro; fold (0 > - a); - generalize (Rge_gt_trans x 0 (- a) r H1); unfold Rgt; + generalize (Rge_gt_trans x 0 (- a) Hge H1); unfold Rgt; intro; split; assumption. Qed. @@ -637,3 +625,51 @@ Proof. intros. now rewrite Rabs_Zabs. Qed. + +Lemma Ropp_Rmax : forall x y, - Rmax x y = Rmin (-x) (-y). +intros x y; apply Rmax_case_strong. + now intros w; rewrite Rmin_left;[ | apply Rge_le, Ropp_le_ge_contravar]. +now intros w; rewrite Rmin_right; [ | apply Rge_le, Ropp_le_ge_contravar]. +Qed. + +Lemma Ropp_Rmin : forall x y, - Rmin x y = Rmax (-x) (-y). +intros x y; apply Rmin_case_strong. + now intros w; rewrite Rmax_left;[ | apply Rge_le, Ropp_le_ge_contravar]. +now intros w; rewrite Rmax_right; [ | apply Rge_le, Ropp_le_ge_contravar]. +Qed. + +Lemma Rmax_assoc : forall a b c, Rmax a (Rmax b c) = Rmax (Rmax a b) c. +Proof. +intros a b c. +unfold Rmax; destruct (Rle_dec b c); destruct (Rle_dec a b); + destruct (Rle_dec a c); destruct (Rle_dec b c); auto with real; + match goal with + | id : ~ ?x <= ?y, id2 : ?x <= ?z |- _ => + case id; apply Rle_trans with z; auto with real + | id : ~ ?x <= ?y, id2 : ~ ?z <= ?x |- _ => + case id; apply Rle_trans with z; auto with real + end. +Qed. + +Lemma Rminmax : forall a b, Rmin a b <= Rmax a b. +Proof. +intros a b; destruct (Rle_dec a b). + rewrite Rmin_left, Rmax_right; assumption. +now rewrite Rmin_right, Rmax_left; assumption || + apply Rlt_le, Rnot_le_gt. +Qed. + +Lemma Rmin_assoc : forall x y z, Rmin x (Rmin y z) = + Rmin (Rmin x y) z. +Proof. +intros a b c. +unfold Rmin; destruct (Rle_dec b c); destruct (Rle_dec a b); + destruct (Rle_dec a c); destruct (Rle_dec b c); auto with real; + match goal with + | id : ~ ?x <= ?y, id2 : ?x <= ?z |- _ => + case id; apply Rle_trans with z; auto with real + | id : ~ ?x <= ?y, id2 : ~ ?z <= ?x |- _ => + case id; apply Rle_trans with z; auto with real + end. +Qed. + diff --git a/theories/Reals/Rcomplete.v b/theories/Reals/Rcomplete.v index 9b896bdd..1766f377 100644 --- a/theories/Reals/Rcomplete.v +++ b/theories/Reals/Rcomplete.v @@ -1,6 +1,6 @@ (************************************************************************) (* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *) +(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *) (* \VV/ **************************************************************) (* // * This file is distributed under the terms of the *) (* * GNU Lesser General Public License Version 2.1 *) @@ -27,21 +27,19 @@ Proof. intros. set (Vn := sequence_minorant Un (cauchy_min Un H)). set (Wn := sequence_majorant Un (cauchy_maj Un H)). - assert (H0 := maj_cv Un H). - fold Wn in H0. - assert (H1 := min_cv Un H). - fold Vn in H1. - elim H0; intros. - elim H1; intros. + pose proof (maj_cv Un H) as (x,p). + fold Wn in p. + pose proof (min_cv Un H) as (x0,p0). + fold Vn in p0. cut (x = x0). - intros. + intros H2. exists x. rewrite <- H2 in p0. unfold Un_cv. intros. unfold Un_cv in p; unfold Un_cv in p0. cut (0 < eps / 3). - intro. + intro H4. elim (p (eps / 3) H4); intros. elim (p0 (eps / 3) H4); intros. exists (max x1 x2). @@ -83,20 +81,20 @@ Proof. [ apply Rabs_triang | ring ]. apply Rlt_le_trans with (eps / 3 + eps / 3 + eps / 3). repeat apply Rplus_lt_compat. - unfold R_dist in H5. - apply H5. + unfold R_dist in H1. + apply H1. unfold ge; apply le_trans with (max x1 x2). apply le_max_l. assumption. rewrite <- Rabs_Ropp. replace (- (x - Vn n)) with (Vn n - x); [ idtac | ring ]. - unfold R_dist in H6. - apply H6. + unfold R_dist in H3. + apply H3. unfold ge; apply le_trans with (max x1 x2). apply le_max_r. assumption. - unfold R_dist in H6. - apply H6. + unfold R_dist in H3. + apply H3. unfold ge; apply le_trans with (max x1 x2). apply le_max_r. assumption. @@ -112,11 +110,11 @@ Proof. intro. unfold Un_cv in p; unfold Un_cv in p0. unfold R_dist in p; unfold R_dist in p0. - elim (p (eps / 5) H3); intros N1 H4. - elim (p0 (eps / 5) H3); intros N2 H5. + elim (p (eps / 5) H1); intros N1 H4. + elim (p0 (eps / 5) H1); intros N2 H5. unfold Cauchy_crit in H. unfold R_dist in H. - elim (H (eps / 5) H3); intros N3 H6. + elim (H (eps / 5) H1); intros N3 H6. set (N := max (max N1 N2) N3). apply Rle_lt_trans with (Rabs (x - Wn N) + Rabs (Wn N - x0)). replace (x - x0) with (x - Wn N + (Wn N - x0)); [ apply Rabs_triang | ring ]. @@ -146,12 +144,11 @@ Proof. cut (Vn N = minorant (fun k:nat => Un (N + k)%nat) (min_ss Un N (cauchy_min Un H))). - intros. - rewrite <- H9; rewrite <- H10. - rewrite <- H9 in H8. - rewrite <- H10 in H7. - elim (H7 (eps / 5) H3); intros k2 H11. - elim (H8 (eps / 5) H3); intros k1 H12. + intros H9 H10. + rewrite <- H9 in H8 |- *. + rewrite <- H10 in H7 |- *. + elim (H7 (eps / 5) H1); intros k2 H11. + elim (H8 (eps / 5) H1); intros k1 H12. apply Rle_lt_trans with (Rabs (Wn N - Un (N + k2)%nat) + Rabs (Un (N + k2)%nat - Vn N)). replace (Wn N - Vn N) with diff --git a/theories/Reals/Rdefinitions.v b/theories/Reals/Rdefinitions.v index 19cc2166..50eb59b2 100644 --- a/theories/Reals/Rdefinitions.v +++ b/theories/Reals/Rdefinitions.v @@ -1,6 +1,6 @@ (************************************************************************) (* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *) +(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *) (* \VV/ **************************************************************) (* // * This file is distributed under the terms of the *) (* * GNU Lesser General Public License Version 2.1 *) diff --git a/theories/Reals/Rderiv.v b/theories/Reals/Rderiv.v index 64b1b0d4..3a332d21 100644 --- a/theories/Reals/Rderiv.v +++ b/theories/Reals/Rderiv.v @@ -1,6 +1,6 @@ (************************************************************************) (* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *) +(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *) (* \VV/ **************************************************************) (* // * This file is distributed under the terms of the *) (* * GNU Lesser General Public License Version 2.1 *) @@ -162,9 +162,9 @@ Proof. (Rplus_lt_compat (Rabs (d x0 * (x1 - x0))) (eps * / 2) (Rabs (x1 - x0) * eps) (eps * / 2) H5 H3); intro; rewrite eps2 in H10; assumption. - unfold Rabs; case (Rcase_abs 2); auto. - intro; cut (0 < 2). - intro ; elim (Rlt_asym 0 2 H7 r). + unfold Rabs; destruct (Rcase_abs 2) as [Hlt|Hge]; auto. + cut (0 < 2). + intro H7; elim (Rlt_asym 0 2 H7 Hlt). fourier. apply Rabs_no_R0. discrR. @@ -193,11 +193,11 @@ Proof. unfold limit_in; intros; simpl; split with eps; split; auto. intros; elim H0; clear H0; intros; unfold D_x in H0; elim H0; intros; - rewrite (Rinv_r (x - x0) (Rminus_eq_contra x x0 (not_eq_sym H3))); - unfold R_dist; rewrite (Rminus_diag_eq 1 1 (eq_refl 1)); - unfold Rabs; case (Rcase_abs 0); intro. + rewrite (Rinv_r (x - x0) (Rminus_eq_contra x x0 (sym_not_eq H3))); + unfold R_dist; rewrite (Rminus_diag_eq 1 1 (refl_equal 1)); + unfold Rabs; case (Rcase_abs 0) as [Hlt|Hge]. absurd (0 < 0); auto. - red; intro; apply (Rlt_irrefl 0 r). + red in |- *; intro; apply (Rlt_irrefl 0 Hlt). unfold Rgt in H; assumption. Qed. diff --git a/theories/Reals/Reals.v b/theories/Reals/Reals.v index 8faa4e25..9cb8a10b 100644 --- a/theories/Reals/Reals.v +++ b/theories/Reals/Reals.v @@ -1,6 +1,6 @@ (************************************************************************) (* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *) +(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *) (* \VV/ **************************************************************) (* // * This file is distributed under the terms of the *) (* * GNU Lesser General Public License Version 2.1 *) diff --git a/theories/Reals/Rfunctions.v b/theories/Reals/Rfunctions.v index ee8988d8..1c353803 100644 --- a/theories/Reals/Rfunctions.v +++ b/theories/Reals/Rfunctions.v @@ -1,6 +1,6 @@ (************************************************************************) (* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *) +(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *) (* \VV/ **************************************************************) (* // * This file is distributed under the terms of the *) (* * GNU Lesser General Public License Version 2.1 *) @@ -489,16 +489,16 @@ Lemma pow_Rabs : forall (x:R) (n:nat), x ^ n <= Rabs x ^ n. Proof. intros; induction n as [| n Hrecn]. right; reflexivity. - simpl; case (Rcase_abs x); intro. + simpl; destruct (Rcase_abs x) as [Hlt|Hle]. apply Rle_trans with (Rabs (x * x ^ n)). apply RRle_abs. rewrite Rabs_mult. apply Rmult_le_compat_l. apply Rabs_pos. - right; symmetry ; apply RPow_abs. - pattern (Rabs x) at 1; rewrite (Rabs_right x r); + right; symmetry; apply RPow_abs. + pattern (Rabs x) at 1; rewrite (Rabs_right x Hle); apply Rmult_le_compat_l. - apply Rge_le; exact r. + apply Rge_le; exact Hle. apply Hrecn. Qed. @@ -520,14 +520,17 @@ Proof. apply Rle_trans with (Rabs y); [ apply Rabs_pos | exact H ]. Qed. +Lemma Rsqr_pow2 : forall x, Rsqr x = x ^ 2. +Proof. +intros; unfold Rsqr; simpl; rewrite Rmult_1_r; reflexivity. +Qed. + + (*******************************) (** * PowerRZ *) (*******************************) (*i Due to L.Thery i*) -Ltac case_eq name := - generalize (eq_refl name); pattern name at -1; case name. - Definition powerRZ (x:R) (n:Z) := match n with | Z0 => 1 @@ -744,10 +747,10 @@ Qed. Lemma R_dist_sym : forall x y:R, R_dist x y = R_dist y x. Proof. unfold R_dist; intros; split_Rabs; try ring. - generalize (Ropp_gt_lt_0_contravar (y - x) r); intro; - rewrite (Ropp_minus_distr y x) in H; generalize (Rlt_asym (x - y) 0 r0); + generalize (Ropp_gt_lt_0_contravar (y - x) Hlt0); intro; + rewrite (Ropp_minus_distr y x) in H; generalize (Rlt_asym (x - y) 0 Hlt); intro; unfold Rgt in H; exfalso; auto. - generalize (minus_Rge y x r); intro; generalize (minus_Rge x y r0); intro; + generalize (minus_Rge y x Hge0); intro; generalize (minus_Rge x y Hge); intro; generalize (Rge_antisym x y H0 H); intro; rewrite H1; ring. Qed. @@ -786,6 +789,13 @@ Proof. ring. Qed. +Lemma R_dist_mult_l : forall a b c, + R_dist (a * b) (a * c) = Rabs a * R_dist b c. +Proof. +unfold R_dist. +intros a b c; rewrite <- Rmult_minus_distr_l, Rabs_mult; reflexivity. +Qed. + (*******************************) (** * Infinite Sum *) (*******************************) diff --git a/theories/Reals/Rgeom.v b/theories/Reals/Rgeom.v index afdf148e..d930c5aa 100644 --- a/theories/Reals/Rgeom.v +++ b/theories/Reals/Rgeom.v @@ -1,6 +1,6 @@ (************************************************************************) (* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *) +(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *) (* \VV/ **************************************************************) (* // * This file is distributed under the terms of the *) (* * GNU Lesser General Public License Version 2.1 *) diff --git a/theories/Reals/RiemannInt.v b/theories/Reals/RiemannInt.v index ce37fcba..856fff80 100644 --- a/theories/Reals/RiemannInt.v +++ b/theories/Reals/RiemannInt.v @@ -1,7 +1,7 @@ (* -*- coding: utf-8 -*- *) (************************************************************************) (* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *) +(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *) (* \VV/ **************************************************************) (* // * This file is distributed under the terms of the *) (* * GNU Lesser General Public License Version 2.1 *) @@ -12,8 +12,6 @@ Require Import SeqSeries. Require Import Ranalysis_reg. Require Import Rbase. Require Import RiemannInt_SF. -Require Import Classical_Prop. -Require Import Classical_Pred_Type. Require Import Max. Local Open Scope R_scope. @@ -51,8 +49,8 @@ Lemma RiemannInt_P1 : forall (f:R -> R) (a b:R), Riemann_integrable f a b -> Riemann_integrable f b a. Proof. - unfold Riemann_integrable; intros; elim (X eps); clear X; intros; - elim p; clear p; intros; exists (mkStepFun (StepFun_P6 (pre x))); + unfold Riemann_integrable; intros; elim (X eps); clear X; intros. + elim p; clear p; intros x0 p; exists (mkStepFun (StepFun_P6 (pre x))); exists (mkStepFun (StepFun_P6 (pre x0))); elim p; clear p; intros; split. intros; apply (H t); elim H1; clear H1; intros; split; @@ -110,12 +108,10 @@ Proof. replace (vn n x + -1 * vn m x) with (vn n x - f x + (f x - vn m x)); [ apply Rabs_triang | ring ]. assert (H12 : Rmin a b = a). - unfold Rmin; case (Rle_dec a b); intro; - [ reflexivity | elim n0; assumption ]. + unfold Rmin; decide (Rle_dec a b) with H0; reflexivity. assert (H13 : Rmax a b = b). - unfold Rmax; case (Rle_dec a b); intro; - [ reflexivity | elim n0; assumption ]. - rewrite <- H12 in H11; pattern b at 2 in H11; rewrite <- H13 in H11; + unfold Rmax; decide (Rle_dec a b) with H0; reflexivity. + rewrite <- H12 in H11; rewrite <- H13 in H11 at 2; rewrite Rmult_1_l; apply Rplus_le_compat. rewrite <- Rabs_Ropp; rewrite Ropp_minus_distr; apply H9. elim H11; intros; split; left; assumption. @@ -142,7 +138,7 @@ Lemma RiemannInt_P3 : Rabs (RiemannInt_SF (wn n)) < un n) -> { l:R | Un_cv (fun N:nat => RiemannInt_SF (vn N)) l }. Proof. - intros; case (Rle_dec a b); intro. + intros; destruct (Rle_dec a b) as [Hle|Hnle]. apply RiemannInt_P2 with f un wn; assumption. assert (H1 : b <= a); auto with real. set (vn' := fun n:nat => mkStepFun (StepFun_P6 (pre (vn n)))); @@ -153,49 +149,48 @@ Proof. (forall t:R, Rmin b a <= t <= Rmax b a -> Rabs (f t - vn' n t) <= wn' n t) /\ Rabs (RiemannInt_SF (wn' n)) < un n). - intro; elim (H0 n0); intros; split. - intros; apply (H2 t); elim H4; clear H4; intros; split; + intro; elim (H0 n); intros; split. + intros t (H4,H5); apply (H2 t); split; [ apply Rle_trans with (Rmin b a); try assumption; right; unfold Rmin | apply Rle_trans with (Rmax b a); try assumption; right; unfold Rmax ]; - (case (Rle_dec a b); case (Rle_dec b a); intros; - try reflexivity || apply Rle_antisym; - [ assumption | assumption | auto with real | auto with real ]). - generalize H3; unfold RiemannInt_SF; case (Rle_dec a b); - case (Rle_dec b a); unfold wn'; intros; + decide (Rle_dec a b) with Hnle; decide (Rle_dec b a) with H1; reflexivity. + generalize H3; unfold RiemannInt_SF; destruct (Rle_dec a b) as [Hleab|Hnleab]; + destruct (Rle_dec b a) as [Hle'|Hnle']; unfold wn'; intros; (replace - (Int_SF (subdivision_val (mkStepFun (StepFun_P6 (pre (wn n0))))) - (subdivision (mkStepFun (StepFun_P6 (pre (wn n0)))))) with - (Int_SF (subdivision_val (wn n0)) (subdivision (wn n0))); + (Int_SF (subdivision_val (mkStepFun (StepFun_P6 (pre (wn n))))) + (subdivision (mkStepFun (StepFun_P6 (pre (wn n)))))) with + (Int_SF (subdivision_val (wn n)) (subdivision (wn n))); [ idtac - | apply StepFun_P17 with (fe (wn n0)) a b; + | apply StepFun_P17 with (fe (wn n)) a b; [ apply StepFun_P1 | apply StepFun_P2; - apply (StepFun_P1 (mkStepFun (StepFun_P6 (pre (wn n0))))) ] ]). + apply (StepFun_P1 (mkStepFun (StepFun_P6 (pre (wn n))))) ] ]). apply H4. rewrite Rabs_Ropp; apply H4. rewrite Rabs_Ropp in H4; apply H4. apply H4. - assert (H3 := RiemannInt_P2 _ _ _ _ H H1 H2); elim H3; intros; + assert (H3 := RiemannInt_P2 _ _ _ _ H H1 H2); elim H3; intros x p; exists (- x); unfold Un_cv; unfold Un_cv in p; intros; elim (p _ H4); intros; exists x0; intros; generalize (H5 _ H6); unfold R_dist, RiemannInt_SF; - case (Rle_dec b a); case (Rle_dec a b); intros. - elim n; assumption. + destruct (Rle_dec b a) as [Hle'|Hnle']; destruct (Rle_dec a b) as [Hle''|Hnle'']; + intros. + elim Hnle; assumption. unfold vn' in H7; - replace (Int_SF (subdivision_val (vn n0)) (subdivision (vn n0))) with - (Int_SF (subdivision_val (mkStepFun (StepFun_P6 (pre (vn n0))))) - (subdivision (mkStepFun (StepFun_P6 (pre (vn n0)))))); + replace (Int_SF (subdivision_val (vn n)) (subdivision (vn n))) with + (Int_SF (subdivision_val (mkStepFun (StepFun_P6 (pre (vn n))))) + (subdivision (mkStepFun (StepFun_P6 (pre (vn n)))))); [ unfold Rminus; rewrite Ropp_involutive; rewrite <- Rabs_Ropp; rewrite Ropp_plus_distr; rewrite Ropp_involutive; apply H7 - | symmetry ; apply StepFun_P17 with (fe (vn n0)) a b; + | symmetry ; apply StepFun_P17 with (fe (vn n)) a b; [ apply StepFun_P1 | apply StepFun_P2; - apply (StepFun_P1 (mkStepFun (StepFun_P6 (pre (vn n0))))) ] ]. - elim n1; assumption. - elim n2; assumption. + apply (StepFun_P1 (mkStepFun (StepFun_P6 (pre (vn n))))) ] ]. + elim Hnle'; assumption. + elim Hnle'; assumption. Qed. Lemma RiemannInt_exists : @@ -244,7 +239,7 @@ Proof. (RiemannInt_SF (phi_sequence vn pr2 n) + -1 * RiemannInt_SF (phi_sequence un pr1 n)); [ idtac | ring ]; rewrite <- StepFun_P30. - case (Rle_dec a b); intro. + destruct (Rle_dec a b) as [Hle|Hnle]. apply Rle_lt_trans with (RiemannInt_SF (mkStepFun @@ -263,13 +258,11 @@ Proof. (phi_sequence vn pr2 n x - f x + (f x - phi_sequence un pr1 n x)); [ apply Rabs_triang | ring ]. assert (H10 : Rmin a b = a). - unfold Rmin; case (Rle_dec a b); intro; - [ reflexivity | elim n0; assumption ]. + unfold Rmin; decide (Rle_dec a b) with Hle; reflexivity. assert (H11 : Rmax a b = b). - unfold Rmax; case (Rle_dec a b); intro; - [ reflexivity | elim n0; assumption ]. + unfold Rmax; decide (Rle_dec a b) with Hle; reflexivity. rewrite (Rplus_comm (psi_un x)); apply Rplus_le_compat. - rewrite <- Rabs_Ropp; rewrite Ropp_minus_distr; elim H5; intros; apply H8. + rewrite <- Rabs_Ropp; rewrite Ropp_minus_distr; destruct H5 as (H8,H9); apply H8. rewrite H10; rewrite H11; elim H7; intros; split; left; assumption. elim H6; intros; apply H8. rewrite H10; rewrite H11; elim H7; intros; split; left; assumption. @@ -319,11 +312,9 @@ Proof. (phi_sequence vn pr2 n x - f x + (f x - phi_sequence un pr1 n x)); [ apply Rabs_triang | ring ]. assert (H10 : Rmin a b = b). - unfold Rmin; case (Rle_dec a b); intro; - [ elim n0; assumption | reflexivity ]. + unfold Rmin; decide (Rle_dec a b) with Hnle; reflexivity. assert (H11 : Rmax a b = a). - unfold Rmax; case (Rle_dec a b); intro; - [ elim n0; assumption | reflexivity ]. + unfold Rmax; decide (Rle_dec a b) with Hnle; reflexivity. apply Rplus_le_compat. rewrite <- Rabs_Ropp; rewrite Ropp_minus_distr; elim H5; intros; apply H8. rewrite H10; rewrite H11; elim H7; intros; split; left; assumption. @@ -388,11 +379,9 @@ Proof. [ idtac | left; change (0 < / (INR n + 1)); apply Rinv_0_lt_compat; assumption ]; apply Rle_lt_trans with (/ (INR x + 1)). - apply Rle_Rinv. + apply Rinv_le_contravar. apply Rplus_le_lt_0_compat; [ apply pos_INR | apply Rlt_0_1 ]. - assumption. - do 2 rewrite <- (Rplus_comm 1); apply Rplus_le_compat_l; apply le_INR; - apply H4. + apply Rplus_le_compat_r; apply le_INR; apply H4. rewrite <- (Rinv_involutive eps). apply Rinv_lt_contravar. apply Rmult_lt_0_compat. @@ -405,6 +394,15 @@ Proof. red; intro; rewrite H6 in H; elim (Rlt_irrefl _ H). Qed. +Lemma Riemann_integrable_ext : + forall f g a b, + (forall x, Rmin a b <= x <= Rmax a b -> f x = g x) -> + Riemann_integrable f a b -> Riemann_integrable g a b. +intros f g a b fg rif eps; destruct (rif eps) as [phi [psi [P1 P2]]]. +exists phi; exists psi;split;[ | assumption ]. +intros t intt; rewrite <- fg;[ | assumption]. +apply P1; assumption. +Qed. (**********) Definition RiemannInt (f:R -> R) (a b:R) (pr:Riemann_integrable f a b) : R := let (a,_) := RiemannInt_exists pr RinvN RinvN_cv in a. @@ -414,10 +412,10 @@ Lemma RiemannInt_P5 : RiemannInt pr1 = RiemannInt pr2. Proof. intros; unfold RiemannInt; - case (RiemannInt_exists pr1 RinvN RinvN_cv); - case (RiemannInt_exists pr2 RinvN RinvN_cv); intros; + case (RiemannInt_exists pr1 RinvN RinvN_cv) as (x,HUn); + case (RiemannInt_exists pr2 RinvN RinvN_cv) as (x0,HUn0); eapply UL_sequence; - [ apply u0 + [ apply HUn | apply RiemannInt_P4 with pr2 RinvN; apply RinvN_cv || assumption ]. Qed. @@ -434,14 +432,13 @@ Proof. exists 0%nat; unfold I; rewrite Rmult_0_l; rewrite Rplus_0_r; assumption. cut (Nbound I). - intro; assert (H2 := Nzorn H0 H1); elim H2; intros; exists x; elim p; intros; + intro; assert (H2 := Nzorn H0 H1); elim H2; intros x p; exists x; elim p; intros; split. apply H3. - case (total_order_T (a + INR (S x) * del) b); intro. - elim s; intro. - assert (H5 := H4 (S x) a0); elim (le_Sn_n _ H5). + destruct (total_order_T (a + INR (S x) * del) b) as [[Hlt|Heq]|Hgt]. + assert (H5 := H4 (S x) Hlt); elim (le_Sn_n _ H5). right; symmetry ; assumption. - left; apply r. + left; apply Hgt. assert (H1 : 0 <= (b - a) / del). unfold Rdiv; apply Rmult_le_pos; [ apply Rge_le; apply Rge_minus; apply Rle_ge; left; apply H @@ -509,22 +506,24 @@ Proof. | apply Rmin_r ] | intros; apply H3; try assumption; apply Rlt_le_trans with (Rmin x (b - a)); [ assumption | apply Rmin_l ] ]. - assert (H3 := completeness E H1 H2); elim H3; intros; cut (0 < x <= b - a). + assert (H3 := completeness E H1 H2); elim H3; intros x p; cut (0 < x <= b - a). intro; elim H4; clear H4; intros; exists (mkposreal _ H4); split. apply H5. unfold is_lub in p; elim p; intros; unfold is_upper_bound in H6; - set (D := Rabs (x0 - y)); elim (classic (exists y : R, D < y /\ E y)); - intro. + set (D := Rabs (x0 - y)). + assert (H11: ((exists y : R, D < y /\ E y) \/ (forall y : R, not (D < y /\ E y)) -> False) -> False). + clear; intros H; apply H. + right; intros y0 H0; apply H. + left; now exists y0. + apply Rnot_le_lt; intros H30. + apply H11; clear H11; intros H11. + revert H30; apply Rlt_not_le. + destruct H11 as [H11|H12]. elim H11; intros; elim H12; clear H12; intros; unfold E in H13; elim H13; intros; apply H15; assumption. - assert (H12 := not_ex_all_not _ (fun y:R => D < y /\ E y) H11); - assert (H13 : is_upper_bound E D). + assert (H13 : is_upper_bound E D). unfold is_upper_bound; intros; assert (H14 := H12 x1); - elim (not_and_or (D < x1) (E x1) H14); intro. - case (Rle_dec x1 D); intro. - assumption. - elim H15; auto with real. - elim H15; assumption. + apply Rnot_lt_le; contradict H14; now split. assert (H14 := H7 _ H13); elim (Rlt_irrefl _ (Rle_lt_trans _ _ _ H14 H10)). unfold is_lub in p; unfold is_upper_bound in p; elim p; clear p; intros; split. @@ -544,17 +543,16 @@ Lemma Heine_cor2 : a <= x <= b -> a <= y <= b -> Rabs (x - y) < delta -> Rabs (f x - f y) < eps }. Proof. - intro f; intros; case (total_order_T a b); intro. - elim s; intro. - assert (H0 := Heine_cor1 a0 H eps); elim H0; intros; exists x; + intro f; intros; destruct (total_order_T a b) as [[Hlt|Heq]|Hgt]. + assert (H0 := Heine_cor1 Hlt H eps); elim H0; intros x p; exists x; elim p; intros; apply H2; assumption. exists (mkposreal _ Rlt_0_1); intros; assert (H3 : x = y); - [ elim H0; elim H1; intros; rewrite b0 in H3; rewrite b0 in H5; + [ elim H0; elim H1; intros; rewrite Heq in H3, H5; apply Rle_antisym; apply Rle_trans with b; assumption | rewrite H3; unfold Rminus; rewrite Rplus_opp_r; rewrite Rabs_R0; apply (cond_pos eps) ]. exists (mkposreal _ Rlt_0_1); intros; elim H0; intros; - elim (Rlt_irrefl _ (Rle_lt_trans _ _ _ (Rle_trans _ _ _ H3 H4) r)). + elim (Rlt_irrefl _ (Rle_lt_trans _ _ _ (Rle_trans _ _ _ H3 H4) Hgt)). Qed. Lemma SubEqui_P1 : @@ -567,7 +565,7 @@ Lemma SubEqui_P2 : forall (a b:R) (del:posreal) (h:a < b), pos_Rl (SubEqui del h) (pred (Rlength (SubEqui del h))) = b. Proof. - intros; unfold SubEqui; case (maxN del h); intros; clear a0; + intros; unfold SubEqui; destruct (maxN del h)as (x,_). cut (forall (x:nat) (a:R) (del:posreal), pos_Rl (SubEquiN (S x) a b del) @@ -623,8 +621,8 @@ Proof. simpl in H; inversion H. rewrite (SubEqui_P6 del h (i:=(max_N del h))). replace (S (max_N del h)) with (pred (Rlength (SubEqui del h))). - rewrite SubEqui_P2; unfold max_N; case (maxN del h); intros; left; - elim a0; intros; assumption. + rewrite SubEqui_P2; unfold max_N; case (maxN del h) as (?&?&?); left; + assumption. rewrite SubEqui_P5; reflexivity. apply lt_n_Sn. repeat rewrite SubEqui_P6. @@ -678,11 +676,11 @@ Proof. | apply Rinv_0_lt_compat; apply Rmult_lt_0_compat; [ prove_sup0 | apply Rlt_Rminus; assumption ] ]. assert (H2 : Rmin a b = a). - unfold Rmin; case (Rle_dec a b); intro; - [ reflexivity | elim n; left; assumption ]. + apply Rlt_le in H. + unfold Rmin; decide (Rle_dec a b) with H; reflexivity. assert (H3 : Rmax a b = b). - unfold Rmax; case (Rle_dec a b); intro; - [ reflexivity | elim n; left; assumption ]. + apply Rlt_le in H. + unfold Rmax; decide (Rle_dec a b) with H; reflexivity. elim (Heine_cor2 H0 (mkposreal _ H1)); intros del H4; elim (SubEqui_P9 del f H); intros phi [H5 H6]; split with phi; split with (mkStepFun (StepFun_P4 a b (eps / (2 * (b - a))))); @@ -727,7 +725,7 @@ Proof. elim (lt_n_O _ H9). unfold co_interval in H10; elim H10; clear H10; intros; rewrite Rabs_right. rewrite SubEqui_P5 in H9; simpl in H9; inversion H9. - apply Rplus_lt_reg_r with (pos_Rl (SubEqui del H) (max_N del H)). + apply Rplus_lt_reg_l with (pos_Rl (SubEqui del H) (max_N del H)). replace (pos_Rl (SubEqui del H) (max_N del H) + (t - pos_Rl (SubEqui del H) (max_N del H))) with t; @@ -738,10 +736,10 @@ Proof. rewrite H13 in H12; rewrite SubEqui_P2 in H12; apply H12. rewrite SubEqui_P6. 2: apply lt_n_Sn. - unfold max_N; case (maxN del H); intros; elim a0; clear a0; - intros _ H13; replace (a + INR x * del + del) with (a + INR (S x) * del); + unfold max_N; destruct (maxN del H) as (?&?&H13); + replace (a + INR x * del + del) with (a + INR (S x) * del); [ assumption | rewrite S_INR; ring ]. - apply Rplus_lt_reg_r with (pos_Rl (SubEqui del H) I); + apply Rplus_lt_reg_l with (pos_Rl (SubEqui del H) I); replace (pos_Rl (SubEqui del H) I + (t - pos_Rl (SubEqui del H) I)) with t; [ idtac | ring ]; replace (pos_Rl (SubEqui del H) I + del) with (pos_Rl (SubEqui del H) (S I)). @@ -759,7 +757,7 @@ Proof. intros; assumption. assert (H4 : Nbound I). unfold Nbound; exists (S (max_N del H)); intros; unfold max_N; - case (maxN del H); intros; elim a0; clear a0; intros _ H5; + destruct (maxN del H) as (?&_&H5); apply INR_le; apply Rmult_le_reg_l with (pos del). apply (cond_pos del). apply Rplus_le_reg_l with a; do 2 rewrite (Rmult_comm del); @@ -767,12 +765,12 @@ Proof. apply Rle_trans with b; try assumption; elim H8; intros; assumption. elim (Nzorn H1 H4); intros N [H5 H6]; assert (H7 : (N < S (max_N del H))%nat). - unfold max_N; case (maxN del H); intros; apply INR_lt; + unfold max_N; case (maxN del H) as (?&?&?); apply INR_lt; apply Rmult_lt_reg_l with (pos del). apply (cond_pos del). - apply Rplus_lt_reg_r with a; do 2 rewrite (Rmult_comm del); + apply Rplus_lt_reg_l with a; do 2 rewrite (Rmult_comm del); apply Rle_lt_trans with t0; unfold I in H5; try assumption; - elim a0; intros; apply Rlt_le_trans with b; try assumption; + apply Rlt_le_trans with b; try assumption; elim H8; intros. elim H11; intro. assumption. @@ -791,8 +789,8 @@ Proof. elim H0; assumption. rewrite SubEqui_P5; reflexivity. rewrite SubEqui_P6. - case (Rle_dec (a + INR (S N) * del) t0); intro. - assert (H11 := H6 (S N) r); elim (le_Sn_n _ H11). + destruct (Rle_dec (a + INR (S N) * del) t0) as [Hle|Hnle]. + assert (H11 := H6 (S N) Hle); elim (le_Sn_n _ H11). auto with real. apply le_lt_n_Sm; assumption. Qed. @@ -805,8 +803,8 @@ Proof. intros; simpl; unfold fct_cte; replace t with a. unfold Rminus; rewrite Rplus_opp_r; rewrite Rabs_R0; right; reflexivity. - generalize H; unfold Rmin, Rmax; case (Rle_dec a a); intros; elim H0; - intros; apply Rle_antisym; assumption. + generalize H; unfold Rmin, Rmax; decide (Rle_dec a a) with (Rle_refl a). + intros (?,?); apply Rle_antisym; assumption. rewrite StepFun_P18; rewrite Rmult_0_l; rewrite Rabs_R0; apply (cond_pos eps). Qed. @@ -815,10 +813,9 @@ Lemma continuity_implies_RiemannInt : a <= b -> (forall x:R, a <= x <= b -> continuity_pt f x) -> Riemann_integrable f a b. Proof. - intros; case (total_order_T a b); intro; - [ elim s; intro; - [ apply RiemannInt_P6; assumption | rewrite b0; apply RiemannInt_P7 ] - | elim (Rlt_irrefl _ (Rle_lt_trans _ _ _ H r)) ]. + intros; destruct (total_order_T a b) as [[Hlt| -> ]|Hgt]; + [ apply RiemannInt_P6; assumption | apply RiemannInt_P7 + | elim (Rlt_irrefl _ (Rle_lt_trans _ _ _ H Hgt)) ]. Qed. Lemma RiemannInt_P8 : @@ -826,9 +823,9 @@ Lemma RiemannInt_P8 : (pr2:Riemann_integrable f b a), RiemannInt pr1 = - RiemannInt pr2. Proof. intro f; intros; eapply UL_sequence. - unfold RiemannInt; case (RiemannInt_exists pr1 RinvN RinvN_cv); - intros; apply u. - unfold RiemannInt; case (RiemannInt_exists pr2 RinvN RinvN_cv); + unfold RiemannInt; destruct (RiemannInt_exists pr1 RinvN RinvN_cv) as (?,HUn); + apply HUn. + unfold RiemannInt; destruct (RiemannInt_exists pr2 RinvN RinvN_cv) as (?,HUn); intros; cut (exists psi1 : nat -> StepFun a b, @@ -857,7 +854,7 @@ Proof. [ assumption | unfold Rminus; rewrite Ropp_0; rewrite Rplus_0_r; apply Rabs_right; left; apply (cond_pos (RinvN n)) ]. - clear H1; unfold Un_cv in u; elim (u _ H3); clear u; intros N1 H1; + clear H1; destruct (HUn _ H3) as (N1,H1); exists (max N0 N1); intros; unfold R_dist; apply Rle_lt_trans with (Rabs @@ -881,7 +878,7 @@ Proof. -1 * RiemannInt_SF (mkStepFun (StepFun_P6 (pre (phi_sequence RinvN pr2 n))))); [ idtac | ring ]; rewrite <- StepFun_P30. - case (Rle_dec a b); intro. + destruct (Rle_dec a b) as [Hle|Hnle]. apply Rle_lt_trans with (RiemannInt_SF (mkStepFun @@ -903,11 +900,9 @@ Proof. (phi_sequence RinvN pr1 n x0 - f x0 + (f x0 - phi_sequence RinvN pr2 n x0)); [ apply Rabs_triang | ring ]. assert (H7 : Rmin a b = a). - unfold Rmin; case (Rle_dec a b); intro; - [ reflexivity | elim n0; assumption ]. + unfold Rmin; decide (Rle_dec a b) with Hle; reflexivity. assert (H8 : Rmax a b = b). - unfold Rmax; case (Rle_dec a b); intro; - [ reflexivity | elim n0; assumption ]. + unfold Rmax; decide (Rle_dec a b) with Hle; reflexivity. apply Rplus_le_compat. elim (H0 n); intros; rewrite <- Rabs_Ropp; rewrite Ropp_minus_distr; apply H9; rewrite H7; rewrite H8. @@ -956,11 +951,9 @@ Proof. (phi_sequence RinvN pr1 n x0 - f x0 + (f x0 - phi_sequence RinvN pr2 n x0)); [ apply Rabs_triang | ring ]. assert (H7 : Rmin a b = b). - unfold Rmin; case (Rle_dec a b); intro; - [ elim n0; assumption | reflexivity ]. + unfold Rmin; decide (Rle_dec a b) with Hnle; reflexivity. assert (H8 : Rmax a b = a). - unfold Rmax; case (Rle_dec a b); intro; - [ elim n0; assumption | reflexivity ]. + unfold Rmax; decide (Rle_dec a b) with Hnle; reflexivity. apply Rplus_le_compat. elim (H0 n); intros; rewrite <- Rabs_Ropp; rewrite Ropp_minus_distr; apply H9; rewrite H7; rewrite H8. @@ -1007,15 +1000,6 @@ Proof. | discrR ]. Qed. -Lemma Req_EM_T : forall r1 r2:R, {r1 = r2} + {r1 <> r2}. -Proof. - intros; elim (total_order_T r1 r2); intros; - [ elim a; intro; - [ right; red; intro; rewrite H in a0; elim (Rlt_irrefl r2 a0) - | left; assumption ] - | right; red; intro; rewrite H in b; elim (Rlt_irrefl r2 b) ]. -Qed. - (* L1([a,b]) is a vectorial space *) Lemma RiemannInt_P10 : forall (f g:R -> R) (a b l:R), @@ -1023,10 +1007,9 @@ Lemma RiemannInt_P10 : Riemann_integrable g a b -> Riemann_integrable (fun x:R => f x + l * g x) a b. Proof. - unfold Riemann_integrable; intros f g; intros; case (Req_EM_T l 0); - intro. - elim (X eps); intros; split with x; elim p; intros; split with x0; elim p0; - intros; split; try assumption; rewrite e; intros; + unfold Riemann_integrable; intros f g; intros; destruct (Req_EM_T l 0) as [Heq|Hneq]. + elim (X eps); intros x p; split with x; elim p; intros x0 p0; split with x0; elim p0; + intros; split; try assumption; rewrite Heq; intros; rewrite Rmult_0_l; rewrite Rplus_0_r; apply H; assumption. assert (H : 0 < eps / 2). unfold Rdiv; apply Rmult_lt_0_compat; @@ -1036,9 +1019,9 @@ Proof. [ apply (cond_pos eps) | apply Rinv_0_lt_compat; apply Rmult_lt_0_compat; [ prove_sup0 | apply Rabs_pos_lt; assumption ] ]. - elim (X (mkposreal _ H)); intros; elim (X0 (mkposreal _ H0)); intros; + elim (X (mkposreal _ H)); intros x p; elim (X0 (mkposreal _ H0)); intros x0 p0; split with (mkStepFun (StepFun_P28 l x x0)); elim p0; - elim p; intros; split with (mkStepFun (StepFun_P28 (Rabs l) x1 x2)); + elim p; intros x1 p1 x2 p2. split with (mkStepFun (StepFun_P28 (Rabs l) x1 x2)); elim p1; elim p2; clear p1 p2 p0 p X X0; intros; split. intros; simpl; apply Rle_trans with (Rabs (f t - x t) + Rabs (l * (g t - x0 t))). @@ -1113,18 +1096,14 @@ Proof. rewrite (Rplus_comm (psi1 n x)); apply Rplus_le_compat. rewrite <- Rabs_Ropp; rewrite Ropp_minus_distr; elim (H1 n); intros; apply H7. assert (H10 : Rmin a b = a). - unfold Rmin; case (Rle_dec a b); intro; - [ reflexivity | elim n0; assumption ]. + unfold Rmin; decide (Rle_dec a b) with Hyp; reflexivity. assert (H11 : Rmax a b = b). - unfold Rmax; case (Rle_dec a b); intro; - [ reflexivity | elim n0; assumption ]. + unfold Rmax; decide (Rle_dec a b) with Hyp; reflexivity. rewrite H10; rewrite H11; elim H6; intros; split; left; assumption. elim (H0 n); intros; apply H7; assert (H10 : Rmin a b = a). - unfold Rmin; case (Rle_dec a b); intro; - [ reflexivity | elim n0; assumption ]. + unfold Rmin; decide (Rle_dec a b) with Hyp; reflexivity. assert (H11 : Rmax a b = b). - unfold Rmax; case (Rle_dec a b); intro; - [ reflexivity | elim n0; assumption ]. + unfold Rmax; decide (Rle_dec a b) with Hyp; reflexivity. rewrite H10; rewrite H11; elim H6; intros; split; left; assumption. rewrite StepFun_P30; rewrite Rmult_1_l; rewrite double; apply Rplus_lt_compat. apply Rlt_trans with (pos (un n)). @@ -1256,10 +1235,10 @@ Lemma RiemannInt_P12 : Proof. intro f; intros; case (Req_dec l 0); intro. pattern l at 2; rewrite H0; rewrite Rmult_0_l; rewrite Rplus_0_r; - unfold RiemannInt; case (RiemannInt_exists pr3 RinvN RinvN_cv); - case (RiemannInt_exists pr1 RinvN RinvN_cv); intros; + unfold RiemannInt; destruct (RiemannInt_exists pr3 RinvN RinvN_cv) as (?,HUn_cv); + destruct (RiemannInt_exists pr1 RinvN RinvN_cv) as (?,HUn_cv0); intros. eapply UL_sequence; - [ apply u0 + [ apply HUn_cv | set (psi1 := fun n:nat => proj1_sig (phi_sequence_prop RinvN pr1 n)); set (psi2 := fun n:nat => proj1_sig (phi_sequence_prop RinvN pr3 n)); apply RiemannInt_P11 with f RinvN (phi_sequence RinvN pr1) psi1 psi2; @@ -1278,22 +1257,22 @@ Proof. [ apply H2; assumption | rewrite H0; ring ] ] | assumption ] ]. eapply UL_sequence. - unfold RiemannInt; case (RiemannInt_exists pr3 RinvN RinvN_cv); - intros; apply u. + unfold RiemannInt; destruct (RiemannInt_exists pr3 RinvN RinvN_cv) as (?,HUn_cv); + intros; apply HUn_cv. unfold Un_cv; intros; unfold RiemannInt; - case (RiemannInt_exists pr1 RinvN RinvN_cv); - case (RiemannInt_exists pr2 RinvN RinvN_cv); unfold Un_cv; + case (RiemannInt_exists pr1 RinvN RinvN_cv) as (x0,HUn_cv0); + case (RiemannInt_exists pr2 RinvN RinvN_cv) as (x,HUn_cv); unfold Un_cv; intros; assert (H2 : 0 < eps / 5). unfold Rdiv; apply Rmult_lt_0_compat; [ assumption | apply Rinv_0_lt_compat; prove_sup0 ]. - elim (u0 _ H2); clear u0; intros N0 H3; assert (H4 := RinvN_cv); + elim (HUn_cv0 _ H2); clear HUn_cv0; intros N0 H3; assert (H4 := RinvN_cv); unfold Un_cv in H4; elim (H4 _ H2); clear H4 H2; intros N1 H4; assert (H5 : 0 < eps / (5 * Rabs l)). unfold Rdiv; apply Rmult_lt_0_compat; [ assumption | apply Rinv_0_lt_compat; apply Rmult_lt_0_compat; [ prove_sup0 | apply Rabs_pos_lt; assumption ] ]. - elim (u _ H5); clear u; intros N2 H6; assert (H7 := RinvN_cv); + elim (HUn_cv _ H5); clear HUn_cv; intros N2 H6; assert (H7 := RinvN_cv); unfold Un_cv in H7; elim (H7 _ H5); clear H7 H5; intros N3 H5; unfold R_dist in H3, H4, H5, H6; set (N := max (max N0 N1) (max N2 N3)). assert (H7 : forall n:nat, (n >= N1)%nat -> RinvN n < eps / 5). @@ -1381,11 +1360,9 @@ Proof. (RiemannInt_SF (phi_sequence RinvN pr1 n) + l * RiemannInt_SF (phi_sequence RinvN pr2 n))); [ idtac | ring ]; do 2 rewrite <- StepFun_P30; assert (H10 : Rmin a b = a). - unfold Rmin; case (Rle_dec a b); intro; - [ reflexivity | elim n0; assumption ]. + unfold Rmin; decide (Rle_dec a b) with H; reflexivity. assert (H11 : Rmax a b = b). - unfold Rmax; case (Rle_dec a b); intro; - [ reflexivity | elim n0; assumption ]. + unfold Rmax; decide (Rle_dec a b) with H; reflexivity. rewrite H10 in H7; rewrite H10 in H8; rewrite H10 in H9; rewrite H11 in H7; rewrite H11 in H8; rewrite H11 in H9; apply Rle_lt_trans with @@ -1495,7 +1472,7 @@ Lemma RiemannInt_P13 : (pr3:Riemann_integrable (fun x:R => f x + l * g x) a b), RiemannInt pr3 = RiemannInt pr1 + l * RiemannInt pr2. Proof. - intros; case (Rle_dec a b); intro; + intros; destruct (Rle_dec a b) as [Hle|Hnle]; [ apply RiemannInt_P12; assumption | assert (H : b <= a); [ auto with real @@ -1526,9 +1503,9 @@ Lemma RiemannInt_P15 : forall (a b c:R) (pr:Riemann_integrable (fct_cte c) a b), RiemannInt pr = c * (b - a). Proof. - intros; unfold RiemannInt; case (RiemannInt_exists pr RinvN RinvN_cv); + intros; unfold RiemannInt; destruct (RiemannInt_exists pr RinvN RinvN_cv) as (?,HUn_cv); intros; eapply UL_sequence. - apply u. + apply HUn_cv. set (phi1 := fun N:nat => phi_sequence RinvN pr N); change (Un_cv (fun N:nat => RiemannInt_SF (phi1 N)) (c * (b - a))); set (f := fct_cte c); @@ -1574,18 +1551,18 @@ Lemma Rle_cv_lim : forall (Un Vn:nat -> R) (l1 l2:R), (forall n:nat, Un n <= Vn n) -> Un_cv Un l1 -> Un_cv Vn l2 -> l1 <= l2. Proof. - intros; case (Rle_dec l1 l2); intro. + intros; destruct (Rle_dec l1 l2) as [Hle|Hnle]. assumption. assert (H2 : l2 < l1). auto with real. - clear n; assert (H3 : 0 < (l1 - l2) / 2). + assert (H3 : 0 < (l1 - l2) / 2). unfold Rdiv; apply Rmult_lt_0_compat; [ apply Rlt_Rminus; assumption | apply Rinv_0_lt_compat; prove_sup0 ]. elim (H1 _ H3); elim (H0 _ H3); clear H0 H1; unfold R_dist; intros; set (N := max x x0); cut (Vn N < Un N). intro; elim (Rlt_irrefl _ (Rle_lt_trans _ _ _ (H N) H4)). apply Rlt_trans with ((l1 + l2) / 2). - apply Rplus_lt_reg_r with (- l2); + apply Rplus_lt_reg_l with (- l2); replace (- l2 + (l1 + l2) / 2) with ((l1 - l2) / 2). rewrite Rplus_comm; apply Rle_lt_trans with (Rabs (Vn N - l2)). apply RRle_abs. @@ -1596,7 +1573,7 @@ Proof. repeat rewrite Rmult_assoc; rewrite <- Rinv_l_sym; [ ring | discrR ] | discrR ]. - apply Ropp_lt_cancel; apply Rplus_lt_reg_r with l1; + apply Ropp_lt_cancel; apply Rplus_lt_reg_l with l1; replace (l1 + - ((l1 + l2) / 2)) with ((l1 - l2) / 2). apply Rle_lt_trans with (Rabs (Un N - l1)). rewrite <- Rabs_Ropp; rewrite Ropp_minus_distr; apply RRle_abs. @@ -1615,9 +1592,9 @@ Lemma RiemannInt_P17 : a <= b -> Rabs (RiemannInt pr1) <= RiemannInt pr2. Proof. intro f; intros; unfold RiemannInt; - case (RiemannInt_exists pr1 RinvN RinvN_cv); - case (RiemannInt_exists pr2 RinvN RinvN_cv); intros; - set (phi1 := phi_sequence RinvN pr1) in u0; + case (RiemannInt_exists pr1 RinvN RinvN_cv) as (x0,HUn_cv0); + case (RiemannInt_exists pr2 RinvN RinvN_cv) as (x,HUn_cv); + set (phi1 := phi_sequence RinvN pr1) in HUn_cv0; set (phi2 := fun N:nat => mkStepFun (StepFun_P32 (phi1 N))); apply Rle_cv_lim with (fun N:nat => Rabs (RiemannInt_SF (phi1 N))) @@ -1672,10 +1649,10 @@ Lemma RiemannInt_P18 : (forall x:R, a < x < b -> f x = g x) -> RiemannInt pr1 = RiemannInt pr2. Proof. intro f; intros; unfold RiemannInt; - case (RiemannInt_exists pr1 RinvN RinvN_cv); - case (RiemannInt_exists pr2 RinvN RinvN_cv); intros; + case (RiemannInt_exists pr1 RinvN RinvN_cv) as (x0,HUn_cv0); + case (RiemannInt_exists pr2 RinvN RinvN_cv) as (x,HUn_cv); eapply UL_sequence. - apply u0. + apply HUn_cv0. set (phi1 := fun N:nat => phi_sequence RinvN pr1 N); change (Un_cv (fun N:nat => RiemannInt_SF (phi1 N)) x); assert @@ -1718,48 +1695,48 @@ Proof. apply RinvN_cv. intro; elim (H2 n); intros; split; try assumption. intros; unfold phi2_m; simpl; unfold phi2_aux; - case (Req_EM_T t a); case (Req_EM_T t b); intros. - rewrite e0; unfold Rminus; rewrite Rplus_opp_r; rewrite Rabs_R0; + destruct (Req_EM_T t a) as [Heqa|Hneqa]; destruct (Req_EM_T t b) as [Heqb|Hneqb]. + rewrite Heqa; unfold Rminus; rewrite Rplus_opp_r; rewrite Rabs_R0; apply Rle_trans with (Rabs (g t - phi2 n t)). apply Rabs_pos. - pattern a at 3; rewrite <- e0; apply H3; assumption. - rewrite e; unfold Rminus; rewrite Rplus_opp_r; rewrite Rabs_R0; + pattern a at 3; rewrite <- Heqa; apply H3; assumption. + rewrite Heqa; unfold Rminus; rewrite Rplus_opp_r; rewrite Rabs_R0; apply Rle_trans with (Rabs (g t - phi2 n t)). apply Rabs_pos. - pattern a at 3; rewrite <- e; apply H3; assumption. - rewrite e; unfold Rminus; rewrite Rplus_opp_r; rewrite Rabs_R0; + pattern a at 3; rewrite <- Heqa; apply H3; assumption. + rewrite Heqb; unfold Rminus; rewrite Rplus_opp_r; rewrite Rabs_R0; apply Rle_trans with (Rabs (g t - phi2 n t)). apply Rabs_pos. - pattern b at 3; rewrite <- e; apply H3; assumption. + pattern b at 3; rewrite <- Heqb; apply H3; assumption. replace (f t) with (g t). apply H3; assumption. symmetry ; apply H0; elim H5; clear H5; intros. assert (H7 : Rmin a b = a). - unfold Rmin; case (Rle_dec a b); intro; - [ reflexivity | elim n2; assumption ]. + unfold Rmin; destruct (Rle_dec a b) as [Heqab|Hneqab]; + [ reflexivity | elim Hneqab; assumption ]. assert (H8 : Rmax a b = b). - unfold Rmax; case (Rle_dec a b); intro; - [ reflexivity | elim n2; assumption ]. + unfold Rmax; destruct (Rle_dec a b) as [Heqab|Hneqab]; + [ reflexivity | elim Hneqab; assumption ]. rewrite H7 in H5; rewrite H8 in H6; split. - elim H5; intro; [ assumption | elim n1; symmetry ; assumption ]. - elim H6; intro; [ assumption | elim n0; assumption ]. + elim H5; intro; [ assumption | elim Hneqa; symmetry ; assumption ]. + elim H6; intro; [ assumption | elim Hneqb; assumption ]. cut (forall N:nat, RiemannInt_SF (phi2_m N) = RiemannInt_SF (phi2 N)). - intro; unfold Un_cv; intros; elim (u _ H4); intros; exists x1; intros; + intro; unfold Un_cv; intros; elim (HUn_cv _ H4); intros; exists x1; intros; rewrite (H3 n); apply H5; assumption. intro; apply Rle_antisym. apply StepFun_P37; try assumption. intros; unfold phi2_m; simpl; unfold phi2_aux; - case (Req_EM_T x1 a); case (Req_EM_T x1 b); intros. - elim H3; intros; rewrite e0 in H4; elim (Rlt_irrefl _ H4). - elim H3; intros; rewrite e in H4; elim (Rlt_irrefl _ H4). - elim H3; intros; rewrite e in H5; elim (Rlt_irrefl _ H5). + destruct (Req_EM_T x1 a) as [Heqa|Hneqa]; destruct (Req_EM_T x1 b) as [Heqb|Hneqb]. + elim H3; intros; rewrite Heqa in H4; elim (Rlt_irrefl _ H4). + elim H3; intros; rewrite Heqa in H4; elim (Rlt_irrefl _ H4). + elim H3; intros; rewrite Heqb in H5; elim (Rlt_irrefl _ H5). right; reflexivity. apply StepFun_P37; try assumption. intros; unfold phi2_m; simpl; unfold phi2_aux; - case (Req_EM_T x1 a); case (Req_EM_T x1 b); intros. - elim H3; intros; rewrite e0 in H4; elim (Rlt_irrefl _ H4). - elim H3; intros; rewrite e in H4; elim (Rlt_irrefl _ H4). - elim H3; intros; rewrite e in H5; elim (Rlt_irrefl _ H5). + destruct (Req_EM_T x1 a) as [ -> |Hneqa]. + elim H3; intros; elim (Rlt_irrefl _ H4). + destruct (Req_EM_T x1 b) as [ -> |Hneqb]. + elim H3; intros; elim (Rlt_irrefl _ H5). right; reflexivity. intro; assert (H2 := pre (phi2 N)); unfold IsStepFun in H2; unfold is_subdivision in H2; elim H2; clear H2; intros l [lf H2]; @@ -1775,21 +1752,19 @@ Proof. apply le_O_n. apply lt_trans with (pred (Rlength l)); [ assumption | apply lt_pred_n_n ]. apply neq_O_lt; intro; rewrite <- H12 in H6; discriminate. - unfold Rmin; case (Rle_dec a b); intro; - [ reflexivity | elim n; assumption ]. + unfold Rmin; decide (Rle_dec a b) with H; reflexivity. assert (H11 : pos_Rl l (S i) <= b). replace b with (Rmax a b). rewrite <- H4; elim (RList_P6 l); intros; apply H11. assumption. apply lt_le_S; assumption. apply lt_pred_n_n; apply neq_O_lt; intro; rewrite <- H13 in H6; discriminate. - unfold Rmax; case (Rle_dec a b); intro; - [ reflexivity | elim n; assumption ]. - elim H7; clear H7; intros; unfold phi2_aux; case (Req_EM_T x1 a); - case (Req_EM_T x1 b); intros. - rewrite e in H12; elim (Rlt_irrefl _ (Rle_lt_trans _ _ _ H11 H12)). - rewrite e in H7; elim (Rlt_irrefl _ (Rle_lt_trans _ _ _ H10 H7)). - rewrite e in H12; elim (Rlt_irrefl _ (Rle_lt_trans _ _ _ H11 H12)). + unfold Rmax; decide (Rle_dec a b) with H; reflexivity. + elim H7; clear H7; intros; unfold phi2_aux; destruct (Req_EM_T x1 a) as [Heq|Hneq]; + destruct (Req_EM_T x1 b) as [Heq'|Hneq']. + rewrite Heq' in H12; elim (Rlt_irrefl _ (Rle_lt_trans _ _ _ H11 H12)). + rewrite Heq in H7; elim (Rlt_irrefl _ (Rle_lt_trans _ _ _ H10 H7)). + rewrite Heq' in H12; elim (Rlt_irrefl _ (Rle_lt_trans _ _ _ H11 H12)). reflexivity. Qed. @@ -1852,17 +1827,17 @@ Proof. intros; replace (primitive h pr a) with 0. replace (RiemannInt pr0) with (primitive h pr b). ring. - unfold primitive; case (Rle_dec a b); case (Rle_dec b b); intros; + unfold primitive; destruct (Rle_dec a b) as [Hle|[]]; destruct (Rle_dec b b) as [Hle'|Hnle']; [ apply RiemannInt_P5 - | elim n; right; reflexivity - | elim n; assumption - | elim n0; assumption ]. - symmetry ; unfold primitive; case (Rle_dec a a); - case (Rle_dec a b); intros; + | destruct Hnle'; right; reflexivity + | assumption + | assumption]. + symmetry ; unfold primitive; destruct (Rle_dec a a) as [Hle|[]]; + destruct (Rle_dec a b) as [Hle'|Hnle']; [ apply RiemannInt_P9 - | elim n; assumption - | elim n; right; reflexivity - | elim n0; right; reflexivity ]. + | elim Hnle'; assumption + | right; reflexivity + | right; reflexivity ]. Qed. Lemma RiemannInt_P21 : @@ -1906,34 +1881,29 @@ Proof. intro; cut (IsStepFun psi3 a c). intro; split with (mkStepFun X); split with (mkStepFun X2); simpl; split. - intros; unfold phi3, psi3; case (Rle_dec t b); case (Rle_dec a t); - intros. + intros; unfold phi3, psi3; case (Rle_dec t b) as [|Hnle]; case (Rle_dec a t) as [|Hnle']. elim H1; intros; apply H3. replace (Rmin a b) with a. replace (Rmax a b) with b. split; assumption. - unfold Rmax; case (Rle_dec a b); intro; - [ reflexivity | elim n; assumption ]. - unfold Rmin; case (Rle_dec a b); intro; - [ reflexivity | elim n; assumption ]. - elim n; replace a with (Rmin a c). + unfold Rmax; decide (Rle_dec a b) with Hyp1; reflexivity. + unfold Rmin; decide (Rle_dec a b) with Hyp1; reflexivity. + elim Hnle'; replace a with (Rmin a c). elim H0; intros; assumption. - unfold Rmin; case (Rle_dec a c); intro; - [ reflexivity | elim n0; apply Rle_trans with b; assumption ]. + unfold Rmin; case (Rle_dec a c) as [|[]]; + [ reflexivity | apply Rle_trans with b; assumption ]. elim H2; intros; apply H3. replace (Rmax b c) with (Rmax a c). elim H0; intros; split; try assumption. replace (Rmin b c) with b. auto with real. - unfold Rmin; case (Rle_dec b c); intro; - [ reflexivity | elim n0; assumption ]. - unfold Rmax; case (Rle_dec a c); case (Rle_dec b c); intros; - try (elim n0; assumption || elim n0; apply Rle_trans with b; assumption). - reflexivity. - elim n; replace a with (Rmin a c). + unfold Rmin; decide (Rle_dec b c) with Hyp2; reflexivity. + unfold Rmax; decide (Rle_dec b c) with Hyp2; case (Rle_dec a c) as [|[]]; + [ reflexivity | apply Rle_trans with b; assumption ]. + elim Hnle'; replace a with (Rmin a c). elim H0; intros; assumption. - unfold Rmin; case (Rle_dec a c); intro; - [ reflexivity | elim n1; apply Rle_trans with b; assumption ]. + unfold Rmin; case (Rle_dec a c) as [|[]]; + [ reflexivity | apply Rle_trans with b; assumption ]. rewrite <- (StepFun_P43 X0 X1 X2). apply Rle_lt_trans with (Rabs (RiemannInt_SF (mkStepFun X0)) + Rabs (RiemannInt_SF (mkStepFun X1))). @@ -1947,33 +1917,33 @@ Proof. apply Rle_antisym. apply StepFun_P37; try assumption. simpl; intros; unfold psi3; elim H0; clear H0; intros; - case (Rle_dec a x); case (Rle_dec x b); intros; - [ elim (Rlt_irrefl _ (Rle_lt_trans _ _ _ r H0)) + destruct (Rle_dec a x) as [Hle|Hnle]; destruct (Rle_dec x b) as [Hle'|Hnle']; + [ elim (Rlt_irrefl _ (Rle_lt_trans _ _ _ Hle' H0)) | right; reflexivity - | elim n; apply Rle_trans with b; [ assumption | left; assumption ] - | elim n0; apply Rle_trans with b; [ assumption | left; assumption ] ]. + | elim Hnle; apply Rle_trans with b; [ assumption | left; assumption ] + | elim Hnle; apply Rle_trans with b; [ assumption | left; assumption ] ]. apply StepFun_P37; try assumption. simpl; intros; unfold psi3; elim H0; clear H0; intros; - case (Rle_dec a x); case (Rle_dec x b); intros; - [ elim (Rlt_irrefl _ (Rle_lt_trans _ _ _ r H0)) + destruct (Rle_dec a x) as [Hle|Hnle]; destruct (Rle_dec x b) as [Hle'|Hnle']; + [ elim (Rlt_irrefl _ (Rle_lt_trans _ _ _ Hle' H0)) | right; reflexivity - | elim n; apply Rle_trans with b; [ assumption | left; assumption ] - | elim n0; apply Rle_trans with b; [ assumption | left; assumption ] ]. + | elim Hnle; apply Rle_trans with b; [ assumption | left; assumption ] + | elim Hnle; apply Rle_trans with b; [ assumption | left; assumption ] ]. apply Rle_antisym. apply StepFun_P37; try assumption. simpl; intros; unfold psi3; elim H0; clear H0; intros; - case (Rle_dec a x); case (Rle_dec x b); intros; + destruct (Rle_dec a x) as [Hle|Hnle]; destruct (Rle_dec x b) as [Hle'|Hnle']; [ right; reflexivity - | elim n; left; assumption - | elim n; left; assumption - | elim n0; left; assumption ]. + | elim Hnle'; left; assumption + | elim Hnle; left; assumption + | elim Hnle; left; assumption ]. apply StepFun_P37; try assumption. simpl; intros; unfold psi3; elim H0; clear H0; intros; - case (Rle_dec a x); case (Rle_dec x b); intros; + destruct (Rle_dec a x) as [Hle|Hnle]; destruct (Rle_dec x b) as [Hle'|Hnle']; [ right; reflexivity - | elim n; left; assumption - | elim n; left; assumption - | elim n0; left; assumption ]. + | elim Hnle'; left; assumption + | elim Hnle; left; assumption + | elim Hnle; left; assumption ]. apply StepFun_P46 with b; assumption. assert (H3 := pre psi2); unfold IsStepFun in H3; unfold is_subdivision in H3; elim H3; clear H3; intros l1 [lf1 H3]; split with l1; @@ -1990,14 +1960,14 @@ Proof. apply lt_trans with (pred (Rlength l1)); try assumption; apply lt_pred_n_n; apply neq_O_lt; red; intro; rewrite <- H12 in H6; discriminate. - unfold Rmin; case (Rle_dec b c); intro; - [ reflexivity | elim n; assumption ]. + unfold Rmin; decide (Rle_dec b c) with Hyp2; + reflexivity. elim H7; intros; assumption. - case (Rle_dec a x); case (Rle_dec x b); intros; - [ elim (Rlt_irrefl _ (Rle_lt_trans _ _ _ r H10)) + destruct (Rle_dec a x) as [Hle|Hnle]; destruct (Rle_dec x b) as [Hle'|Hnle']; + [ elim (Rlt_irrefl _ (Rle_lt_trans _ _ _ Hle' H10)) | reflexivity - | elim n; apply Rle_trans with b; [ assumption | left; assumption ] - | elim n0; apply Rle_trans with b; [ assumption | left; assumption ] ]. + | elim Hnle; apply Rle_trans with b; [ assumption | left; assumption ] + | elim Hnle; apply Rle_trans with b; [ assumption | left; assumption ] ]. assert (H3 := pre psi1); unfold IsStepFun in H3; unfold is_subdivision in H3; elim H3; clear H3; intros l1 [lf1 H3]; split with l1; split with lf1; unfold adapted_couple in H3; decompose [and] H3; @@ -2012,8 +1982,7 @@ Proof. rewrite <- H4; elim (RList_P6 l1); intros; apply H10; try assumption. apply lt_pred_n_n; apply neq_O_lt; red; intro; rewrite <- H12 in H6; discriminate. - unfold Rmax; case (Rle_dec a b); intro; - [ reflexivity | elim n; assumption ]. + unfold Rmax; decide (Rle_dec a b) with Hyp1; reflexivity. assert (H11 : a <= x). apply Rle_trans with (pos_Rl l1 i). replace a with (Rmin a b). @@ -2022,11 +1991,9 @@ Proof. apply lt_trans with (pred (Rlength l1)); try assumption; apply lt_pred_n_n; apply neq_O_lt; red; intro; rewrite <- H13 in H6; discriminate. - unfold Rmin; case (Rle_dec a b); intro; - [ reflexivity | elim n; assumption ]. + unfold Rmin; decide (Rle_dec a b) with Hyp1; reflexivity. left; elim H7; intros; assumption. - case (Rle_dec a x); case (Rle_dec x b); intros; reflexivity || elim n; - assumption. + decide (Rle_dec a x) with H11; decide (Rle_dec x b) with H10; reflexivity. apply StepFun_P46 with b. assert (H3 := pre phi1); unfold IsStepFun in H3; unfold is_subdivision in H3; elim H3; clear H3; intros l1 [lf1 H3]; split with l1; @@ -2042,8 +2009,7 @@ Proof. rewrite <- H4; elim (RList_P6 l1); intros; apply H10; try assumption. apply lt_pred_n_n; apply neq_O_lt; red; intro; rewrite <- H12 in H6; discriminate. - unfold Rmax; case (Rle_dec a b); intro; - [ reflexivity | elim n; assumption ]. + unfold Rmax; decide (Rle_dec a b) with Hyp1; reflexivity. assert (H11 : a <= x). apply Rle_trans with (pos_Rl l1 i). replace a with (Rmin a b). @@ -2052,10 +2018,9 @@ Proof. apply lt_trans with (pred (Rlength l1)); try assumption; apply lt_pred_n_n; apply neq_O_lt; red; intro; rewrite <- H13 in H6; discriminate. - unfold Rmin; case (Rle_dec a b); intro; - [ reflexivity | elim n; assumption ]. + unfold Rmin; decide (Rle_dec a b) with Hyp1; reflexivity. left; elim H7; intros; assumption. - unfold phi3; case (Rle_dec a x); case (Rle_dec x b); intros; + unfold phi3; decide (Rle_dec a x) with H11; decide (Rle_dec x b) with H10; reflexivity || elim n; assumption. assert (H3 := pre phi2); unfold IsStepFun in H3; unfold is_subdivision in H3; elim H3; clear H3; intros l1 [lf1 H3]; split with l1; @@ -2072,14 +2037,13 @@ Proof. apply lt_trans with (pred (Rlength l1)); try assumption; apply lt_pred_n_n; apply neq_O_lt; red; intro; rewrite <- H12 in H6; discriminate. - unfold Rmin; case (Rle_dec b c); intro; - [ reflexivity | elim n; assumption ]. + unfold Rmin; decide (Rle_dec b c) with Hyp2; reflexivity. elim H7; intros; assumption. - unfold phi3; case (Rle_dec a x); case (Rle_dec x b); intros; - [ elim (Rlt_irrefl _ (Rle_lt_trans _ _ _ r H10)) + unfold phi3; destruct (Rle_dec a x) as [Hle|Hnle]; destruct (Rle_dec x b) as [Hle'|Hnle']; intros; + [ elim (Rlt_irrefl _ (Rle_lt_trans _ _ _ Hle' H10)) | reflexivity - | elim n; apply Rle_trans with b; [ assumption | left; assumption ] - | elim n0; apply Rle_trans with b; [ assumption | left; assumption ] ]. + | elim Hnle; apply Rle_trans with b; [ assumption | left; assumption ] + | elim Hnle; apply Rle_trans with b; [ assumption | left; assumption ] ]. Qed. Lemma RiemannInt_P22 : @@ -2098,21 +2062,10 @@ Proof. split; assumption. split with (mkStepFun H3); split with (mkStepFun H4); split. simpl; intros; apply H. - replace (Rmin a b) with (Rmin a c). - elim H5; intros; split; try assumption. + replace (Rmin a b) with (Rmin a c) by (rewrite 2!Rmin_left; eauto using Rle_trans). + destruct H5; split; try assumption. apply Rle_trans with (Rmax a c); try assumption. - replace (Rmax a b) with b. - replace (Rmax a c) with c. - assumption. - unfold Rmax; case (Rle_dec a c); intro; - [ reflexivity | elim n; assumption ]. - unfold Rmax; case (Rle_dec a b); intro; - [ reflexivity | elim n; apply Rle_trans with c; assumption ]. - unfold Rmin; case (Rle_dec a c); case (Rle_dec a b); intros; - [ reflexivity - | elim n; apply Rle_trans with c; assumption - | elim n; assumption - | elim n0; assumption ]. + apply Rle_max_compat_l; assumption. rewrite Rabs_right. assert (H5 : IsStepFun psi c b). apply StepFun_P46 with a. @@ -2130,15 +2083,11 @@ Proof. apply Rle_trans with (Rabs (f x - phi x)). apply Rabs_pos. apply H. - replace (Rmin a b) with a. - replace (Rmax a b) with b. - elim H6; intros; split; left. + rewrite Rmin_left; eauto using Rle_trans. + rewrite Rmax_right; eauto using Rle_trans. + destruct H6; split; left. apply Rle_lt_trans with c; assumption. assumption. - unfold Rmax; case (Rle_dec a b); intro; - [ reflexivity | elim n; apply Rle_trans with c; assumption ]. - unfold Rmin; case (Rle_dec a b); intro; - [ reflexivity | elim n; apply Rle_trans with c; assumption ]. rewrite StepFun_P18; ring. apply Rle_lt_trans with (Rabs (RiemannInt_SF psi)). apply RRle_abs. @@ -2160,15 +2109,11 @@ Proof. apply Rle_trans with (Rabs (f x - phi x)). apply Rabs_pos. apply H. - replace (Rmin a b) with a. - replace (Rmax a b) with b. - elim H5; intros; split; left. + rewrite Rmin_left; eauto using Rle_trans. + rewrite Rmax_right; eauto using Rle_trans. + destruct H5; split; left. assumption. apply Rlt_le_trans with c; assumption. - unfold Rmax; case (Rle_dec a b); intro; - [ reflexivity | elim n; apply Rle_trans with c; assumption ]. - unfold Rmin; case (Rle_dec a b); intro; - [ reflexivity | elim n; apply Rle_trans with c; assumption ]. rewrite StepFun_P18; ring. Qed. @@ -2191,18 +2136,10 @@ Proof. replace (Rmax a b) with (Rmax c b). elim H5; intros; split; try assumption. apply Rle_trans with (Rmin c b); try assumption. - replace (Rmin a b) with a. - replace (Rmin c b) with c. - assumption. - unfold Rmin; case (Rle_dec c b); intro; - [ reflexivity | elim n; assumption ]. - unfold Rmin; case (Rle_dec a b); intro; - [ reflexivity | elim n; apply Rle_trans with c; assumption ]. - unfold Rmax; case (Rle_dec c b); case (Rle_dec a b); intros; - [ reflexivity - | elim n; apply Rle_trans with c; assumption - | elim n; assumption - | elim n0; assumption ]. + rewrite Rmin_left; eauto using Rle_trans. + rewrite Rmin_left; eauto using Rle_trans. + rewrite Rmax_right; eauto using Rle_trans. + rewrite Rmax_right; eauto using Rle_trans. rewrite Rabs_right. assert (H5 : IsStepFun psi a c). apply StepFun_P46 with b. @@ -2220,15 +2157,11 @@ Proof. apply Rle_trans with (Rabs (f x - phi x)). apply Rabs_pos. apply H. - replace (Rmin a b) with a. - replace (Rmax a b) with b. - elim H6; intros; split; left. + rewrite Rmin_left; eauto using Rle_trans. + rewrite Rmax_right; eauto using Rle_trans. + destruct H6; split; left. assumption. apply Rlt_le_trans with c; assumption. - unfold Rmax; case (Rle_dec a b); intro; - [ reflexivity | elim n; apply Rle_trans with c; assumption ]. - unfold Rmin; case (Rle_dec a b); intro; - [ reflexivity | elim n; apply Rle_trans with c; assumption ]. rewrite StepFun_P18; ring. apply Rle_lt_trans with (Rabs (RiemannInt_SF psi)). apply RRle_abs. @@ -2250,15 +2183,11 @@ Proof. apply Rle_trans with (Rabs (f x - phi x)). apply Rabs_pos. apply H. - replace (Rmin a b) with a. - replace (Rmax a b) with b. - elim H5; intros; split; left. + rewrite Rmin_left; eauto using Rle_trans. + rewrite Rmax_right; eauto using Rle_trans. + destruct H5; split; left. apply Rle_lt_trans with c; assumption. assumption. - unfold Rmax; case (Rle_dec a b); intro; - [ reflexivity | elim n; apply Rle_trans with c; assumption ]. - unfold Rmin; case (Rle_dec a b); intro; - [ reflexivity | elim n; apply Rle_trans with c; assumption ]. rewrite StepFun_P18; ring. Qed. @@ -2291,16 +2220,15 @@ Lemma RiemannInt_P25 : a <= b -> b <= c -> RiemannInt pr1 + RiemannInt pr2 = RiemannInt pr3. Proof. intros f a b c pr1 pr2 pr3 Hyp1 Hyp2; unfold RiemannInt; - case (RiemannInt_exists pr1 RinvN RinvN_cv); - case (RiemannInt_exists pr2 RinvN RinvN_cv); - case (RiemannInt_exists pr3 RinvN RinvN_cv); intros; + case (RiemannInt_exists pr1 RinvN RinvN_cv) as (x1,HUn_cv1); + case (RiemannInt_exists pr2 RinvN RinvN_cv) as (x0,HUn_cv0); + case (RiemannInt_exists pr3 RinvN RinvN_cv) as (x,HUn_cv); symmetry ; eapply UL_sequence. - apply u. + apply HUn_cv. unfold Un_cv; intros; assert (H0 : 0 < eps / 3). unfold Rdiv; apply Rmult_lt_0_compat; [ assumption | apply Rinv_0_lt_compat; prove_sup0 ]. - elim (u1 _ H0); clear u1; intros N1 H1; elim (u0 _ H0); clear u0; - intros N2 H2; + destruct (HUn_cv1 _ H0) as (N1,H1); clear HUn_cv1; destruct (HUn_cv0 _ H0) as (N2,H2); clear HUn_cv0; cut (Un_cv (fun n:nat => @@ -2357,7 +2285,7 @@ Proof. do 2 rewrite (Rmult_comm 3); repeat rewrite Rmult_assoc; rewrite <- Rinv_l_sym; [ ring | discrR ] | discrR ]. - clear x u x0 x1 eps H H0 N1 H1 N2 H2; + clear x HUn_cv x0 x1 eps H H0 N1 H1 N2 H2; assert (H1 : exists psi1 : nat -> StepFun a b, @@ -2477,25 +2405,17 @@ Proof. apply Rplus_le_compat. apply H1. elim H14; intros; split. - replace (Rmin a c) with a. + rewrite Rmin_left; eauto using Rle_trans. apply Rle_trans with b; try assumption. left; assumption. - unfold Rmin; case (Rle_dec a c); intro; - [ reflexivity | elim n0; apply Rle_trans with b; assumption ]. - replace (Rmax a c) with c. + rewrite Rmax_right; eauto using Rle_trans. left; assumption. - unfold Rmax; case (Rle_dec a c); intro; - [ reflexivity | elim n0; apply Rle_trans with b; assumption ]. apply H3. elim H14; intros; split. - replace (Rmin b c) with b. + rewrite Rmin_left; eauto using Rle_trans. left; assumption. - unfold Rmin; case (Rle_dec b c); intro; - [ reflexivity | elim n0; assumption ]. - replace (Rmax b c) with c. + rewrite Rmax_right; eauto using Rle_trans. left; assumption. - unfold Rmax; case (Rle_dec b c); intro; - [ reflexivity | elim n0; assumption ]. do 2 rewrite <- (Rplus_comm @@ -2509,26 +2429,18 @@ Proof. apply Rplus_le_compat. apply H1. elim H14; intros; split. - replace (Rmin a c) with a. + rewrite Rmin_left; eauto using Rle_trans. left; assumption. - unfold Rmin; case (Rle_dec a c); intro; - [ reflexivity | elim n0; apply Rle_trans with b; assumption ]. - replace (Rmax a c) with c. + rewrite Rmax_right; eauto using Rle_trans. apply Rle_trans with b. left; assumption. assumption. - unfold Rmax; case (Rle_dec a c); intro; - [ reflexivity | elim n0; apply Rle_trans with b; assumption ]. apply H8. elim H14; intros; split. - replace (Rmin a b) with a. + rewrite Rmin_left; trivial. left; assumption. - unfold Rmin; case (Rle_dec a b); intro; - [ reflexivity | elim n0; assumption ]. - replace (Rmax a b) with b. + rewrite Rmax_right; trivial. left; assumption. - unfold Rmax; case (Rle_dec a b); intro; - [ reflexivity | elim n0; assumption ]. do 2 rewrite StepFun_P30. do 2 rewrite Rmult_1_l; replace @@ -2571,27 +2483,27 @@ Lemma RiemannInt_P26 : (pr2:Riemann_integrable f b c) (pr3:Riemann_integrable f a c), RiemannInt pr1 + RiemannInt pr2 = RiemannInt pr3. Proof. - intros; case (Rle_dec a b); case (Rle_dec b c); intros. + intros; destruct (Rle_dec a b) as [Hle|Hnle]; destruct (Rle_dec b c) as [Hle'|Hnle']. apply RiemannInt_P25; assumption. - case (Rle_dec a c); intro. + destruct (Rle_dec a c) as [Hle''|Hnle'']. assert (H : c <= b). auto with real. - rewrite <- (RiemannInt_P25 pr3 (RiemannInt_P1 pr2) pr1 r0 H); + rewrite <- (RiemannInt_P25 pr3 (RiemannInt_P1 pr2) pr1 Hle'' H); rewrite (RiemannInt_P8 pr2 (RiemannInt_P1 pr2)); ring. assert (H : c <= a). auto with real. rewrite (RiemannInt_P8 pr2 (RiemannInt_P1 pr2)); - rewrite <- (RiemannInt_P25 (RiemannInt_P1 pr3) pr1 (RiemannInt_P1 pr2) H r); + rewrite <- (RiemannInt_P25 (RiemannInt_P1 pr3) pr1 (RiemannInt_P1 pr2) H Hle); rewrite (RiemannInt_P8 pr3 (RiemannInt_P1 pr3)); ring. assert (H : b <= a). auto with real. - case (Rle_dec a c); intro. - rewrite <- (RiemannInt_P25 (RiemannInt_P1 pr1) pr3 pr2 H r0); + destruct (Rle_dec a c) as [Hle''|Hnle'']. + rewrite <- (RiemannInt_P25 (RiemannInt_P1 pr1) pr3 pr2 H Hle''); rewrite (RiemannInt_P8 pr1 (RiemannInt_P1 pr1)); ring. assert (H0 : c <= a). auto with real. rewrite (RiemannInt_P8 pr1 (RiemannInt_P1 pr1)); - rewrite <- (RiemannInt_P25 pr2 (RiemannInt_P1 pr3) (RiemannInt_P1 pr1) r H0); + rewrite <- (RiemannInt_P25 pr2 (RiemannInt_P1 pr3) (RiemannInt_P1 pr1) Hle' H0); rewrite (RiemannInt_P8 pr3 (RiemannInt_P1 pr3)); ring. rewrite (RiemannInt_P8 pr1 (RiemannInt_P1 pr1)); rewrite (RiemannInt_P8 pr2 (RiemannInt_P1 pr2)); @@ -2616,13 +2528,13 @@ Proof. assert (H4 : 0 < del). unfold del; unfold Rmin; case (Rle_dec (b - x) (x - a)); intro. - case (Rle_dec x0 (b - x)); intro; + destruct (Rle_dec x0 (b - x)) as [Hle|Hnle]; [ elim H3; intros; assumption | apply Rlt_Rminus; assumption ]. - case (Rle_dec x0 (x - a)); intro; + destruct (Rle_dec x0 (x - a)) as [Hle'|Hnle']; [ elim H3; intros; assumption | apply Rlt_Rminus; assumption ]. split with (mkposreal _ H4); intros; assert (H7 : Riemann_integrable f x (x + h0)). - case (Rle_dec x (x + h0)); intro. + destruct (Rle_dec x (x + h0)) as [Hle''|Hnle'']. apply continuity_implies_RiemannInt; try assumption. intros; apply C0; elim H7; intros; split. apply Rle_trans with x; [ left; assumption | assumption ]. @@ -2659,7 +2571,7 @@ Proof. with ((RiemannInt H7 - RiemannInt (RiemannInt_P14 x (x + h0) (f x))) / h0). replace (RiemannInt H7 - RiemannInt (RiemannInt_P14 x (x + h0) (f x))) with (RiemannInt (RiemannInt_P10 (-1) H7 (RiemannInt_P14 x (x + h0) (f x)))). - unfold Rdiv; rewrite Rabs_mult; case (Rle_dec x (x + h0)); intro. + unfold Rdiv; rewrite Rabs_mult; destruct (Rle_dec x (x + h0)) as [Hle|Hnle]. apply Rle_lt_trans with (RiemannInt (RiemannInt_P16 @@ -2678,14 +2590,14 @@ Proof. apply Rabs_pos. apply RiemannInt_P19; try assumption. intros; replace (f x1 + -1 * fct_cte (f x) x1) with (f x1 - f x). - unfold fct_cte; case (Req_dec x x1); intro. + unfold fct_cte; destruct (Req_dec x x1) as [H9|H9]. rewrite H9; unfold Rminus; rewrite Rplus_opp_r; rewrite Rabs_R0; left; assumption. - elim H3; intros; left; apply H11. + elim H3; intros; left; apply H11. repeat split. assumption. rewrite Rabs_right. - apply Rplus_lt_reg_r with x; replace (x + (x1 - x)) with x1; [ idtac | ring ]. + apply Rplus_lt_reg_l with x; replace (x + (x1 - x)) with x1; [ idtac | ring ]. apply Rlt_le_trans with (x + h0). elim H8; intros; assumption. apply Rplus_le_compat_l; apply Rle_trans with del. @@ -2707,8 +2619,8 @@ Proof. apply Rinv_r_sym. assumption. apply Rle_ge; left; apply Rinv_0_lt_compat. - elim r; intro. - apply Rplus_lt_reg_r with x; rewrite Rplus_0_r; assumption. + elim Hle; intro. + apply Rplus_lt_reg_l with x; rewrite Rplus_0_r; assumption. elim H5; symmetry ; apply Rplus_eq_reg_l with x; rewrite Rplus_0_r; assumption. apply Rle_lt_trans with @@ -2748,7 +2660,7 @@ Proof. repeat split. assumption. rewrite Rabs_left. - apply Rplus_lt_reg_r with (x1 - x0); replace (x1 - x0 + x0) with x1; + apply Rplus_lt_reg_l with (x1 - x0); replace (x1 - x0 + x0) with x1; [ idtac | ring ]. replace (x1 - x0 + - (x1 - x)) with (x - x0); [ idtac | ring ]. apply Rle_lt_trans with (x + h0). @@ -2758,7 +2670,7 @@ Proof. apply Rle_trans with del; [ left; assumption | unfold del; apply Rmin_l ]. elim H8; intros; assumption. - apply Rplus_lt_reg_r with x; rewrite Rplus_0_r; + apply Rplus_lt_reg_l with x; rewrite Rplus_0_r; replace (x + (x1 - x)) with x1; [ elim H8; intros; assumption | ring ]. unfold fct_cte; ring. rewrite RiemannInt_P15. @@ -2778,7 +2690,7 @@ Proof. apply Rinv_lt_0_compat. assert (H8 : x + h0 < x). auto with real. - apply Rplus_lt_reg_r with x; rewrite Rplus_0_r; assumption. + apply Rplus_lt_reg_l with x; rewrite Rplus_0_r; assumption. rewrite (RiemannInt_P13 H7 (RiemannInt_P14 x (x + h0) (f x)) (RiemannInt_P10 (-1) H7 (RiemannInt_P14 x (x + h0) (f x)))) @@ -2792,9 +2704,11 @@ Proof. cut (a <= x + h0). cut (x + h0 <= b). intros; unfold primitive. - case (Rle_dec a (x + h0)); case (Rle_dec (x + h0) b); case (Rle_dec a x); - case (Rle_dec x b); intros; try (elim n; assumption || left; assumption). - rewrite <- (RiemannInt_P26 (FTC_P1 h C0 r0 r) H7 (FTC_P1 h C0 r2 r1)); ring. + assert (H10: a <= x) by (left; assumption). + assert (H11: x <= b) by (left; assumption). + decide (Rle_dec a (x + h0)) with H9; decide (Rle_dec (x + h0) b) with H8; + decide (Rle_dec a x) with H10; decide (Rle_dec x b) with H11. + rewrite <- (RiemannInt_P26 (FTC_P1 h C0 H10 H11) H7 (FTC_P1 h C0 H9 H8)); ring. apply Rplus_le_reg_l with (- x); replace (- x + (x + h0)) with h0; [ idtac | ring ]. rewrite Rplus_comm; apply Rle_trans with (Rabs h0). @@ -2854,11 +2768,11 @@ Proof. unfold R_dist; intros; set (del := Rmin x0 (Rmin x1 (b - a))); assert (H10 : 0 < del). unfold del; unfold Rmin; case (Rle_dec x1 (b - a)); intros. - case (Rle_dec x0 x1); intro; + destruct (Rle_dec x0 x1) as [Hle|Hnle]; [ apply (cond_pos x0) | elim H9; intros; assumption ]. - case (Rle_dec x0 (b - a)); intro; + destruct (Rle_dec x0 (b - a)) as [Hle'|Hnle']; [ apply (cond_pos x0) | apply Rlt_Rminus; assumption ]. - split with (mkposreal _ H10); intros; case (Rcase_abs h0); intro. + split with (mkposreal _ H10); intros; destruct (Rcase_abs h0) as [Hle|Hnle]. assert (H14 : b + h0 < b). pattern b at 2; rewrite <- Rplus_0_r; apply Rplus_lt_compat_l; assumption. @@ -2914,7 +2828,7 @@ Proof. repeat split. assumption. rewrite <- Rabs_Ropp; rewrite Ropp_minus_distr; rewrite Rabs_right. - apply Rplus_lt_reg_r with (x2 - x1); + apply Rplus_lt_reg_l with (x2 - x1); replace (x2 - x1 + (b - x2)) with (b - x1); [ idtac | ring ]. replace (x2 - x1 + x1) with x2; [ idtac | ring ]. apply Rlt_le_trans with (b + h0). @@ -2957,11 +2871,11 @@ Proof. | assumption ]. cut (a <= b + h0). cut (b + h0 <= b). - intros; unfold primitive; case (Rle_dec a (b + h0)); - case (Rle_dec (b + h0) b); case (Rle_dec a b); case (Rle_dec b b); - intros; try (elim n; right; reflexivity) || (elim n; left; assumption). - rewrite <- (RiemannInt_P26 (FTC_P1 h C0 r3 r2) H13 (FTC_P1 h C0 r1 r0)); ring. - elim n; assumption. + intros; unfold primitive; destruct (Rle_dec a (b + h0)) as [Hle'|Hnle']; + destruct (Rle_dec (b + h0) b) as [Hle''|[]]; destruct (Rle_dec a b) as [Hleab|[]]; destruct (Rle_dec b b) as [Hlebb|[]]; + assumption || (right; reflexivity) || (try (left; assumption)). + rewrite <- (RiemannInt_P26 (FTC_P1 h C0 Hle' Hle'') H13 (FTC_P1 h C0 Hleab Hlebb)); ring. + elim Hnle'; assumption. left; assumption. apply Rplus_le_reg_l with (- a - h0). replace (- a - h0 + a) with (- h0); [ idtac | ring ]. @@ -2979,22 +2893,22 @@ Proof. [ assumption | unfold del; apply Rmin_l ]. assert (H14 : b < b + h0). pattern b at 1; rewrite <- Rplus_0_r; apply Rplus_lt_compat_l. - assert (H14 := Rge_le _ _ r); elim H14; intro. + assert (H14 := Rge_le _ _ Hnle); elim H14; intro. assumption. elim H11; symmetry ; assumption. - unfold primitive; case (Rle_dec a (b + h0)); - case (Rle_dec (b + h0) b); intros; - [ elim (Rlt_irrefl _ (Rle_lt_trans _ _ _ r0 H14)) + unfold primitive; destruct (Rle_dec a (b + h0)) as [Hle|[]]; + destruct (Rle_dec (b + h0) b) as [Hle'|Hnle']; + [ elim (Rlt_irrefl _ (Rle_lt_trans _ _ _ Hle' H14)) | unfold f_b; reflexivity - | elim n; left; apply Rlt_trans with b; assumption - | elim n0; left; apply Rlt_trans with b; assumption ]. + | left; apply Rlt_trans with b; assumption + | left; apply Rlt_trans with b; assumption ]. unfold f_b; unfold Rminus; rewrite Rplus_opp_r; rewrite Rmult_0_r; rewrite Rplus_0_l; unfold primitive; - case (Rle_dec a b); case (Rle_dec b b); intros; + destruct (Rle_dec a b) as [Hle'|Hnle']; destruct (Rle_dec b b) as [Hle''|[]]; [ apply RiemannInt_P5 - | elim n; right; reflexivity - | elim n; left; assumption - | elim n; right; reflexivity ]. + | right; reflexivity + | elim Hnle'; left; assumption + | right; reflexivity ]. (*****) set (f_a := fun x:R => f a * (x - a)); rewrite <- H2; assert (H3 : derivable_pt_lim f_a a (f a)). @@ -3028,16 +2942,18 @@ Proof. apply (cond_pos x0). apply Rlt_Rminus; assumption. split with (mkposreal _ H9). - intros; case (Rcase_abs h0); intro. + intros; destruct (Rcase_abs h0) as [Hle|Hnle]. assert (H12 : a + h0 < a). pattern a at 2; rewrite <- Rplus_0_r; apply Rplus_lt_compat_l; assumption. unfold primitive. - case (Rle_dec a (a + h0)); case (Rle_dec (a + h0) b); case (Rle_dec a a); - case (Rle_dec a b); intros; - try (elim n; left; assumption) || (elim n; right; reflexivity). - elim (Rlt_irrefl _ (Rle_lt_trans _ _ _ r3 H12)). - elim n; left; apply Rlt_trans with a; assumption. + destruct (Rle_dec a (a + h0)) as [Hle'|Hnle']; + destruct (Rle_dec (a + h0) b) as [Hle''|Hnle'']; + destruct (Rle_dec a a) as [Hleaa|[]]; + destruct (Rle_dec a b) as [Hleab|[]]; + try (left; assumption) || (right; reflexivity). + elim (Rlt_irrefl _ (Rle_lt_trans _ _ _ Hle' H12)). + elim Hnle''; left; apply Rlt_trans with a; assumption. rewrite RiemannInt_P9; replace 0 with (f_a a). replace (f a * (a + h0 - a)) with (f_a (a + h0)). apply H5; try assumption. @@ -3045,10 +2961,10 @@ Proof. [ assumption | unfold del; apply Rmin_l ]. unfold f_a; ring. unfold f_a; ring. - elim n; left; apply Rlt_trans with a; assumption. + elim Hnle''; left; apply Rlt_trans with a; assumption. assert (H12 : a < a + h0). pattern a at 1; rewrite <- Rplus_0_r; apply Rplus_lt_compat_l. - assert (H12 := Rge_le _ _ r); elim H12; intro. + assert (H12 := Rge_le _ _ Hnle); elim H12; intro. assumption. elim H10; symmetry ; assumption. assert (H13 : Riemann_integrable f a (a + h0)). @@ -3097,7 +3013,7 @@ Proof. elim H8; intros; left; apply H17; repeat split. assumption. rewrite Rabs_right. - apply Rplus_lt_reg_r with a; replace (a + (x2 - a)) with x2; [ idtac | ring ]. + apply Rplus_lt_reg_l with a; replace (a + (x2 - a)) with x2; [ idtac | ring ]. apply Rlt_le_trans with (a + h0). elim H14; intros; assumption. apply Rplus_le_compat_l; left; apply Rle_lt_trans with (Rabs h0). @@ -3121,7 +3037,7 @@ Proof. rewrite Rplus_comm; unfold Rminus; rewrite Rplus_assoc; rewrite Rplus_opp_r; rewrite Rplus_0_r; rewrite <- Rinv_r_sym; [ reflexivity | assumption ]. - apply Rle_ge; left; apply Rinv_0_lt_compat; assert (H14 := Rge_le _ _ r); + apply Rle_ge; left; apply Rinv_0_lt_compat; assert (H14 := Rge_le _ _ Hnle); elim H14; intro. assumption. elim H10; symmetry ; assumption. @@ -3136,13 +3052,13 @@ Proof. rewrite Rmult_assoc; rewrite <- Rinv_r_sym; [ ring | assumption ]. cut (a <= a + h0). cut (a + h0 <= b). - intros; unfold primitive; case (Rle_dec a (a + h0)); - case (Rle_dec (a + h0) b); case (Rle_dec a a); case (Rle_dec a b); - intros; try (elim n; right; reflexivity) || (elim n; left; assumption). + intros; unfold primitive. + decide (Rle_dec (a+h0) b) with H14. + decide (Rle_dec a a) with (Rle_refl a). + decide (Rle_dec a (a+h0)) with H15. + decide (Rle_dec a b) with h. rewrite RiemannInt_P9; unfold Rminus; rewrite Ropp_0; rewrite Rplus_0_r; apply RiemannInt_P5. - elim n; assumption. - elim n; assumption. 2: left; assumption. apply Rplus_le_reg_l with (- a); replace (- a + (a + h0)) with h0; [ idtac | ring ]. @@ -3189,18 +3105,18 @@ Proof. unfold derivable_pt_lim; intros; elim (H2 _ H4); intros; elim (H3 _ H4); intros; set (del := Rmin x0 x1). assert (H7 : 0 < del). - unfold del; unfold Rmin; case (Rle_dec x0 x1); intro. + unfold del; unfold Rmin; destruct (Rle_dec x0 x1) as [Hle|Hnle]. apply (cond_pos x0). apply (cond_pos x1). - split with (mkposreal _ H7); intros; case (Rcase_abs h0); intro. + split with (mkposreal _ H7); intros; destruct (Rcase_abs h0) as [Hle|Hnle]. assert (H10 : a + h0 < a). pattern a at 2; rewrite <- Rplus_0_r; apply Rplus_lt_compat_l; assumption. - rewrite H1; unfold primitive; case (Rle_dec a (a + h0)); - case (Rle_dec (a + h0) b); case (Rle_dec a a); case (Rle_dec a b); - intros; try (elim n; right; assumption || reflexivity). - elim (Rlt_irrefl _ (Rle_lt_trans _ _ _ r3 H10)). - elim (Rlt_irrefl _ (Rle_lt_trans _ _ _ r2 H10)). + rewrite H1; unfold primitive. + apply (decide_left (Rle_dec a b) h); intro h'. + assert (H11:~ a<=a+h0) by auto using Rlt_not_le. + decide (Rle_dec a (a+h0)) with H11. + decide (Rle_dec a a) with (Rle_refl a). rewrite RiemannInt_P9; replace 0 with (f_a a). replace (f a * (a + h0 - a)) with (f_a (a + h0)). apply H5; try assumption. @@ -3208,27 +3124,26 @@ Proof. unfold del; apply Rmin_l. unfold f_a; ring. unfold f_a; ring. - elim n; rewrite <- H0; left; assumption. assert (H10 : a < a + h0). pattern a at 1; rewrite <- Rplus_0_r; apply Rplus_lt_compat_l. - assert (H10 := Rge_le _ _ r); elim H10; intro. + assert (H10 := Rge_le _ _ Hnle); elim H10; intro. assumption. elim H8; symmetry ; assumption. - rewrite H0 in H1; rewrite H1; unfold primitive; - case (Rle_dec a (b + h0)); case (Rle_dec (b + h0) b); - case (Rle_dec a b); case (Rle_dec b b); intros; - try (elim n; right; assumption || reflexivity). - rewrite H0 in H10; elim (Rlt_irrefl _ (Rle_lt_trans _ _ _ r2 H10)). - repeat rewrite RiemannInt_P9. - replace (RiemannInt (FTC_P1 h C0 r1 r0)) with (f_b b). + rewrite H0 in H1; rewrite H1; unfold primitive. + decide (Rle_dec a b) with h. + decide (Rle_dec b b) with (Rle_refl b). + assert (H12 : a<=b+h0) by (eauto using Rlt_le_trans with real). + decide (Rle_dec a (b+h0)) with H12. + rewrite H0 in H10. + assert (H13 : ~b+h0<=b) by (auto using Rlt_not_le). + decide (Rle_dec (b+h0) b) with H13. + replace (RiemannInt (FTC_P1 h C0 hbis H11)) with (f_b b). fold (f_b (b + h0)). apply H6; try assumption. apply Rlt_le_trans with del; try assumption. unfold del; apply Rmin_r. unfold f_b; unfold Rminus; rewrite Rplus_opp_r; rewrite Rmult_0_r; rewrite Rplus_0_l; apply RiemannInt_P5. - elim n; rewrite <- H0; left; assumption. - elim n0; rewrite <- H0; left; assumption. Qed. Lemma RiemannInt_P29 : @@ -3266,7 +3181,7 @@ Qed. Lemma RiemannInt_P32 : forall (f:C1_fun) (a b:R), Riemann_integrable (derive f (diff0 f)) a b. Proof. - intro f; intros; case (Rle_dec a b); intro; + intro f; intros; destruct (Rle_dec a b) as [Hle|Hnle]; [ apply continuity_implies_RiemannInt; try assumption; intros; apply (cont1 f) | assert (H : b <= a); @@ -3296,10 +3211,45 @@ Lemma FTC_Riemann : forall (f:C1_fun) (a b:R) (pr:Riemann_integrable (derive f (diff0 f)) a b), RiemannInt pr = f b - f a. Proof. - intro f; intros; case (Rle_dec a b); intro; + intro f; intros; destruct (Rle_dec a b) as [Hle|Hnle]; [ apply RiemannInt_P33; assumption | assert (H : b <= a); [ auto with real | assert (H0 := RiemannInt_P1 pr); rewrite (RiemannInt_P8 pr H0); rewrite (RiemannInt_P33 _ H0 H); ring ] ]. Qed. + +(* RiemannInt *) +Lemma RiemannInt_const_bound : + forall f a b l u (h : Riemann_integrable f a b), a <= b -> + (forall x, a < x < b -> l <= f x <= u) -> + l * (b - a) <= RiemannInt h <= u * (b - a). +intros f a b l u ri ab intf. +rewrite <- !(fun l => RiemannInt_P15 (RiemannInt_P14 a b l)). +split; apply RiemannInt_P19; try assumption; + intros x intx; unfold fct_cte; destruct (intf x intx); assumption. +Qed. + +Lemma Riemann_integrable_scal : + forall f a b k, + Riemann_integrable f a b -> + Riemann_integrable (fun x => k * f x) a b. +intros f a b k ri. +apply Riemann_integrable_ext with + (f := fun x => 0 + k * f x). + intros; ring. +apply (RiemannInt_P10 _ (RiemannInt_P14 _ _ 0) ri). +Qed. + +Arguments Riemann_integrable_scal [f a b] k _ eps. + +Lemma Riemann_integrable_Ropp : + forall f a b, Riemann_integrable f a b -> + Riemann_integrable (fun x => - f x) a b. +intros ff a b h. +apply Riemann_integrable_ext with (f := fun x => (-1) * ff x). +intros; ring. +apply Riemann_integrable_scal; assumption. +Qed. + +Arguments Riemann_integrable_Ropp [f a b] _ eps. diff --git a/theories/Reals/RiemannInt_SF.v b/theories/Reals/RiemannInt_SF.v index 8eb49bf3..1484ab2a 100644 --- a/theories/Reals/RiemannInt_SF.v +++ b/theories/Reals/RiemannInt_SF.v @@ -1,6 +1,6 @@ (************************************************************************) (* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *) +(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *) (* \VV/ **************************************************************) (* // * This file is distributed under the terms of the *) (* * GNU Lesser General Public License Version 2.1 *) @@ -40,26 +40,25 @@ Proof. assert (H2 : exists x : R, E x). elim H; intros; exists (INR x); unfold E; exists x; split; [ assumption | reflexivity ]. - assert (H3 := completeness E H1 H2); elim H3; intros; unfold is_lub in p; - elim p; clear p; intros; unfold is_upper_bound in H4, H5; + destruct (completeness E H1 H2) as (x,(H4,H5)); unfold is_upper_bound in H4, H5; assert (H6 : 0 <= x). - elim H2; intros; unfold E in H6; elim H6; intros; elim H7; intros; + destruct H2 as (x0,H6). remember H6 as H7. destruct H7 as (x1,(H8,H9)). apply Rle_trans with x0; [ rewrite <- H9; change (INR 0 <= INR x1); apply le_INR; apply le_O_n | apply H4; assumption ]. assert (H7 := archimed x); elim H7; clear H7; intros; assert (H9 : x <= IZR (up x) - 1). - apply H5; intros; assert (H10 := H4 _ H9); unfold E in H9; elim H9; intros; - elim H11; intros; rewrite <- H13; apply Rplus_le_reg_l with 1; + apply H5; intros x0 H9. assert (H10 := H4 _ H9); unfold E in H9; elim H9; intros x1 (H12,<-). + apply Rplus_le_reg_l with 1; replace (1 + (IZR (up x) - 1)) with (IZR (up x)); [ idtac | ring ]; replace (1 + INR x1) with (INR (S x1)); [ idtac | rewrite S_INR; ring ]. assert (H14 : (0 <= up x)%Z). apply le_IZR; apply Rle_trans with x; [ apply H6 | left; assumption ]. - assert (H15 := IZN _ H14); elim H15; clear H15; intros; rewrite H15; - rewrite <- INR_IZR_INZ; apply le_INR; apply lt_le_S; - apply INR_lt; rewrite H13; apply Rle_lt_trans with x; + destruct (IZN _ H14) as (x2,H15). + rewrite H15, <- INR_IZR_INZ; apply le_INR; apply lt_le_S. + apply INR_lt; apply Rle_lt_trans with x; [ assumption | rewrite INR_IZR_INZ; rewrite <- H15; assumption ]. assert (H10 : x = IZR (up x) - 1). apply Rle_antisym; @@ -70,32 +69,32 @@ Proof. [ assumption | ring ] ]. assert (H11 : (0 <= up x)%Z). apply le_IZR; apply Rle_trans with x; [ apply H6 | left; assumption ]. - assert (H12 := IZN_var H11); elim H12; clear H12; intros; assert (H13 : E x). + assert (H12 := IZN_var H11); elim H12; clear H12; intros x0 H8; assert (H13 : E x). elim (classic (E x)); intro; try assumption. cut (forall y:R, E y -> y <= x - 1). - intro; assert (H14 := H5 _ H13); cut (x - 1 < x). - intro; elim (Rlt_irrefl _ (Rle_lt_trans _ _ _ H14 H15)). + intro H13; assert (H14 := H5 _ H13); cut (x - 1 < x). + intro H15; elim (Rlt_irrefl _ (Rle_lt_trans _ _ _ H14 H15)). apply Rminus_lt; replace (x - 1 - x) with (-1); [ idtac | ring ]; rewrite <- Ropp_0; apply Ropp_lt_gt_contravar; apply Rlt_0_1. - intros; assert (H14 := H4 _ H13); elim H14; intro; unfold E in H13; elim H13; - intros; elim H16; intros; apply Rplus_le_reg_l with 1. + intros y H13; assert (H14 := H4 _ H13); elim H14; intro H15; unfold E in H13; elim H13; + intros x1 H16; elim H16; intros H17 H18; apply Rplus_le_reg_l with 1. replace (1 + (x - 1)) with x; [ idtac | ring ]; rewrite <- H18; replace (1 + INR x1) with (INR (S x1)); [ idtac | rewrite S_INR; ring ]. cut (x = INR (pred x0)). - intro; rewrite H19; apply le_INR; apply lt_le_S; apply INR_lt; rewrite H18; + intro H19; rewrite H19; apply le_INR; apply lt_le_S; apply INR_lt; rewrite H18; rewrite <- H19; assumption. - rewrite H10; rewrite p; rewrite <- INR_IZR_INZ; replace 1 with (INR 1); + rewrite H10; rewrite H8; rewrite <- INR_IZR_INZ; replace 1 with (INR 1); [ idtac | reflexivity ]; rewrite <- minus_INR. replace (x0 - 1)%nat with (pred x0); [ reflexivity | case x0; [ reflexivity | intro; simpl; apply minus_n_O ] ]. - induction x0 as [| x0 Hrecx0]; - [ rewrite p in H7; rewrite <- INR_IZR_INZ in H7; simpl in H7; - elim (Rlt_irrefl _ (Rle_lt_trans _ _ _ H6 H7)) - | apply le_n_S; apply le_O_n ]. - rewrite H15 in H13; elim H12; assumption. + induction x0 as [|x0 Hrecx0]. + rewrite H8 in H3. rewrite <- INR_IZR_INZ in H3; simpl in H3. + elim (Rlt_irrefl _ (Rle_lt_trans _ _ _ H6 H3)). + apply le_n_S; apply le_O_n. + rewrite H15 in H13; elim H12; assumption. split with (pred x0); unfold E in H13; elim H13; intros; elim H12; intros; - rewrite H10 in H15; rewrite p in H15; rewrite <- INR_IZR_INZ in H15; + rewrite H10 in H15; rewrite H8 in H15; rewrite <- INR_IZR_INZ in H15; assert (H16 : INR x0 = INR x1 + 1). rewrite H15; ring. rewrite <- S_INR in H16; assert (H17 := INR_eq _ _ H16); rewrite H17; @@ -144,7 +143,7 @@ Definition subdivision (a b:R) (f:StepFun a b) : Rlist := projT1 (pre f). Definition subdivision_val (a b:R) (f:StepFun a b) : Rlist := match projT2 (pre f) with - | existT a b => a + | existT _ a b => a end. Fixpoint Int_SF (l k:Rlist) : R := @@ -173,8 +172,8 @@ Lemma StepFun_P1 : forall (a b:R) (f:StepFun a b), adapted_couple f a b (subdivision f) (subdivision_val f). Proof. - intros a b f; unfold subdivision_val; case (projT2 (pre f)); intros; - apply a0. + intros a b f; unfold subdivision_val; case (projT2 (pre f)) as (x,H); + apply H. Qed. Lemma StepFun_P2 : @@ -201,19 +200,17 @@ Proof. intros; unfold adapted_couple; repeat split. unfold ordered_Rlist; intros; simpl in H0; inversion H0; [ simpl; assumption | elim (le_Sn_O _ H2) ]. - simpl; unfold Rmin; case (Rle_dec a b); intro; - [ reflexivity | elim n; assumption ]. - simpl; unfold Rmax; case (Rle_dec a b); intro; - [ reflexivity | elim n; assumption ]. + simpl; unfold Rmin; decide (Rle_dec a b) with H; reflexivity. + simpl; unfold Rmax; decide (Rle_dec a b) with H; reflexivity. unfold constant_D_eq, open_interval; intros; simpl in H0; inversion H0; [ reflexivity | elim (le_Sn_O _ H3) ]. Qed. Lemma StepFun_P4 : forall a b c:R, IsStepFun (fct_cte c) a b. Proof. - intros; unfold IsStepFun; case (Rle_dec a b); intro. + intros; unfold IsStepFun; destruct (Rle_dec a b) as [Hle|Hnle]. apply existT with (cons a (cons b nil)); unfold is_subdivision; - apply existT with (cons c nil); apply (StepFun_P3 c r). + apply existT with (cons c nil); apply (StepFun_P3 c Hle). apply existT with (cons b (cons a nil)); unfold is_subdivision; apply existT with (cons c nil); apply StepFun_P2; apply StepFun_P3; auto with real. @@ -244,17 +241,15 @@ Lemma StepFun_P7 : Proof. unfold adapted_couple; intros; decompose [and] H0; clear H0; assert (H5 : Rmax a b = b). - unfold Rmax; case (Rle_dec a b); intro; - [ reflexivity | elim n; assumption ]. + unfold Rmax; decide (Rle_dec a b) with H; reflexivity. assert (H7 : r2 <= b). rewrite H5 in H2; rewrite <- H2; apply RList_P7; [ assumption | simpl; right; left; reflexivity ]. repeat split. apply RList_P4 with r1; assumption. - rewrite H5 in H2; unfold Rmin; case (Rle_dec r2 b); intro; - [ reflexivity | elim n; assumption ]. - unfold Rmax; case (Rle_dec r2 b); intro; - [ rewrite H5 in H2; rewrite <- H2; reflexivity | elim n; assumption ]. + rewrite H5 in H2; unfold Rmin; decide (Rle_dec r2 b) with H7; reflexivity. + unfold Rmax; decide (Rle_dec r2 b) with H7. + rewrite H5 in H2; rewrite <- H2; reflexivity. simpl in H4; simpl; apply INR_eq; apply Rplus_eq_reg_l with 1; do 2 rewrite (Rplus_comm 1); do 2 rewrite <- S_INR; rewrite H4; reflexivity. @@ -340,33 +335,28 @@ Proof. apply H6. rewrite <- Hyp_eq; rewrite H3 in H1; unfold adapted_couple in H1; decompose [and] H1; clear H1; simpl in H9; rewrite H9; - unfold Rmin; case (Rle_dec a b); intro; - [ reflexivity | elim n; assumption ]. + unfold Rmin; decide (Rle_dec a b) with H0; reflexivity. elim H6; clear H6; intros l' [lf' H6]; case (Req_dec t2 b); intro. exists (cons a (cons b nil)); exists (cons r1 nil); unfold adapted_couple_opt; unfold adapted_couple; repeat split. unfold ordered_Rlist; intros; simpl in H8; inversion H8; [ simpl; assumption | elim (le_Sn_O _ H10) ]. - simpl; unfold Rmin; case (Rle_dec a b); intro; - [ reflexivity | elim n; assumption ]. - simpl; unfold Rmax; case (Rle_dec a b); intro; - [ reflexivity | elim n; assumption ]. + simpl; unfold Rmin; decide (Rle_dec a b) with H0; reflexivity. + simpl; unfold Rmax; decide (Rle_dec a b) with H0; reflexivity. intros; simpl in H8; inversion H8. unfold constant_D_eq, open_interval; intros; simpl; simpl in H9; rewrite H3 in H1; unfold adapted_couple in H1; decompose [and] H1; apply (H16 0%nat). simpl; apply lt_O_Sn. unfold open_interval; simpl; rewrite H7; simpl in H13; - rewrite H13; unfold Rmin; case (Rle_dec a b); - intro; [ assumption | elim n; assumption ]. + rewrite H13; unfold Rmin; decide (Rle_dec a b) with H0; assumption. elim (le_Sn_O _ H10). intros; simpl in H8; elim (lt_n_O _ H8). intros; simpl in H8; inversion H8; [ simpl; assumption | elim (le_Sn_O _ H10) ]. assert (Hyp_min : Rmin t2 b = t2). - unfold Rmin; case (Rle_dec t2 b); intro; - [ reflexivity | elim n; assumption ]. + unfold Rmin; decide (Rle_dec t2 b) with H5; reflexivity. unfold adapted_couple in H6; elim H6; clear H6; intros; elim (RList_P20 _ (StepFun_P9 H6 H7)); intros s1 [s2 [s3 H9]]; induction lf' as [| r2 lf' Hreclf']. @@ -391,18 +381,16 @@ Proof. apply (H16 (S i)); simpl; assumption. simpl; simpl in H14; rewrite H14; reflexivity. simpl; simpl in H18; rewrite H18; unfold Rmax; - case (Rle_dec a b); case (Rle_dec t2 b); intros; reflexivity || elim n; - assumption. + decide (Rle_dec a b) with H0; decide (Rle_dec t2 b) with H5; reflexivity. simpl; simpl in H20; apply H20. intros; simpl in H1; unfold constant_D_eq, open_interval; intros; induction i as [| i Hreci]. - simpl; simpl in H6; case (total_order_T x t2); intro. - elim s; intro. + simpl; simpl in H6; destruct (total_order_T x t2) as [[Hlt|Heq]|Hgt]. apply (H17 0%nat); [ simpl; apply lt_O_Sn | unfold open_interval; simpl; elim H6; intros; split; assumption ]. - rewrite b0; assumption. + rewrite Heq; assumption. rewrite H10; apply (H22 0%nat); [ simpl; apply lt_O_Sn | unfold open_interval; simpl; replace s1 with t2; @@ -440,8 +428,7 @@ Proof. assumption. simpl; simpl in H19; apply H19. rewrite H9; simpl; simpl in H13; rewrite H13; unfold Rmax; - case (Rle_dec t2 b); case (Rle_dec a b); intros; reflexivity || elim n; - assumption. + decide (Rle_dec t2 b) with H5; decide (Rle_dec a b) with H0; reflexivity. rewrite H9; simpl; simpl in H15; rewrite H15; reflexivity. intros; simpl in H1; unfold constant_D_eq, open_interval; intros; induction i as [| i Hreci]. @@ -483,8 +470,7 @@ Proof. assumption. simpl; simpl in H18; apply H18. rewrite H9; simpl; simpl in H12; rewrite H12; unfold Rmax; - case (Rle_dec t2 b); case (Rle_dec a b); intros; reflexivity || elim n; - assumption. + decide (Rle_dec t2 b) with H5; decide (Rle_dec a b) with H0; reflexivity. rewrite H9; simpl; simpl in H14; rewrite H14; reflexivity. intros; simpl in H1; unfold constant_D_eq, open_interval; intros; induction i as [| i Hreci]. @@ -511,8 +497,7 @@ Proof. clear H1; clear H H7 H9; cut (Rmax a b = b); [ intro; rewrite H in H5; rewrite <- H5; apply RList_P7; [ assumption | simpl; right; left; reflexivity ] - | unfold Rmax; case (Rle_dec a b); intro; - [ reflexivity | elim n; assumption ] ]. + | unfold Rmax; decide (Rle_dec a b) with H0; reflexivity ]. Qed. Lemma StepFun_P11 : @@ -528,7 +513,7 @@ Proof. simpl in H10; simpl in H5; rewrite H10; rewrite H5; reflexivity. assert (H14 := H3 0%nat (lt_O_Sn _)); simpl in H14; elim H14; intro. assert (H15 := H7 0%nat (lt_O_Sn _)); simpl in H15; elim H15; intro. - rewrite <- H12 in H1; case (Rle_dec r1 s2); intro; try assumption. + rewrite <- H12 in H1; destruct (Rle_dec r1 s2) as [Hle|Hnle]; try assumption. assert (H16 : s2 < r1); auto with real. induction s3 as [| r0 s3 Hrecs3]. simpl in H9; rewrite H9 in H16; cut (r1 <= Rmax a b). @@ -662,12 +647,11 @@ Lemma StepFun_P13 : adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1) -> adapted_couple_opt f a b (cons s1 (cons s2 s3)) (cons r4 lf2) -> r1 <= s2. Proof. - intros; case (total_order_T a b); intro. - elim s; intro. - eapply StepFun_P11; [ apply a0 | apply H0 | apply H1 ]. + intros; destruct (total_order_T a b) as [[Hlt|Heq]|Hgt]. + eapply StepFun_P11; [ apply Hlt | apply H0 | apply H1 ]. elim H; assumption. eapply StepFun_P11; - [ apply r0 | apply StepFun_P2; apply H0 | apply StepFun_P12; apply H1 ]. + [ apply Hgt | apply StepFun_P2; apply H0 | apply StepFun_P12; apply H1 ]. Qed. Lemma StepFun_P14 : @@ -689,11 +673,9 @@ Proof. case (Req_dec a b); intro. rewrite (StepFun_P8 H2 H4); rewrite (StepFun_P8 H H4); reflexivity. assert (Hyp_min : Rmin a b = a). - unfold Rmin; case (Rle_dec a b); intro; - [ reflexivity | elim n; assumption ]. + unfold Rmin; decide (Rle_dec a b) with H1; reflexivity. assert (Hyp_max : Rmax a b = b). - unfold Rmax; case (Rle_dec a b); intro; - [ reflexivity | elim n; assumption ]. + unfold Rmax; decide (Rle_dec a b) with H1; reflexivity. elim (RList_P20 _ (StepFun_P9 H H4)); intros s1 [s2 [s3 H5]]; rewrite H5 in H; rewrite H5; induction lf1 as [| r3 lf1 Hreclf1]. unfold adapted_couple in H2; decompose [and] H2; @@ -883,8 +865,8 @@ Lemma StepFun_P15 : adapted_couple f a b l1 lf1 -> adapted_couple_opt f a b l2 lf2 -> Int_SF lf1 l1 = Int_SF lf2 l2. Proof. - intros; case (Rle_dec a b); intro; - [ apply (StepFun_P14 r H H0) + intros; destruct (Rle_dec a b) as [Hle|Hnle]; + [ apply (StepFun_P14 Hle H H0) | assert (H1 : b <= a); [ auto with real | eapply StepFun_P14; @@ -897,8 +879,8 @@ Lemma StepFun_P16 : exists l' : Rlist, (exists lf' : Rlist, adapted_couple_opt f a b l' lf'). Proof. - intros; case (Rle_dec a b); intro; - [ apply (StepFun_P10 r H) + intros; destruct (Rle_dec a b) as [Hle|Hnle]; + [ apply (StepFun_P10 Hle H) | assert (H1 : b <= a); [ auto with real | assert (H2 := StepFun_P10 H1 (StepFun_P2 H)); elim H2; @@ -961,9 +943,8 @@ Lemma StepFun_P21 : forall (a b:R) (f:R -> R) (l:Rlist), is_subdivision f a b l -> adapted_couple f a b l (FF l f). Proof. - intros; unfold adapted_couple; unfold is_subdivision in X; - unfold adapted_couple in X; elim X; clear X; intros; - decompose [and] p; clear p; repeat split; try assumption. + intros * (x & H & H1 & H0 & H2 & H4). + repeat split; try assumption. apply StepFun_P20; rewrite H2; apply lt_O_Sn. intros; assert (H5 := H4 _ H3); unfold constant_D_eq, open_interval in H5; unfold constant_D_eq, open_interval; intros; @@ -1003,11 +984,9 @@ Lemma StepFun_P22 : Proof. unfold is_subdivision; intros a b f g lf lg Hyp X X0; elim X; elim X0; clear X X0; intros lg0 p lf0 p0; assert (Hyp_min : Rmin a b = a). - unfold Rmin; case (Rle_dec a b); intro; - [ reflexivity | elim n; assumption ]. + unfold Rmin; decide (Rle_dec a b) with Hyp; reflexivity. assert (Hyp_max : Rmax a b = b). - unfold Rmax; case (Rle_dec a b); intro; - [ reflexivity | elim n; assumption ]. + unfold Rmax; decide (Rle_dec a b) with Hyp; reflexivity. apply existT with (FF (cons_ORlist lf lg) f); unfold adapted_couple in p, p0; decompose [and] p; decompose [and] p0; clear p p0; rewrite Hyp_min in H6; rewrite Hyp_min in H1; rewrite Hyp_max in H0; @@ -1221,13 +1200,13 @@ Proof. [ apply lt_n_S; assumption | symmetry ; apply S_pred with 0%nat; apply neq_O_lt; red; intro; rewrite <- H22 in H21; elim (lt_n_O _ H21) ]. - elim (Rle_dec (pos_Rl lf (S x0)) (pos_Rl (cons_ORlist lf lg) i)); intro. + elim (Rle_dec (pos_Rl lf (S x0)) (pos_Rl (cons_ORlist lf lg) i)); intro a0. assert (H23 : (S x0 <= x0)%nat). apply H20; unfold I; split; assumption. elim (le_Sn_n _ H23). assert (H23 : pos_Rl (cons_ORlist lf lg) i < pos_Rl lf (S x0)). auto with real. - clear b0; apply RList_P17; try assumption. + clear a0; apply RList_P17; try assumption. apply RList_P2; assumption. elim (RList_P9 lf lg (pos_Rl lf (S x0))); intros; apply H25; left; elim (RList_P3 lf (pos_Rl lf (S x0))); intros; apply H27; @@ -1255,11 +1234,9 @@ Lemma StepFun_P24 : Proof. unfold is_subdivision; intros a b f g lf lg Hyp X X0; elim X; elim X0; clear X X0; intros lg0 p lf0 p0; assert (Hyp_min : Rmin a b = a). - unfold Rmin; case (Rle_dec a b); intro; - [ reflexivity | elim n; assumption ]. + unfold Rmin; decide (Rle_dec a b) with Hyp; reflexivity. assert (Hyp_max : Rmax a b = b). - unfold Rmax; case (Rle_dec a b); intro; - [ reflexivity | elim n; assumption ]. + unfold Rmax; decide (Rle_dec a b) with Hyp; reflexivity. apply existT with (FF (cons_ORlist lf lg) g); unfold adapted_couple in p, p0; decompose [and] p; decompose [and] p0; clear p p0; rewrite Hyp_min in H1; rewrite Hyp_min in H6; rewrite Hyp_max in H0; @@ -1471,12 +1448,12 @@ Proof. apply lt_n_S; assumption. symmetry ; apply S_pred with 0%nat; apply neq_O_lt; red; intro; rewrite <- H22 in H21; elim (lt_n_O _ H21). - elim (Rle_dec (pos_Rl lg (S x0)) (pos_Rl (cons_ORlist lf lg) i)); intro. + elim (Rle_dec (pos_Rl lg (S x0)) (pos_Rl (cons_ORlist lf lg) i)); intro a0. assert (H23 : (S x0 <= x0)%nat); [ apply H20; unfold I; split; assumption | elim (le_Sn_n _ H23) ]. assert (H23 : pos_Rl (cons_ORlist lf lg) i < pos_Rl lg (S x0)). auto with real. - clear b0; apply RList_P17; try assumption; + clear a0; apply RList_P17; try assumption; [ apply RList_P2; assumption | elim (RList_P9 lf lg (pos_Rl lg (S x0))); intros; apply H25; right; elim (RList_P3 lg (pos_Rl lg (S x0))); intros; @@ -1652,7 +1629,7 @@ Lemma StepFun_P34 : a <= b -> Rabs (RiemannInt_SF f) <= RiemannInt_SF (mkStepFun (StepFun_P32 f)). Proof. - intros; unfold RiemannInt_SF; case (Rle_dec a b); intro. + intros; unfold RiemannInt_SF; decide (Rle_dec a b) with H. replace (Int_SF (subdivision_val (mkStepFun (StepFun_P32 f))) (subdivision (mkStepFun (StepFun_P32 f)))) with @@ -1663,7 +1640,6 @@ Proof. apply StepFun_P17 with (fun x:R => Rabs (f x)) a b; [ apply StepFun_P31; apply StepFun_P1 | apply (StepFun_P1 (mkStepFun (StepFun_P32 f))) ]. - elim n; assumption. Qed. Lemma StepFun_P35 : @@ -1741,24 +1717,21 @@ Lemma StepFun_P36 : (forall x:R, a < x < b -> f x <= g x) -> RiemannInt_SF f <= RiemannInt_SF g. Proof. - intros; unfold RiemannInt_SF; case (Rle_dec a b); intro. + intros; unfold RiemannInt_SF; decide (Rle_dec a b) with H. replace (Int_SF (subdivision_val f) (subdivision f)) with (Int_SF (FF l f) l). replace (Int_SF (subdivision_val g) (subdivision g)) with (Int_SF (FF l g) l). unfold is_subdivision in X; elim X; clear X; intros; unfold adapted_couple in p; decompose [and] p; clear p; assert (H5 : Rmin a b = a); - [ unfold Rmin; case (Rle_dec a b); intro; - [ reflexivity | elim n; assumption ] + [ unfold Rmin; decide (Rle_dec a b) with H; reflexivity | assert (H7 : Rmax a b = b); - [ unfold Rmax; case (Rle_dec a b); intro; - [ reflexivity | elim n; assumption ] + [ unfold Rmax; decide (Rle_dec a b) with H; reflexivity | rewrite H5 in H3; rewrite H7 in H2; eapply StepFun_P35 with a b; assumption ] ]. apply StepFun_P17 with (fe g) a b; [ apply StepFun_P21; assumption | apply StepFun_P1 ]. apply StepFun_P17 with (fe f) a b; [ apply StepFun_P21; assumption | apply StepFun_P1 ]. - elim n; assumption. Qed. Lemma StepFun_P37 : @@ -1819,8 +1792,7 @@ Proof. induction i as [| i Hreci]. simpl; rewrite H12; replace (Rmin r1 b) with r1. simpl in H0; rewrite <- H0; apply (H 0%nat); simpl; apply lt_O_Sn. - unfold Rmin; case (Rle_dec r1 b); intro; - [ reflexivity | elim n; assumption ]. + unfold Rmin; decide (Rle_dec r1 b) with H7; reflexivity. apply (H10 i); apply lt_S_n. replace (S (pred (Rlength lg))) with (Rlength lg). apply H9. @@ -1829,8 +1801,7 @@ Proof. simpl; assert (H14 : a <= b). rewrite <- H1; simpl in H0; rewrite <- H0; apply RList_P7; [ assumption | left; reflexivity ]. - unfold Rmin; case (Rle_dec a b); intro; - [ reflexivity | elim n; assumption ]. + unfold Rmin; decide (Rle_dec a b) with H14; reflexivity. assert (H14 : a <= b). rewrite <- H1; simpl in H0; rewrite <- H0; apply RList_P7; [ assumption | left; reflexivity ]. @@ -1838,14 +1809,13 @@ Proof. rewrite <- H11; induction lg as [| r0 lg Hreclg]. simpl in H13; discriminate. reflexivity. - unfold Rmax; case (Rle_dec a b); case (Rle_dec r1 b); intros; - reflexivity || elim n; assumption. + unfold Rmax; decide (Rle_dec a b) with H14; decide (Rle_dec r1 b) with H7; + reflexivity. simpl; rewrite H13; reflexivity. intros; simpl in H9; induction i as [| i Hreci]. unfold constant_D_eq, open_interval; simpl; intros; assert (H16 : Rmin r1 b = r1). - unfold Rmin; case (Rle_dec r1 b); intro; - [ reflexivity | elim n; assumption ]. + unfold Rmin; decide (Rle_dec r1 b) with H7; reflexivity. rewrite H16 in H12; rewrite H12 in H14; elim H14; clear H14; intros _ H14; unfold g'; case (Rle_dec r1 x); intro r3. elim (Rlt_irrefl _ (Rle_lt_trans _ _ _ r3 H14)). @@ -1862,9 +1832,9 @@ Proof. assert (H18 := H16 H17); unfold constant_D_eq, open_interval in H18; unfold constant_D_eq, open_interval; intros; assert (H19 := H18 _ H14); rewrite <- H19; unfold g'; - case (Rle_dec r1 x); intro. + case (Rle_dec r1 x) as [|[]]. reflexivity. - elim n; replace r1 with (Rmin r1 b). + replace r1 with (Rmin r1 b). rewrite <- H12; elim H14; clear H14; intros H14 _; left; apply Rle_lt_trans with (pos_Rl lg i); try assumption. apply RList_P5. @@ -1874,12 +1844,9 @@ Proof. apply lt_trans with (pred (Rlength lg)); try assumption. apply lt_pred_n_n; apply neq_O_lt; red; intro; rewrite <- H22 in H17; elim (lt_n_O _ H17). - unfold Rmin; case (Rle_dec r1 b); intro; - [ reflexivity | elim n0; assumption ]. + unfold Rmin; decide (Rle_dec r1 b) with H7; reflexivity. exists (mkStepFun H8); split. - simpl; unfold g'; case (Rle_dec r1 b); intro. - assumption. - elim n; assumption. + simpl; unfold g'; decide (Rle_dec r1 b) with H7; assumption. intros; simpl in H9; induction i as [| i Hreci]. unfold constant_D_eq, co_interval; simpl; intros; simpl in H0; rewrite H0; elim H10; clear H10; intros; unfold g'; @@ -1896,9 +1863,9 @@ Proof. assert (H12 := H10 H11); unfold constant_D_eq, co_interval in H12; unfold constant_D_eq, co_interval; intros; rewrite <- (H12 _ H13); simpl; unfold g'; - case (Rle_dec r1 x); intro. + case (Rle_dec r1 x) as [|[]]. reflexivity. - elim n; elim H13; clear H13; intros; + elim H13; clear H13; intros; apply Rle_trans with (pos_Rl (cons r1 l) i); try assumption; change (pos_Rl (cons r1 l) 0 <= pos_Rl (cons r1 l) i); elim (RList_P6 (cons r1 l)); intros; apply H15; @@ -1954,24 +1921,22 @@ Proof. unfold adapted_couple; decompose [and] H1; decompose [and] H2; clear H1 H2; repeat split. apply RList_P25; try assumption. - rewrite H10; rewrite H4; unfold Rmin, Rmax; case (Rle_dec a b); - case (Rle_dec b c); intros; - (right; reflexivity) || (elim n; left; assumption). + rewrite H10; rewrite H4; unfold Rmin, Rmax; case (Rle_dec a b) as [|[]]; + case (Rle_dec b c) as [|[]]; + (right; reflexivity) || (left; assumption). rewrite RList_P22. - rewrite H5; unfold Rmin, Rmax; case (Rle_dec a b); case (Rle_dec a c); - intros; + rewrite H5; unfold Rmin, Rmax; case (Rle_dec a c) as [|[]]; case (Rle_dec a b) as [|[]]; [ reflexivity - | elim n; apply Rle_trans with b; left; assumption - | elim n; left; assumption - | elim n0; left; assumption ]. + | left; assumption + | apply Rle_trans with b; left; assumption + | left; assumption ]. red; intro; rewrite H1 in H6; discriminate. rewrite RList_P24. - rewrite H9; unfold Rmin, Rmax; case (Rle_dec b c); case (Rle_dec a c); - intros; + rewrite H9; unfold Rmin, Rmax; case (Rle_dec a c) as [|[]]; case (Rle_dec b c) as [|[]]; [ reflexivity - | elim n; apply Rle_trans with b; left; assumption - | elim n; left; assumption - | elim n0; left; assumption ]. + | left; assumption + | apply Rle_trans with b; left; assumption + | left; assumption ]. red; intro; rewrite H1 in H11; discriminate. apply StepFun_P20. rewrite RList_P23; apply neq_O_lt; red; intro. @@ -2061,7 +2026,7 @@ Proof. assert (H16 : pos_Rl (cons_Rlist l1 l2) (S i) = b). rewrite RList_P29. rewrite H15; rewrite <- minus_n_n; rewrite H10; unfold Rmin; - case (Rle_dec b c); intro; [ reflexivity | elim n; left; assumption ]. + case (Rle_dec b c) as [|[]]; [ reflexivity | left; assumption ]. rewrite H15; apply le_n. induction l1 as [| r l1 Hrecl1]. simpl in H15; discriminate. @@ -2069,8 +2034,8 @@ Proof. assert (H17 : pos_Rl (cons_Rlist l1 l2) i = b). rewrite RList_P26. replace i with (pred (Rlength l1)); - [ rewrite H4; unfold Rmax; case (Rle_dec a b); intro; - [ reflexivity | elim n; left; assumption ] + [ rewrite H4; unfold Rmax; case (Rle_dec a b) as [|[]]; + [ reflexivity | left; assumption ] | rewrite H15; reflexivity ]. rewrite H15; apply lt_n_Sn. rewrite H16 in H2; rewrite H17 in H2; elim H2; intros; @@ -2095,8 +2060,8 @@ Proof. discriminate. clear Hrecl1; induction l1 as [| r0 l1 Hrecl1]. simpl in H5; simpl in H4; assert (H0 : Rmin a b < Rmax a b). - unfold Rmin, Rmax; case (Rle_dec a b); intro; - [ assumption | elim n; left; assumption ]. + unfold Rmin, Rmax; case (Rle_dec a b) as [|[]]; + [ assumption | left; assumption ]. rewrite <- H5 in H0; rewrite <- H4 in H0; elim (Rlt_irrefl _ H0). clear Hrecl1; simpl; repeat apply le_n_S; apply le_O_n. elim (RList_P20 _ H18); intros r1 [r2 [r3 H19]]; rewrite H19; @@ -2222,9 +2187,9 @@ Proof. | left _ => Int_SF lf3 l3 | right _ => - Int_SF lf3 l3 end. - case (Rle_dec a b); case (Rle_dec b c); case (Rle_dec a c); intros. - elim r1; intro. - elim r0; intro. + case (Rle_dec a b) as [Hle|Hnle]; case (Rle_dec b c) as [Hle'|Hnle']; case (Rle_dec a c) as [Hle''|Hnle'']. + elim Hle; intro. + elim Hle'; intro. replace (Int_SF lf3 l3) with (Int_SF (FF (cons_Rlist l1 l2) f) (cons_Rlist l1 l2)). replace (Int_SF lf1 l1) with (Int_SF (FF l1 f) l1). @@ -2232,8 +2197,7 @@ Proof. symmetry ; apply StepFun_P42. unfold adapted_couple in H1, H2; decompose [and] H1; decompose [and] H2; clear H1 H2; rewrite H11; rewrite H5; unfold Rmax, Rmin; - case (Rle_dec a b); case (Rle_dec b c); intros; reflexivity || elim n; - assumption. + decide (Rle_dec a b) with Hle; decide (Rle_dec b c) with Hle'; reflexivity. eapply StepFun_P17; [ apply StepFun_P21; unfold is_subdivision; split with lf2; apply H2; assumption @@ -2250,13 +2214,13 @@ Proof. rewrite Rplus_0_l; eapply StepFun_P17; [ apply H2 | rewrite H in H3; apply H3 ]. symmetry ; eapply StepFun_P8; [ apply H1 | assumption ]. - elim n; apply Rle_trans with b; assumption. + elim Hnle''; apply Rle_trans with b; assumption. apply Rplus_eq_reg_l with (Int_SF lf2 l2); replace (Int_SF lf2 l2 + (Int_SF lf1 l1 + - Int_SF lf2 l2)) with (Int_SF lf1 l1); [ idtac | ring ]. assert (H : c < b). auto with real. - elim r; intro. + elim Hle''; intro. rewrite Rplus_comm; replace (Int_SF lf1 l1) with (Int_SF (FF (cons_Rlist l3 l2) f) (cons_Rlist l3 l2)). @@ -2264,12 +2228,9 @@ Proof. replace (Int_SF lf2 l2) with (Int_SF (FF l2 f) l2). apply StepFun_P42. unfold adapted_couple in H2, H3; decompose [and] H2; decompose [and] H3; - clear H3 H2; rewrite H10; rewrite H6; unfold Rmax, Rmin; - case (Rle_dec a c); case (Rle_dec b c); intros; - [ elim n; assumption - | reflexivity - | elim n0; assumption - | elim n1; assumption ]. + clear H3 H2; rewrite H10; rewrite H6; unfold Rmax, Rmin. + decide (Rle_dec a c) with Hle''; decide (Rle_dec b c) with Hnle'; + reflexivity. eapply StepFun_P17; [ apply StepFun_P21; unfold is_subdivision; split with lf2; apply H2 | assumption ]. @@ -2284,7 +2245,7 @@ Proof. symmetry ; eapply StepFun_P8; [ apply H3 | assumption ]. replace (Int_SF lf2 l2) with (Int_SF lf3 l3 + Int_SF lf1 l1). ring. - elim r; intro. + elim Hle; intro. replace (Int_SF lf2 l2) with (Int_SF (FF (cons_Rlist l3 l1) f) (cons_Rlist l3 l1)). replace (Int_SF lf3 l3) with (Int_SF (FF l3 f) l3). @@ -2292,11 +2253,7 @@ Proof. symmetry ; apply StepFun_P42. unfold adapted_couple in H1, H3; decompose [and] H1; decompose [and] H3; clear H3 H1; rewrite H9; rewrite H5; unfold Rmax, Rmin; - case (Rle_dec a c); case (Rle_dec a b); intros; - [ elim n; assumption - | elim n1; assumption - | reflexivity - | elim n1; assumption ]. + decide (Rle_dec a c) with Hnle''; decide (Rle_dec a b) with Hle; reflexivity. eapply StepFun_P17; [ apply StepFun_P21; unfold is_subdivision; split with lf1; apply H1 | assumption ]. @@ -2316,7 +2273,7 @@ Proof. auto with real. replace (Int_SF lf2 l2) with (Int_SF lf3 l3 + Int_SF lf1 l1). ring. - rewrite Rplus_comm; elim r; intro. + rewrite Rplus_comm; elim Hle''; intro. replace (Int_SF lf2 l2) with (Int_SF (FF (cons_Rlist l1 l3) f) (cons_Rlist l1 l3)). replace (Int_SF lf3 l3) with (Int_SF (FF l3 f) l3). @@ -2324,11 +2281,8 @@ Proof. symmetry ; apply StepFun_P42. unfold adapted_couple in H1, H3; decompose [and] H1; decompose [and] H3; clear H3 H1; rewrite H11; rewrite H5; unfold Rmax, Rmin; - case (Rle_dec a c); case (Rle_dec a b); intros; - [ elim n; assumption - | reflexivity - | elim n0; assumption - | elim n1; assumption ]. + decide (Rle_dec a c) with Hle''; decide (Rle_dec a b) with Hnle; + reflexivity. eapply StepFun_P17; [ apply StepFun_P21; unfold is_subdivision; split with lf1; apply H1 | assumption ]. @@ -2346,7 +2300,7 @@ Proof. auto with real. replace (Int_SF lf1 l1) with (Int_SF lf2 l2 + Int_SF lf3 l3). ring. - elim r; intro. + elim Hle'; intro. replace (Int_SF lf1 l1) with (Int_SF (FF (cons_Rlist l2 l3) f) (cons_Rlist l2 l3)). replace (Int_SF lf3 l3) with (Int_SF (FF l3 f) l3). @@ -2354,11 +2308,8 @@ Proof. symmetry ; apply StepFun_P42. unfold adapted_couple in H2, H3; decompose [and] H2; decompose [and] H3; clear H3 H2; rewrite H11; rewrite H5; unfold Rmax, Rmin; - case (Rle_dec a c); case (Rle_dec b c); intros; - [ elim n; assumption - | elim n1; assumption - | reflexivity - | elim n1; assumption ]. + decide (Rle_dec a c) with Hnle''; decide (Rle_dec b c) with Hle'; + reflexivity. eapply StepFun_P17; [ apply StepFun_P21; unfold is_subdivision; split with lf2; apply H2 | assumption ]. @@ -2371,8 +2322,8 @@ Proof. replace (Int_SF lf2 l2) with 0. rewrite Rplus_0_l; eapply StepFun_P17; [ apply H3 | rewrite H0 in H1; apply H1 ]. - symmetry ; eapply StepFun_P8; [ apply H2 | assumption ]. - elim n; apply Rle_trans with a; try assumption. + symmetry; eapply StepFun_P8; [ apply H2 | assumption ]. + elim Hnle'; apply Rle_trans with a; try assumption. auto with real. assert (H : c < b). auto with real. @@ -2387,11 +2338,8 @@ Proof. symmetry ; apply StepFun_P42. unfold adapted_couple in H2, H1; decompose [and] H2; decompose [and] H1; clear H1 H2; rewrite H11; rewrite H5; unfold Rmax, Rmin; - case (Rle_dec a b); case (Rle_dec b c); intros; - [ elim n1; assumption - | elim n1; assumption - | elim n0; assumption - | reflexivity ]. + decide (Rle_dec a b) with Hnle; decide (Rle_dec b c) with Hnle'; + reflexivity. eapply StepFun_P17; [ apply StepFun_P21; unfold is_subdivision; split with lf2; apply H2 | assumption ]. @@ -2463,10 +2411,8 @@ Proof. replace a with (Rmin a b). pattern b at 2; replace b with (Rmax a b). rewrite <- H2; rewrite H3; reflexivity. - unfold Rmax; case (Rle_dec a b); intro; - [ reflexivity | elim n; assumption ]. - unfold Rmin; case (Rle_dec a b); intro; - [ reflexivity | elim n; assumption ]. + unfold Rmax; decide (Rle_dec a b) with H7; reflexivity. + unfold Rmin; decide (Rle_dec a b) with H7; reflexivity. split with (cons r nil); split with lf1; assert (H2 : c = b). rewrite H1 in H0; elim H0; intros; apply Rle_antisym; assumption. rewrite H2; assumption. @@ -2475,20 +2421,18 @@ Proof. discriminate. clear Hreclf1; assert (H1 : {c <= r1} + {r1 < c}). case (Rle_dec c r1); intro; [ left; assumption | right; auto with real ]. - elim H1; intro. + elim H1; intro a0. split with (cons r (cons c nil)); split with (cons r3 nil); unfold adapted_couple in H; decompose [and] H; clear H; assert (H6 : r = a). - simpl in H4; rewrite H4; unfold Rmin; case (Rle_dec a b); intro; + simpl in H4; rewrite H4; unfold Rmin; case (Rle_dec a b) as [|[]]; [ reflexivity - | elim n; elim H0; intros; apply Rle_trans with c; assumption ]. + | elim H0; intros; apply Rle_trans with c; assumption ]. elim H0; clear H0; intros; unfold adapted_couple; repeat split. rewrite H6; unfold ordered_Rlist; intros; simpl in H8; inversion H8; [ simpl; assumption | elim (le_Sn_O _ H10) ]. - simpl; unfold Rmin; case (Rle_dec a c); intro; - [ assumption | elim n; assumption ]. - simpl; unfold Rmax; case (Rle_dec a c); intro; - [ reflexivity | elim n; assumption ]. + simpl; unfold Rmin; decide (Rle_dec a c) with H; assumption. + simpl; unfold Rmax; decide (Rle_dec a c) with H; reflexivity. unfold constant_D_eq, open_interval; intros; simpl in H8; inversion H8. simpl; assert (H10 := H7 0%nat); @@ -2508,8 +2452,8 @@ Proof. assert (H14 : a <= b). elim H0; intros; apply Rle_trans with c; assumption. assert (H16 : r = a). - simpl in H7; rewrite H7; unfold Rmin; case (Rle_dec a b); intro; - [ reflexivity | elim n; assumption ]. + simpl in H7; rewrite H7; unfold Rmin; decide (Rle_dec a b) with H14; + reflexivity. induction l1' as [| r4 l1' Hrecl1']. simpl in H13; discriminate. clear Hrecl1'; unfold adapted_couple; repeat split. @@ -2517,18 +2461,18 @@ Proof. simpl; replace r4 with r1. apply (H5 0%nat). simpl; apply lt_O_Sn. - simpl in H12; rewrite H12; unfold Rmin; case (Rle_dec r1 c); intro; - [ reflexivity | elim n; left; assumption ]. + simpl in H12; rewrite H12; unfold Rmin; case (Rle_dec r1 c) as [|[]]; + [ reflexivity | left; assumption ]. apply (H9 i); simpl; apply lt_S_n; assumption. - simpl; unfold Rmin; case (Rle_dec a c); intro; - [ assumption | elim n; elim H0; intros; assumption ]. + simpl; unfold Rmin; case (Rle_dec a c) as [|[]]; + [ assumption | elim H0; intros; assumption ]. replace (Rmax a c) with (Rmax r1 c). rewrite <- H11; reflexivity. - unfold Rmax; case (Rle_dec r1 c); case (Rle_dec a c); intros; - [ reflexivity - | elim n; elim H0; intros; assumption - | elim n; left; assumption - | elim n0; left; assumption ]. + unfold Rmax; case (Rle_dec a c) as [|[]]; case (Rle_dec r1 c) as [|[]]; + [ reflexivity + | left; assumption + | elim H0; intros; assumption + | left; assumption ]. simpl; simpl in H13; rewrite H13; reflexivity. intros; simpl in H; unfold constant_D_eq, open_interval; intros; induction i as [| i Hreci]. @@ -2539,8 +2483,8 @@ Proof. elim H4; clear H4; intros; split; try assumption; replace r1 with r4. assumption. - simpl in H12; rewrite H12; unfold Rmin; case (Rle_dec r1 c); intro; - [ reflexivity | elim n; left; assumption ]. + simpl in H12; rewrite H12; unfold Rmin; case (Rle_dec r1 c) as [|[]]; + [ reflexivity | left; assumption ]. clear Hreci; simpl; apply H15. simpl; apply lt_S_n; assumption. unfold open_interval; apply H4. @@ -2578,10 +2522,8 @@ Proof. replace a with (Rmin a b). pattern b at 2; replace b with (Rmax a b). rewrite <- H2; rewrite H3; reflexivity. - unfold Rmax; case (Rle_dec a b); intro; - [ reflexivity | elim n; assumption ]. - unfold Rmin; case (Rle_dec a b); intro; - [ reflexivity | elim n; assumption ]. + unfold Rmax; decide (Rle_dec a b) with H7; reflexivity. + unfold Rmin; decide (Rle_dec a b) with H7; reflexivity. split with (cons r nil); split with lf1; assert (H2 : c = b). rewrite H1 in H0; elim H0; intros; apply Rle_antisym; assumption. rewrite <- H2 in H1; rewrite <- H1; assumption. @@ -2590,22 +2532,22 @@ Proof. discriminate. clear Hreclf1; assert (H1 : {c <= r1} + {r1 < c}). case (Rle_dec c r1); intro; [ left; assumption | right; auto with real ]. - elim H1; intro. + elim H1; intro a0. split with (cons c (cons r1 r2)); split with (cons r3 lf1); unfold adapted_couple in H; decompose [and] H; clear H; unfold adapted_couple; repeat split. unfold ordered_Rlist; intros; simpl in H; induction i as [| i Hreci]. simpl; assumption. clear Hreci; apply (H2 (S i)); simpl; assumption. - simpl; unfold Rmin; case (Rle_dec c b); intro; - [ reflexivity | elim n; elim H0; intros; assumption ]. + simpl; unfold Rmin; case (Rle_dec c b) as [|[]]; + [ reflexivity | elim H0; intros; assumption ]. replace (Rmax c b) with (Rmax a b). rewrite <- H3; reflexivity. - unfold Rmax; case (Rle_dec a b); case (Rle_dec c b); intros; + unfold Rmax; case (Rle_dec c b) as [|[]]; case (Rle_dec a b) as [|[]]; [ reflexivity - | elim n; elim H0; intros; assumption - | elim n; elim H0; intros; apply Rle_trans with c; assumption - | elim n0; elim H0; intros; apply Rle_trans with c; assumption ]. + | elim H0; intros; apply Rle_trans with c; assumption + | elim H0; intros; assumption + | elim H0; intros; apply Rle_trans with c; assumption ]. simpl; simpl in H5; apply H5. intros; simpl in H; induction i as [| i Hreci]. unfold constant_D_eq, open_interval; intros; simpl; @@ -2615,9 +2557,9 @@ Proof. intros; split; try assumption; apply Rle_lt_trans with c; try assumption; replace r with a. elim H0; intros; assumption. - simpl in H4; rewrite H4; unfold Rmin; case (Rle_dec a b); intros; + simpl in H4; rewrite H4; unfold Rmin; case (Rle_dec a b) as [|[]]; [ reflexivity - | elim n; elim H0; intros; apply Rle_trans with c; assumption ]. + | elim H0; intros; apply Rle_trans with c; assumption ]. clear Hreci; apply (H7 (S i)); simpl; assumption. cut (adapted_couple f r1 b (cons r1 r2) lf1). cut (r1 <= c <= b). diff --git a/theories/Reals/Rlimit.v b/theories/Reals/Rlimit.v index c3020611..c8887dfb 100644 --- a/theories/Reals/Rlimit.v +++ b/theories/Reals/Rlimit.v @@ -1,6 +1,6 @@ (************************************************************************) (* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *) +(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *) (* \VV/ **************************************************************) (* // * This file is distributed under the terms of the *) (* * GNU Lesser General Public License Version 2.1 *) @@ -164,7 +164,7 @@ Definition limit_in (X X':Metric_Space) (f:Base X -> Base X') eps > 0 -> exists alp : R, alp > 0 /\ - (forall x:Base X, D x /\ dist X x x0 < alp -> dist X' (f x) l < eps). + (forall x:Base X, D x /\ X.(dist) x x0 < alp -> X'.(dist) (f x) l < eps). (*******************************) (** ** R is a metric space *) @@ -174,6 +174,8 @@ Definition limit_in (X X':Metric_Space) (f:Base X -> Base X') Definition R_met : Metric_Space := Build_Metric_Space R R_dist R_dist_pos R_dist_sym R_dist_refl R_dist_tri. +Declare Equivalent Keys dist R_dist. + (*******************************) (** * Limit 1 arg *) (*******************************) @@ -191,9 +193,9 @@ Lemma tech_limit : Proof. intros f D l x0 H H0. case (Rabs_pos (f x0 - l)); intros H1. - absurd (dist R_met (f x0) l < dist R_met (f x0) l). + absurd (R_met.(@dist) (f x0) l < R_met.(@dist) (f x0) l). apply Rlt_irrefl. - case (H0 (dist R_met (f x0) l)); auto. + case (H0 (R_met.(@dist) (f x0) l)); auto. intros alpha1 [H2 H3]; apply H3; auto; split; auto. case (dist_refl R_met x0 x0); intros Hr1 Hr2; rewrite Hr2; auto. case (dist_refl R_met (f x0) l); intros Hr1 Hr2; symmetry; auto. @@ -312,7 +314,7 @@ Proof. rewrite (Rplus_comm 1 (Rabs l)); unfold Rgt; apply Rle_lt_0_plus_1; exact (Rabs_pos l). unfold R_dist in H9; - apply (Rplus_lt_reg_r (- Rabs l) (Rabs (f x2)) (1 + Rabs l)). + apply (Rplus_lt_reg_l (- Rabs l) (Rabs (f x2)) (1 + Rabs l)). rewrite <- (Rplus_assoc (- Rabs l) 1 (Rabs l)); rewrite (Rplus_comm (- Rabs l) 1); rewrite (Rplus_assoc 1 (- Rabs l) (Rabs l)); rewrite (Rplus_opp_l (Rabs l)); @@ -345,18 +347,19 @@ Lemma single_limit : adhDa D x0 -> limit1_in f D l x0 -> limit1_in f D l' x0 -> l = l'. Proof. unfold limit1_in; unfold limit_in; intros. + simpl in *. cut (forall eps:R, eps > 0 -> dist R_met l l' < 2 * eps). - clear H0 H1; unfold dist; unfold R_met; unfold R_dist; - unfold Rabs; case (Rcase_abs (l - l')); intros. + clear H0 H1; unfold dist in |- *; unfold R_met; unfold R_dist in |- *; + unfold Rabs; case (Rcase_abs (l - l')) as [Hlt|Hge]; intros. cut (forall eps:R, eps > 0 -> - (l - l') < eps). intro; generalize (prop_eps (- (l - l')) H1); intro; - generalize (Ropp_gt_lt_0_contravar (l - l') r); intro; + generalize (Ropp_gt_lt_0_contravar (l - l') Hlt); intro; unfold Rgt in H3; generalize (Rgt_not_le (- (l - l')) 0 H3); intro; exfalso; auto. intros; cut (eps * / 2 > 0). intro; generalize (H0 (eps * / 2) H2); rewrite (Rmult_comm eps (/ 2)); rewrite <- (Rmult_assoc 2 (/ 2) eps); rewrite (Rinv_r 2). - elim (Rmult_ne eps); intros a b; rewrite b; clear a b; trivial. + elim (Rmult_ne eps); intros a b; rewrite b; clear a b; trivial. apply (Rlt_dichotomy_converse 2 0); right; generalize Rlt_0_1; intro; unfold Rgt; generalize (Rplus_lt_compat_l 1 0 1 H3); intro; elim (Rplus_ne 1); intros a b; rewrite a in H4; @@ -374,7 +377,7 @@ Proof. intros a b; clear b; apply (Rminus_diag_uniq l l'); apply a; split. assumption. - apply (Rge_le (l - l') 0 r). + apply (Rge_le (l - l') 0 Hge). intros; cut (eps * / 2 > 0). intro; generalize (H0 (eps * / 2) H2); rewrite (Rmult_comm eps (/ 2)); rewrite <- (Rmult_assoc 2 (/ 2) eps); rewrite (Rinv_r 2). diff --git a/theories/Reals/Rlogic.v b/theories/Reals/Rlogic.v index 14dea1c6..07792942 100644 --- a/theories/Reals/Rlogic.v +++ b/theories/Reals/Rlogic.v @@ -1,261 +1,137 @@ (************************************************************************) (* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *) +(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *) (* \VV/ **************************************************************) (* // * This file is distributed under the terms of the *) (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) -(** * This module proves some logical properties of the axiomatics of Reals +(** This module proves some logical properties of the axiomatic of Reals. -1. Decidablity of arithmetical statements from - the axiom that the order of the real numbers is decidable. - -2. Derivability of the archimedean "axiom" +- Decidability of arithmetical statements. +- Derivability of the Archimedean "axiom". +- Decidability of negated formulas. *) -(** 1- Proof of the decidablity of arithmetical statements from -excluded middle and the axiom that the order of the real numbers is -decidable. *) +Require Import RIneq. -(** Assuming a decidable predicate [P n], A series is constructed whose -[n]th term is 1/2^n if [P n] holds and 0 otherwise. This sum reaches 2 -only if [P n] holds for all [n], otherwise the sum is less than 2. -Comparing the sum to 2 decides if [forall n, P n] or [~forall n, P n] *) +(** * Decidability of arithmetical statements *) (** One can iterate this lemma and use classical logic to decide any statement in the arithmetical hierarchy. *) -(** Contributed by Cezary Kaliszyk and Russell O'Connor *) - -Require Import ConstructiveEpsilon. -Require Import Rfunctions. -Require Import PartSum. -Require Import SeqSeries. -Require Import RiemannInt. -Require Import Fourier. - Section Arithmetical_dec. Variable P : nat -> Prop. Hypothesis HP : forall n, {P n} + {~P n}. -Let ge_fun_sums_ge_lemma : (forall (m n : nat) (f : nat -> R), (lt m n) -> (forall i : nat, 0 <= f i) -> sum_f_R0 f m <= sum_f_R0 f n). -Proof. -intros m n f mn fpos. -replace (sum_f_R0 f m) with (sum_f_R0 f m + 0) by ring. -rewrite (tech2 f m n mn). -apply Rplus_le_compat_l. - induction (n - S m)%nat; simpl in *. - apply fpos. -replace 0 with (0 + 0) by ring. -apply (Rplus_le_compat _ _ _ _ IHn0 (fpos (S (m + S n0)%nat))). -Qed. - -Let ge_fun_sums_ge : (forall (m n : nat) (f : nat -> R), (le m n) -> (forall i : nat, 0 <= f i) -> sum_f_R0 f m <= sum_f_R0 f n). -Proof. -intros m n f mn pos. - elim (le_lt_or_eq _ _ mn). - intro; apply ge_fun_sums_ge_lemma; assumption. -intro H; rewrite H; auto with *. -Qed. - -Let f:=fun n => (if HP n then (1/2)^n else 0)%R. - -Lemma cauchy_crit_geometric_dec_fun : Cauchy_crit_series f. +Lemma sig_forall_dec : {n | ~P n} + {forall n, P n}. Proof. -intros e He. -assert (X:(Pser (fun n:nat => 1) (1/2) (/ (1 - (1/2))))%R). - apply GP_infinite. - apply Rabs_def1; fourier. -assert (He':e/2 > 0) by fourier. -destruct (X _ He') as [N HN]. -clear X. -exists N. -intros n m Hn Hm. -replace e with (e/2 + e/2)%R by field. -set (g:=(fun n0 : nat => 1 * (1 / 2) ^ n0)) in *. -assert (R_dist (sum_f_R0 g n) (sum_f_R0 g m) < e / 2 + e / 2). - apply Rle_lt_trans with (R_dist (sum_f_R0 g n) 2+R_dist 2 (sum_f_R0 g m))%R. - apply R_dist_tri. - replace (/(1 - 1/2)) with 2 in HN by field. - cut (forall n, (n >= N)%nat -> R_dist (sum_f_R0 g n) 2 < e/2)%R. - intros Z. - apply Rplus_lt_compat. - apply Z; assumption. - rewrite R_dist_sym. - apply Z; assumption. - clear - HN He. - intros n Hn. - apply HN. - auto. -eapply Rle_lt_trans;[|apply H]. -clear -ge_fun_sums_ge n. -cut (forall n m, (m <= n)%nat -> R_dist (sum_f_R0 f n) (sum_f_R0 f m) <= R_dist (sum_f_R0 g n) (sum_f_R0 g m)). - intros H. - destruct (le_lt_dec m n). - apply H; assumption. - rewrite R_dist_sym. - rewrite (R_dist_sym (sum_f_R0 g n)). - apply H; auto with *. -clear n m. -intros n m Hnm. -unfold R_dist. -cut (forall i : nat, (1 / 2) ^ i >= 0). intro RPosPow. -rewrite Rabs_pos_eq. - rewrite Rabs_pos_eq. - cut (sum_f_R0 g m - sum_f_R0 f m <= sum_f_R0 g n - sum_f_R0 f n). - intros; fourier. - do 2 rewrite <- minus_sum. - apply (ge_fun_sums_ge m n (fun i : nat => g i - f i) Hnm). - intro i. - unfold f, g. - elim (HP i); intro; ring_simplify; auto with *. - cut (sum_f_R0 g m <= sum_f_R0 g n). - intro; fourier. - apply (ge_fun_sums_ge m n g Hnm). - intro. unfold g. - ring_simplify. - apply Rge_le. - apply RPosPow. - cut (sum_f_R0 f m <= sum_f_R0 f n). - intro; fourier. - apply (ge_fun_sums_ge m n f Hnm). - intro; unfold f. - elim (HP i); intro; simpl. - apply Rge_le. - apply RPosPow. - auto with *. -intro i. -apply Rle_ge. -apply pow_le. -fourier. -Qed. - -Lemma forall_dec : {forall n, P n} + {~forall n, P n}. -Proof. -destruct (cv_cauchy_2 _ cauchy_crit_geometric_dec_fun). - cut (2 <= x <-> forall n : nat, P n). - intro H. - elim (Rle_dec 2 x); intro X. - left; tauto. - right; tauto. -assert (A:Rabs(1/2) < 1) by (apply Rabs_def1; fourier). -assert (A0:=(GP_infinite (1/2) A)). -symmetry. - split; intro. - replace 2 with (/ (1 - (1 / 2))) by field. - unfold Pser, infinite_sum in A0. - eapply Rle_cv_lim;[|unfold Un_cv; apply A0 |apply u]. - intros n. - clear -n H. - induction n; unfold f;simpl. - destruct (HP 0); auto with *. - elim n; auto. - apply Rplus_le_compat; auto. - destruct (HP (S n)); auto with *. - elim n0; auto. -intros n. -destruct (HP n); auto. -elim (RIneq.Rle_not_lt _ _ H). -assert (B:0< (1/2)^n). - apply pow_lt. - fourier. -apply Rle_lt_trans with (2-(1/2)^n);[|fourier]. -replace (/(1-1/2))%R with 2 in A0 by field. -set (g:= fun m => if (eq_nat_dec m n) then (1/2)^n else 0). -assert (Z: Un_cv (fun N : nat => sum_f_R0 g N) ((1/2)^n)). - intros e He. - exists n. - intros a Ha. - replace (sum_f_R0 g a) with ((1/2)^n). - rewrite (R_dist_eq); assumption. - symmetry. - cut (forall a : nat, ((a >= n)%nat -> sum_f_R0 g a = (1 / 2) ^ n) /\ ((a < n)%nat -> sum_f_R0 g a = 0))%R. - intros H0. - destruct (H0 a). - auto. - clear - g. - induction a. - split; - intros H; - simpl; unfold g; - destruct (eq_nat_dec 0 n) as [t|f]; try reflexivity. - elim f; auto with *. - exfalso; omega. - destruct IHa as [IHa0 IHa1]. - split; - intros H; - simpl; unfold g at 2; - destruct (eq_nat_dec (S a) n). - rewrite IHa1. - ring. - omega. - ring_simplify. - apply IHa0. - omega. - exfalso; omega. - ring_simplify. - apply IHa1. - omega. -assert (C:=CV_minus _ _ _ _ A0 Z). -eapply Rle_cv_lim;[|apply u |apply C]. -clear - n0 B. -intros m. -simpl. -induction m. - simpl. - unfold f, g. - destruct (eq_nat_dec 0 n). - destruct (HP 0). - elim n0. - congruence. - clear -n. - induction n; simpl; fourier. - destruct (HP); simpl; fourier. -cut (f (S m) <= 1 * ((1 / 2) ^ (S m)) - g (S m)). - intros L. - eapply Rle_trans. +assert (Hi: (forall n, 0 < INR n + 1)%R). + intros n. + apply Rle_lt_0_plus_1, pos_INR. +set (u n := (if HP n then 0 else / (INR n + 1))%R). +assert (Bu: forall n, (u n <= 1)%R). + intros n. + unfold u. + case HP ; intros _. + apply Rle_0_1. + rewrite <- S_INR, <- Rinv_1. + apply Rinv_le_contravar with (1 := Rlt_0_1). + apply (le_INR 1), le_n_S, le_0_n. +set (E y := exists n, y = u n). +destruct (completeness E) as [l [ub lub]]. + exists R1. + intros y [n ->]. + apply Bu. + exists (u O). + now exists O. +assert (Hnp: forall n, not (P n) -> ((/ (INR n + 1) <= l)%R)). + intros n Hp. + apply ub. + exists n. + unfold u. + now destruct (HP n). +destruct (Rle_lt_dec l 0) as [Hl|Hl]. + right. + intros n. + destruct (HP n) as [H|H]. + exact H. + exfalso. + apply Rle_not_lt with (1 := Hl). + apply Rlt_le_trans with (/ (INR n + 1))%R. + now apply Rinv_0_lt_compat. + now apply Hnp. +left. +set (N := Zabs_nat (up (/l) - 2)). +assert (H1l: (1 <= /l)%R). + rewrite <- Rinv_1. + apply Rinv_le_contravar with (1 := Hl). + apply lub. + now intros y [m ->]. +assert (HN: (INR N + 1 = IZR (up (/ l)) - 1)%R). + unfold N. + rewrite INR_IZR_INZ. + rewrite inj_Zabs_nat. + replace (IZR (up (/ l)) - 1)%R with (IZR (up (/ l) - 2) + 1)%R. + apply (f_equal (fun v => IZR v + 1)%R). + apply Zabs_eq. + apply Zle_minus_le_0. + apply (Zlt_le_succ 1). + apply lt_IZR. + apply Rle_lt_trans with (1 := H1l). + apply archimed. + rewrite minus_IZR. simpl. - apply Rplus_le_compat. - apply IHm. - apply L. - simpl; fourier. -unfold f, g. -destruct (eq_nat_dec (S m) n). - destruct (HP (S m)). - elim n0. - congruence. - rewrite e. - fourier. -destruct (HP (S m)). - fourier. + ring. +assert (Hl': (/ (INR (S N) + 1) < l)%R). + rewrite <- (Rinv_involutive l) by now apply Rgt_not_eq. + apply Rinv_1_lt_contravar with (1 := H1l). + rewrite S_INR. + rewrite HN. + ring_simplify. + apply archimed. +exists N. +intros H. +apply Rle_not_lt with (2 := Hl'). +apply lub. +intros y [n ->]. +unfold u. +destruct (HP n) as [_|Hp]. + apply Rlt_le. + now apply Rinv_0_lt_compat. +apply Rinv_le_contravar. +apply Hi. +apply Rplus_le_compat_r. +apply le_INR. +destruct (le_or_lt n N) as [Hn|Hn]. + 2: now apply lt_le_S. +exfalso. +destruct (le_lt_or_eq _ _ Hn) as [Hn'| ->]. +2: now apply Hp. +apply Rlt_not_le with (2 := Hnp _ Hp). +rewrite <- (Rinv_involutive l) by now apply Rgt_not_eq. +apply Rinv_1_lt_contravar. +rewrite <- S_INR. +apply (le_INR 1), le_n_S, le_0_n. +apply Rlt_le_trans with (INR N + 1)%R. +apply Rplus_lt_compat_r. +now apply lt_INR. +rewrite HN. +apply Rplus_le_reg_r with (-/l + 1)%R. ring_simplify. -apply pow_le. -fourier. -Qed. - -Lemma sig_forall_dec : {n | ~P n}+{forall n, P n}. -Proof. -destruct forall_dec. - right; assumption. -left. -apply constructive_indefinite_ground_description_nat; auto. - clear - HP. - firstorder. -apply Classical_Pred_Type.not_all_ex_not. -assumption. +apply archimed. Qed. End Arithmetical_dec. -(** 2- Derivability of the Archimedean axiom *) +(** * Derivability of the Archimedean axiom *) -(* This is a standard proof (it has been taken from PlanetMath). It is +(** This is a standard proof (it has been taken from PlanetMath). It is formulated negatively so as to avoid the need for classical -logic. Using a proof of {n | ~P n}+{forall n, P n} (the one above or a -variant of it that does not need classical axioms) , we can in -principle also derive [up] and its [specification] *) +logic. Using a proof of [{n | ~P n}+{forall n, P n}], we can in +principle also derive [up] and its specification. The proof above +cannot be used for that purpose, since it relies on the [archimed] axiom. *) Theorem not_not_archimedean : forall r : R, ~ (forall n : nat, (INR n <= r)%R). @@ -296,3 +172,33 @@ rewrite (Rplus_comm (INR n) 0) in H6. rewrite Rplus_0_l in H6. assumption. Qed. + +(** * Decidability of negated formulas *) + +Lemma sig_not_dec : forall P : Prop, {not (not P)} + {not P}. +Proof. +intros P. +set (E := fun x => x = R0 \/ (x = R1 /\ P)). +destruct (completeness E) as [x H]. + exists R1. + intros x [->|[-> _]]. + apply Rle_0_1. + apply Rle_refl. + exists R0. + now left. +destruct (Rle_lt_dec 1 x) as [H'|H']. +- left. + intros HP. + elim Rle_not_lt with (1 := H'). + apply Rle_lt_trans with (2 := Rlt_0_1). + apply H. + intros y [->|[_ Hy]]. + apply Rle_refl. + now elim HP. +- right. + intros HP. + apply Rlt_not_le with (1 := H'). + apply H. + right. + now split. +Qed. diff --git a/theories/Reals/Rminmax.v b/theories/Reals/Rminmax.v index 9121ccc2..ba1fe90f 100644 --- a/theories/Reals/Rminmax.v +++ b/theories/Reals/Rminmax.v @@ -1,6 +1,6 @@ (************************************************************************) (* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *) +(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *) (* \VV/ **************************************************************) (* // * This file is distributed under the terms of the *) (* * GNU Lesser General Public License Version 2.1 *) diff --git a/theories/Reals/Rpow_def.v b/theories/Reals/Rpow_def.v index 0116e29a..1d697f3c 100644 --- a/theories/Reals/Rpow_def.v +++ b/theories/Reals/Rpow_def.v @@ -1,6 +1,6 @@ (************************************************************************) (* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *) +(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *) (* \VV/ **************************************************************) (* // * This file is distributed under the terms of the *) (* * GNU Lesser General Public License Version 2.1 *) diff --git a/theories/Reals/Rpower.v b/theories/Reals/Rpower.v index 014d7025..e30ea334 100644 --- a/theories/Reals/Rpower.v +++ b/theories/Reals/Rpower.v @@ -1,6 +1,6 @@ (************************************************************************) (* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *) +(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *) (* \VV/ **************************************************************) (* // * This file is distributed under the terms of the *) (* * GNU Lesser General Public License Version 2.1 *) @@ -20,8 +20,10 @@ Require Import Ranalysis1. Require Import Exp_prop. Require Import Rsqrt_def. Require Import R_sqrt. +Require Import Sqrt_reg. Require Import MVT. Require Import Ranalysis4. +Require Import Fourier. Local Open Scope R_scope. Lemma P_Rmin : forall (P:R -> Prop) (x y:R), P x -> P y -> P (Rmin x y). @@ -43,7 +45,7 @@ Proof. rewrite Rmult_1_r; rewrite <- (Rmult_comm 3); rewrite <- Rmult_assoc; rewrite <- Rinv_l_sym. rewrite Rmult_1_l; replace (/ exp 1) with (exp (-1)). - unfold exp; case (exist_exp (-1)); intros; simpl; + unfold exp; case (exist_exp (-1)) as (?,e); simpl in |- *; unfold exp_in in e; assert (H := alternated_series_ineq (fun i:nat => / INR (fact i)) x 1). cut @@ -137,7 +139,7 @@ Qed. Lemma exp_ineq1 : forall x:R, 0 < x -> 1 + x < exp x. Proof. - intros; apply Rplus_lt_reg_r with (- exp 0); rewrite <- (Rplus_comm (exp x)); + intros; apply Rplus_lt_reg_l with (- exp 0); rewrite <- (Rplus_comm (exp x)); assert (H0 := MVT_cor1 exp 0 x derivable_exp H); elim H0; intros; elim H1; intros; unfold Rminus in H2; rewrite H2; rewrite Ropp_0; rewrite Rplus_0_r; @@ -178,13 +180,13 @@ Qed. (**********) Lemma ln_exists : forall y:R, 0 < y -> { z:R | y = exp z }. Proof. - intros; case (Rle_dec 1 y); intro. - apply (ln_exists1 _ r). + intros; destruct (Rle_dec 1 y) as [Hle|Hnle]. + apply (ln_exists1 _ Hle). assert (H0 : 1 <= / y). apply Rmult_le_reg_l with y. apply H. rewrite <- Rinv_r_sym. - rewrite Rmult_1_r; left; apply (Rnot_le_lt _ _ n). + rewrite Rmult_1_r; left; apply (Rnot_le_lt _ _ Hnle). red; intro; rewrite H0 in H; elim (Rlt_irrefl _ H). destruct (ln_exists1 _ H0) as (x,p); exists (- x); apply Rmult_eq_reg_l with (exp x / y). @@ -213,12 +215,10 @@ Definition ln (x:R) : R := Lemma exp_ln : forall x:R, 0 < x -> exp (ln x) = x. Proof. - intros; unfold ln; case (Rlt_dec 0 x); intro. + intros; unfold ln; decide (Rlt_dec 0 x) with H. unfold Rln; - case (ln_exists (mkposreal x r) (cond_pos (mkposreal x r))); - intros. - simpl in e; symmetry ; apply e. - elim n; apply H. + case (ln_exists (mkposreal x H) (cond_pos (mkposreal x H))) as (?,Hex). + symmetry; apply Hex. Qed. Theorem exp_inv : forall x y:R, exp x = exp y -> x = y. @@ -313,12 +313,12 @@ Proof. red; apply P_Rmin. apply Rmult_lt_0_compat. assumption. - apply Rplus_lt_reg_r with 1. + apply Rplus_lt_reg_l with 1. rewrite Rplus_0_r; replace (1 + (exp eps - 1)) with (exp eps); [ apply H1 | ring ]. apply Rmult_lt_0_compat. assumption. - apply Rplus_lt_reg_r with (exp (- eps)). + apply Rplus_lt_reg_l with (exp (- eps)). rewrite Rplus_0_r; replace (exp (- eps) + (1 - exp (- eps))) with 1; [ apply H2 | ring ]. unfold dist, R_met, R_dist; simpl. @@ -335,7 +335,7 @@ Proof. apply H. rewrite Hxyy. apply Ropp_lt_cancel. - apply Rplus_lt_reg_r with (r := y). + apply Rplus_lt_reg_l with (r := y). replace (y + - (y * exp (- eps))) with (y * (1 - exp (- eps))); [ idtac | ring ]. replace (y + - x) with (Rabs (x - y)). @@ -358,7 +358,7 @@ Proof. apply Rmult_lt_reg_l with (r := y). apply H. rewrite Hxyy. - apply Rplus_lt_reg_r with (r := - y). + apply Rplus_lt_reg_l with (r := - y). replace (- y + y * exp eps) with (y * (exp eps - 1)); [ idtac | ring ]. replace (- y + x) with (Rabs (x - y)). apply Rlt_le_trans with (1 := H5); apply Rmin_l. @@ -610,7 +610,7 @@ Proof. replace h with (x + h - x); [ idtac | ring ]. apply H3; split. unfold D_x; split. - case (Rcase_abs h); intro. + destruct (Rcase_abs h) as [Hlt|Hgt]. assert (H7 : Rabs h < x / 2). apply Rlt_le_trans with alp. apply H6. @@ -619,13 +619,13 @@ Proof. unfold Rdiv; apply Rmult_lt_0_compat; [ assumption | apply Rinv_0_lt_compat; prove_sup0 ]. rewrite Rabs_left in H7. - apply Rplus_lt_reg_r with (- h - x / 2). + apply Rplus_lt_reg_l with (- h - x / 2). replace (- h - x / 2 + x / 2) with (- h); [ idtac | ring ]. pattern x at 2; rewrite double_var. replace (- h - x / 2 + (x / 2 + x / 2 + h)) with (x / 2); [ apply H7 | ring ]. - apply r. - apply Rplus_lt_le_0_compat; [ assumption | apply Rge_le; apply r ]. - apply (not_eq_sym (A:=R)); apply Rminus_not_eq; replace (x + h - x) with h; + apply Hlt. + apply Rplus_lt_le_0_compat; [ assumption | apply Rge_le; apply Hgt ]. + apply (sym_not_eq (A:=R)); apply Rminus_not_eq; replace (x + h - x) with h; [ apply H5 | ring ]. replace (x + h - x) with h; [ apply Rlt_le_trans with alp; @@ -703,3 +703,128 @@ Proof. ring. apply derivable_pt_lim_exp. Qed. + +(* added later. *) + +Lemma Rpower_mult_distr : + forall x y z, 0 < x -> 0 < y -> + Rpower x z * Rpower y z = Rpower (x * y) z. +intros x y z x0 y0; unfold Rpower. +rewrite <- exp_plus, ln_mult, Rmult_plus_distr_l; auto. +Qed. + +Lemma Rle_Rpower_l a b c: 0 <= c -> 0 < a <= b -> Rpower a c <= Rpower b c. +Proof. +intros [c0 | c0]; + [ | intros; rewrite <- c0, !Rpower_O; [apply Rle_refl | |] ]. + intros [a0 [ab|ab]]. + left; apply exp_increasing. + now apply Rmult_lt_compat_l; auto; apply ln_increasing; fourier. + rewrite ab; apply Rle_refl. + apply Rlt_le_trans with a; tauto. +tauto. +Qed. + +(* arcsinh function *) + +Definition arcsinh x := ln (x + sqrt (x ^ 2 + 1)). + +Lemma arcsinh_sinh : forall x, arcsinh (sinh x) = x. +intros x; unfold sinh, arcsinh. +assert (Rminus_eq_0 : forall r, r - r = 0) by (intros; ring). +pattern 1 at 5; rewrite <- exp_0, <- (Rminus_eq_0 x); unfold Rminus. +rewrite exp_plus. +match goal with |- context[sqrt ?a] => + replace a with (((exp x + exp(-x))/2)^2) by field +end. +rewrite sqrt_pow2; + [|apply Rlt_le, Rmult_lt_0_compat;[apply Rplus_lt_0_compat; apply exp_pos | + apply Rinv_0_lt_compat, Rlt_0_2]]. +match goal with |- context[ln ?a] => replace a with (exp x) by field end. +rewrite ln_exp; reflexivity. +Qed. + +Lemma sinh_arcsinh x : sinh (arcsinh x) = x. +unfold sinh, arcsinh. +assert (cmp : 0 < x + sqrt (x ^ 2 + 1)). + destruct (Rle_dec x 0). + replace (x ^ 2) with ((-x) ^ 2) by ring. + assert (sqrt ((- x) ^ 2) < sqrt ((-x)^2+1)). + apply sqrt_lt_1_alt. + split;[apply pow_le | ]; fourier. + pattern x at 1; replace x with (- (sqrt ((- x) ^ 2))). + assert (t:= sqrt_pos ((-x)^2)); fourier. + simpl; rewrite Rmult_1_r, sqrt_square, Ropp_involutive;[reflexivity | fourier]. + apply Rplus_lt_le_0_compat;[apply Rnot_le_gt; assumption | apply sqrt_pos]. +rewrite exp_ln;[ | assumption]. +rewrite exp_Ropp, exp_ln;[ | assumption]. +assert (Rmult_minus_distr_r : + forall x y z, (x - y) * z = x * z - y * z) by (intros; ring). +apply Rminus_diag_uniq; unfold Rdiv; rewrite Rmult_minus_distr_r. +assert (t: forall x y z, x - z = y -> x - y - z = 0);[ | apply t; clear t]. + intros a b c H; rewrite <- H; ring. +apply Rmult_eq_reg_l with (2 * (x + sqrt (x ^ 2 + 1)));[ | + apply Rgt_not_eq, Rmult_lt_0_compat;[apply Rlt_0_2 | assumption]]. +assert (pow2_sqrt : forall x, 0 <= x -> sqrt x ^ 2 = x) by + (intros; simpl; rewrite Rmult_1_r, sqrt_sqrt; auto). +field_simplify;[rewrite pow2_sqrt;[field | ] | apply Rgt_not_eq; fourier]. +apply Rplus_le_le_0_compat;[simpl; rewrite Rmult_1_r; apply (Rle_0_sqr x)|apply Rlt_le, Rlt_0_1]. +Qed. + +Lemma derivable_pt_lim_arcsinh : + forall x, derivable_pt_lim arcsinh x (/sqrt (x ^ 2 + 1)). +intros x; unfold arcsinh. +assert (0 < x + sqrt (x ^ 2 + 1)). + destruct (Rle_dec x 0); + [ | assert (0 < x) by (apply Rnot_le_gt; assumption); + apply Rplus_lt_le_0_compat; auto; apply sqrt_pos]. + replace (x ^ 2) with ((-x) ^ 2) by ring. + assert (sqrt ((- x) ^ 2) < sqrt ((-x)^2+1)). + apply sqrt_lt_1_alt. + split;[apply pow_le|]; fourier. + pattern x at 1; replace x with (- (sqrt ((- x) ^ 2))). + assert (t:= sqrt_pos ((-x)^2)); fourier. + simpl; rewrite Rmult_1_r, sqrt_square, Ropp_involutive; auto; fourier. +assert (0 < x ^ 2 + 1). + apply Rplus_le_lt_0_compat;[simpl; rewrite Rmult_1_r; apply Rle_0_sqr|fourier]. +replace (/sqrt (x ^ 2 + 1)) with + (/(x + sqrt (x ^ 2 + 1)) * + (1 + (/(2 * sqrt (x ^ 2 + 1)) * (INR 2 * x ^ 1 + 0)))). +apply (derivable_pt_lim_comp (fun x => x + sqrt (x ^ 2 + 1)) ln). + apply (derivable_pt_lim_plus). + apply derivable_pt_lim_id. + apply (derivable_pt_lim_comp (fun x => x ^ 2 + 1) sqrt x). + apply derivable_pt_lim_plus. + apply derivable_pt_lim_pow. + apply derivable_pt_lim_const. + apply derivable_pt_lim_sqrt; assumption. + apply derivable_pt_lim_ln; assumption. + replace (INR 2 * x ^ 1 + 0) with (2 * x) by (simpl; ring). +replace (1 + / (2 * sqrt (x ^ 2 + 1)) * (2 * x)) with + (((sqrt (x ^ 2 + 1) + x))/sqrt (x ^ 2 + 1)); + [ | field; apply Rgt_not_eq, sqrt_lt_R0; assumption]. +apply Rmult_eq_reg_l with (x + sqrt (x ^ 2 + 1)); + [ | apply Rgt_not_eq; assumption]. +rewrite <- Rmult_assoc, Rinv_r;[field | ]; apply Rgt_not_eq; auto; + apply sqrt_lt_R0; assumption. +Qed. + +Lemma arcsinh_lt : forall x y, x < y -> arcsinh x < arcsinh y. +intros x y xy. +case (Rle_dec (arcsinh y) (arcsinh x));[ | apply Rnot_le_lt ]. +intros abs; case (Rlt_not_le _ _ xy). +rewrite <- (sinh_arcsinh y), <- (sinh_arcsinh x). +destruct abs as [lt | q];[| rewrite q; fourier]. +apply Rlt_le, sinh_lt; assumption. +Qed. + +Lemma arcsinh_le : forall x y, x <= y -> arcsinh x <= arcsinh y. +intros x y [xy | xqy]. + apply Rlt_le, arcsinh_lt; assumption. +rewrite xqy; apply Rle_refl. +Qed. + +Lemma arcsinh_0 : arcsinh 0 = 0. + unfold arcsinh; rewrite pow_ne_zero, !Rplus_0_l, sqrt_1, ln_1; + [reflexivity | discriminate]. +Qed. diff --git a/theories/Reals/Rprod.v b/theories/Reals/Rprod.v index 341ec8fd..1ee9410f 100644 --- a/theories/Reals/Rprod.v +++ b/theories/Reals/Rprod.v @@ -1,6 +1,6 @@ (************************************************************************) (* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *) +(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *) (* \VV/ **************************************************************) (* // * This file is distributed under the terms of the *) (* * GNU Lesser General Public License Version 2.1 *) @@ -12,6 +12,7 @@ Require Import Rfunctions. Require Import Rseries. Require Import PartSum. Require Import Binomial. +Require Import Omega. Local Open Scope R_scope. (** TT Ak; 0<=k<=N *) diff --git a/theories/Reals/Rseries.v b/theories/Reals/Rseries.v index c540a931..fd16ea61 100644 --- a/theories/Reals/Rseries.v +++ b/theories/Reals/Rseries.v @@ -1,6 +1,6 @@ (************************************************************************) (* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *) +(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *) (* \VV/ **************************************************************) (* // * This file is distributed under the terms of the *) (* * GNU Lesser General Public License Version 2.1 *) @@ -108,7 +108,7 @@ Section sequence. intros n H4. unfold R_dist. rewrite Rabs_left1, Ropp_minus_distr. - apply Rplus_lt_reg_r with (Un n - eps). + apply Rplus_lt_reg_l with (Un n - eps). apply Rlt_le_trans with (Un N). now replace (Un n - eps + (l - Un n)) with (l - eps) by ring. replace (Un n - eps + eps) with (Un n) by ring. @@ -171,7 +171,7 @@ Section sequence. rewrite H1. apply Rle_trans with (1 := proj2 (Hsum n)). apply Rlt_le. - apply Rplus_lt_reg_r with ((/2)^n - 1). + apply Rplus_lt_reg_l with ((/2)^n - 1). now ring_simplify. exists 0. now exists O. @@ -202,7 +202,7 @@ Section sequence. refine (False_ind _ (Rle_not_lt _ _ (H (l - eps) _) _)). intros x (n, H1). now rewrite H1. - apply Rplus_lt_reg_r with (eps - l). + apply Rplus_lt_reg_l with (eps - l). now ring_simplify. assert (Rabs (/2) < 1). @@ -237,9 +237,9 @@ Section sequence. apply le_n_Sn. rewrite (IHN H6), Rplus_0_l. unfold test. - destruct Rle_lt_dec. + destruct Rle_lt_dec as [Hle|Hlt]. apply eq_refl. - now elim Rlt_not_le with (1 := r). + now elim Rlt_not_le with (1 := Hlt). destruct (le_or_lt N n) as [Hn|Hn]. rewrite le_plus_minus with (1 := Hn). @@ -247,7 +247,7 @@ Section sequence. rewrite Hs, Rplus_0_l. set (k := (N + (n - N))%nat). apply Rlt_le. - apply Rplus_lt_reg_r with ((/2)^k - (/2)^N). + apply Rplus_lt_reg_l with ((/2)^k - (/2)^N). now ring_simplify. apply Rle_trans with (sum N). rewrite le_plus_minus with (1 := Hn). @@ -261,7 +261,7 @@ Section sequence. Lemma Un_cv_crit : Un_growing -> bound EUn -> exists l : R, Un_cv l. Proof. intros Hug Heub. - exists (projT1 (completeness EUn Heub EUn_noempty)). + exists (proj1_sig (completeness EUn Heub EUn_noempty)). destruct (completeness EUn Heub EUn_noempty) as (l, H). now apply Un_cv_crit_lub. Qed. @@ -404,3 +404,26 @@ Proof. apply Rinv_neq_0_compat. assumption. Qed. + +(* Convergence is preserved after shifting the indices. *) +Lemma CV_shift : + forall f k l, Un_cv (fun n => f (n + k)%nat) l -> Un_cv f l. +intros f' k l cvfk eps ep; destruct (cvfk eps ep) as [N Pn]. +exists (N + k)%nat; intros n nN; assert (tmp: (n = (n - k) + k)%nat). + rewrite Nat.sub_add;[ | apply le_trans with (N + k)%nat]; auto with arith. +rewrite tmp; apply Pn; apply Nat.le_add_le_sub_r; assumption. +Qed. + +Lemma CV_shift' : + forall f k l, Un_cv f l -> Un_cv (fun n => f (n + k)%nat) l. +intros f' k l cvf eps ep; destruct (cvf eps ep) as [N Pn]. +exists N; intros n nN; apply Pn; auto with arith. +Qed. + +(* Growing property is preserved after shifting the indices (one way only) *) + +Lemma Un_growing_shift : + forall k un, Un_growing un -> Un_growing (fun n => un (n + k)%nat). +Proof. +intros k un P n; apply P. +Qed. diff --git a/theories/Reals/Rsigma.v b/theories/Reals/Rsigma.v index 0dcb4b25..458d1f8c 100644 --- a/theories/Reals/Rsigma.v +++ b/theories/Reals/Rsigma.v @@ -1,6 +1,6 @@ (************************************************************************) (* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *) +(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *) (* \VV/ **************************************************************) (* // * This file is distributed under the terms of the *) (* * GNU Lesser General Public License Version 2.1 *) @@ -10,6 +10,7 @@ Require Import Rbase. Require Import Rfunctions. Require Import Rseries. Require Import PartSum. +Require Import Omega. Local Open Scope R_scope. Set Implicit Arguments. diff --git a/theories/Reals/Rsqrt_def.v b/theories/Reals/Rsqrt_def.v index 307035ab..b8ec8d3c 100644 --- a/theories/Reals/Rsqrt_def.v +++ b/theories/Reals/Rsqrt_def.v @@ -1,6 +1,6 @@ (************************************************************************) (* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *) +(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *) (* \VV/ **************************************************************) (* // * This file is distributed under the terms of the *) (* * GNU Lesser General Public License Version 2.1 *) @@ -276,8 +276,7 @@ Proof. intros. unfold cv_infty. intro. - case (total_order_T 0 M); intro. - elim s; intro. + destruct (total_order_T 0 M) as [[Hlt|<-]|Hgt]. set (N := up M). cut (0 <= N)%Z. intro. @@ -302,7 +301,6 @@ Proof. assert (H0 := archimed M); elim H0; intros. left; apply Rlt_trans with M; assumption. exists 0%nat; intros. - rewrite <- b. unfold pow_2_n; apply pow_lt; prove_sup0. exists 0%nat; intros. apply Rlt_trans with 0. @@ -342,8 +340,7 @@ Proof. unfold Un_cv; unfold R_dist. intros. assert (H4 := cv_infty_cv_R0 pow_2_n pow_2_n_neq_R0 pow_2_n_infty). - case (total_order_T x y); intro. - elim s; intro. + destruct (total_order_T x y) as [[ Hlt | -> ]|Hgt]. unfold Un_cv in H4; unfold R_dist in H4. cut (0 < y - x). intro Hyp. @@ -373,19 +370,18 @@ Proof. assumption. unfold Rdiv; apply Rmult_lt_0_compat; [ assumption | apply Rinv_0_lt_compat; assumption ]. - apply Rplus_lt_reg_r with x; rewrite Rplus_0_r. + apply Rplus_lt_reg_l with x; rewrite Rplus_0_r. replace (x + (y - x)) with y; [ assumption | ring ]. exists 0%nat; intros. - replace (dicho_lb x y P n - dicho_up x y P n - 0) with - (dicho_lb x y P n - dicho_up x y P n); [ idtac | ring ]. + replace (dicho_lb y y P n - dicho_up y y P n - 0) with + (dicho_lb y y P n - dicho_up y y P n); [ idtac | ring ]. rewrite <- Rabs_Ropp. rewrite Ropp_minus_distr'. rewrite dicho_lb_dicho_up. - rewrite b. unfold Rminus, Rdiv; rewrite Rplus_opp_r; rewrite Rmult_0_l; rewrite Rabs_R0; assumption. assumption. - elim (Rlt_irrefl _ (Rle_lt_trans _ _ _ H r)). + elim (Rlt_irrefl _ (Rle_lt_trans _ _ _ H Hgt)). Qed. Definition cond_positivity (x:R) : bool := @@ -427,18 +423,15 @@ Lemma dicho_lb_car : P x = false -> P (dicho_lb x y P n) = false. Proof. intros. - induction n as [| n Hrecn]. - simpl. - assumption. - simpl. - assert - (X := - sumbool_of_bool (P ((Dichotomy_lb x y P n + Dichotomy_ub x y P n) / 2))). - elim X; intro. - rewrite a. - unfold dicho_lb in Hrecn; assumption. - rewrite b. - assumption. + induction n as [| n Hrecn]. + - assumption. + - simpl. + destruct + (sumbool_of_bool (P ((Dichotomy_lb x y P n + Dichotomy_ub x y P n) / 2))) as [Heq|Heq]. + + rewrite Heq. + unfold dicho_lb in Hrecn; assumption. + + rewrite Heq. + assumption. Qed. Lemma dicho_up_car : @@ -446,18 +439,23 @@ Lemma dicho_up_car : P y = true -> P (dicho_up x y P n) = true. Proof. intros. - induction n as [| n Hrecn]. - simpl. - assumption. - simpl. - assert - (X := - sumbool_of_bool (P ((Dichotomy_lb x y P n + Dichotomy_ub x y P n) / 2))). - elim X; intro. - rewrite a. - unfold dicho_lb in Hrecn; assumption. - rewrite b. - assumption. + induction n as [| n Hrecn]. + - assumption. + - simpl. + destruct + (sumbool_of_bool (P ((Dichotomy_lb x y P n + Dichotomy_ub x y P n) / 2))) as [Heq|Heq]. + + rewrite Heq. + unfold dicho_lb in Hrecn; assumption. + + rewrite Heq. + assumption. +Qed. + +(* A general purpose corollary. *) +Lemma cv_pow_half : forall a, Un_cv (fun n => a/2^n) 0. +intros a; unfold Rdiv; replace 0 with (a * 0) by ring. +apply CV_mult. + intros eps ep; exists 0%nat; rewrite R_dist_eq; intros n _; assumption. +exact (cv_infty_cv_R0 pow_2_n pow_2_n_neq_R0 pow_2_n_infty). Qed. (** Intermediate Value Theorem *) @@ -467,13 +465,9 @@ Lemma IVT : x < y -> f x < 0 -> 0 < f y -> { z:R | x <= z <= y /\ f z = 0 }. Proof. intros. - cut (x <= y). - intro. - generalize (dicho_lb_cv x y (fun z:R => cond_positivity (f z)) H3). - generalize (dicho_up_cv x y (fun z:R => cond_positivity (f z)) H3). - intros X X0. - elim X; intros. - elim X0; intros. + assert (x <= y) by (left; assumption). + destruct (dicho_lb_cv x y (fun z:R => cond_positivity (f z)) H3) as (x1,p0). + destruct (dicho_up_cv x y (fun z:R => cond_positivity (f z)) H3) as (x0,p). assert (H4 := cv_dicho _ _ _ _ _ H3 p0 p). rewrite H4 in p0. exists x0. @@ -490,7 +484,6 @@ Proof. apply dicho_up_decreasing; assumption. assumption. right; reflexivity. - 2: left; assumption. set (Vn := fun n:nat => dicho_lb x y (fun z:R => cond_positivity (f z)) n). set (Wn := fun n:nat => dicho_up x y (fun z:R => cond_positivity (f z)) n). cut ((forall n:nat, f (Vn n) <= 0) -> f x0 <= 0). @@ -515,14 +508,14 @@ Proof. left; assumption. intro. unfold cond_positivity. - case (Rle_dec 0 z); intro. + case (Rle_dec 0 z) as [Hle|Hnle]. split. intro; assumption. intro; reflexivity. split. intro feqt;discriminate feqt. intro. - elim n0; assumption. + contradiction. unfold Vn. cut (forall z:R, cond_positivity z = false <-> z < 0). intros. @@ -536,20 +529,19 @@ Proof. assumption. intro. unfold cond_positivity. - case (Rle_dec 0 z); intro. + case (Rle_dec 0 z) as [Hle|Hnle]. split. intro feqt; discriminate feqt. - intro; elim (Rlt_irrefl _ (Rle_lt_trans _ _ _ r H7)). + intro; elim (Rlt_irrefl _ (Rle_lt_trans _ _ _ Hle H7)). split. intro; auto with real. intro; reflexivity. cut (Un_cv Wn x0). intros. assert (H7 := continuity_seq f Wn x0 (H x0) H5). - case (total_order_T 0 (f x0)); intro. - elim s; intro. + destruct (total_order_T 0 (f x0)) as [[Hlt|<-]|Hgt]. left; assumption. - rewrite <- b; right; reflexivity. + right; reflexivity. unfold Un_cv in H7; unfold R_dist in H7. cut (0 < - f x0). intro. @@ -559,7 +551,7 @@ Proof. rewrite Rabs_right in H11. pattern (- f x0) at 1 in H11; rewrite <- Rplus_0_r in H11. unfold Rminus in H11; rewrite (Rplus_comm (f (Wn x2))) in H11. - assert (H12 := Rplus_lt_reg_r _ _ _ H11). + assert (H12 := Rplus_lt_reg_l _ _ _ H11). assert (H13 := H6 x2). elim (Rlt_irrefl _ (Rle_lt_trans _ _ _ H13 H12)). apply Rle_ge; left; unfold Rminus; apply Rplus_le_lt_0_compat. @@ -570,29 +562,28 @@ Proof. cut (Un_cv Vn x0). intros. assert (H7 := continuity_seq f Vn x0 (H x0) H5). - case (total_order_T 0 (f x0)); intro. - elim s; intro. + destruct (total_order_T 0 (f x0)) as [[Hlt|<-]|Hgt]. unfold Un_cv in H7; unfold R_dist in H7. - elim (H7 (f x0) a); intros. + elim (H7 (f x0) Hlt); intros. cut (x2 >= x2)%nat; [ intro | unfold ge; apply le_n ]. assert (H10 := H8 x2 H9). rewrite Rabs_left in H10. pattern (f x0) at 2 in H10; rewrite <- Rplus_0_r in H10. rewrite Ropp_minus_distr' in H10. unfold Rminus in H10. - assert (H11 := Rplus_lt_reg_r _ _ _ H10). + assert (H11 := Rplus_lt_reg_l _ _ _ H10). assert (H12 := H6 x2). cut (0 < f (Vn x2)). intro. elim (Rlt_irrefl _ (Rlt_le_trans _ _ _ H13 H12)). rewrite <- (Ropp_involutive (f (Vn x2))). apply Ropp_0_gt_lt_contravar; assumption. - apply Rplus_lt_reg_r with (f x0 - f (Vn x2)). + apply Rplus_lt_reg_l with (f x0 - f (Vn x2)). rewrite Rplus_0_r; replace (f x0 - f (Vn x2) + (f (Vn x2) - f x0)) with 0; [ unfold Rminus; apply Rplus_lt_le_0_compat | ring ]. assumption. apply Ropp_0_ge_le_contravar; apply Rle_ge; apply H6. - right; rewrite <- b; reflexivity. + right; reflexivity. left; assumption. unfold Vn; assumption. Qed. @@ -603,31 +594,23 @@ Lemma IVT_cor : x <= y -> f x * f y <= 0 -> { z:R | x <= z <= y /\ f z = 0 }. Proof. intros. - case (total_order_T 0 (f x)); intro. - case (total_order_T 0 (f y)); intro. - elim s; intro. - elim s0; intro. + destruct (total_order_T 0 (f x)) as [[Hltx|Heqx]|Hgtx]. + destruct (total_order_T 0 (f y)) as [[Hlty|Heqy]|Hgty]. cut (0 < f x * f y); [ intro; elim (Rlt_irrefl _ (Rle_lt_trans _ _ _ H1 H2)) | apply Rmult_lt_0_compat; assumption ]. exists y. split. split; [ assumption | right; reflexivity ]. - symmetry ; exact b. - exists x. - split. - split; [ right; reflexivity | assumption ]. - symmetry ; exact b. - elim s; intro. + symmetry ; exact Heqy. cut (x < y). intro. assert (H3 := IVT (- f)%F x y (continuity_opp f H) H2). cut ((- f)%F x < 0). cut (0 < (- f)%F y). intros. - elim (H3 H5 H4); intros. + destruct (H3 H5 H4) as (x0,[]). exists x0. - elim p; intros. split. assumption. unfold opp_fct in H7. @@ -635,25 +618,24 @@ Proof. apply Ropp_eq_0_compat; assumption. unfold opp_fct; apply Ropp_0_gt_lt_contravar; assumption. unfold opp_fct. - apply Rplus_lt_reg_r with (f x); rewrite Rplus_opp_r; rewrite Rplus_0_r; + apply Rplus_lt_reg_l with (f x); rewrite Rplus_opp_r; rewrite Rplus_0_r; assumption. inversion H0. assumption. - rewrite H2 in a. - elim (Rlt_irrefl _ (Rlt_trans _ _ _ r a)). + rewrite H2 in Hltx. + elim (Rlt_irrefl _ (Rlt_trans _ _ _ Hgty Hltx)). exists x. split. split; [ right; reflexivity | assumption ]. symmetry ; assumption. - case (total_order_T 0 (f y)); intro. - elim s; intro. + destruct (total_order_T 0 (f y)) as [[Hlty|Heqy]|Hgty]. cut (x < y). intro. apply IVT; assumption. inversion H0. assumption. - rewrite H2 in r. - elim (Rlt_irrefl _ (Rlt_trans _ _ _ r a)). + rewrite H2 in Hgtx. + elim (Rlt_irrefl _ (Rlt_trans _ _ _ Hlty Hgtx)). exists y. split. split; [ assumption | right; reflexivity ]. @@ -676,8 +658,7 @@ Proof. intro. cut (continuity f). intro. - case (total_order_T y 1); intro. - elim s; intro. + destruct (total_order_T y 1) as [[Hlt| -> ]|Hgt]. cut (0 <= f 1). intro. cut (f 0 * f 1 <= 0). @@ -701,7 +682,7 @@ Proof. exists 1. split. left; apply Rlt_0_1. - rewrite b; symmetry ; apply Rsqr_1. + symmetry; apply Rsqr_1. cut (0 <= f y). intro. cut (f 0 * f y <= 0). @@ -723,7 +704,7 @@ Proof. pattern y at 1; rewrite <- Rmult_1_r. unfold Rsqr; apply Rmult_le_compat_l. assumption. - left; exact r. + left; exact Hgt. replace f with (Rsqr - fct_cte y)%F. apply continuity_minus. apply derivable_continuous; apply derivable_Rsqr. @@ -743,39 +724,31 @@ Definition Rsqrt (y:nonnegreal) : R := Lemma Rsqrt_positivity : forall x:nonnegreal, 0 <= Rsqrt x. Proof. intro. - assert (X := Rsqrt_exists (nonneg x) (cond_nonneg x)). - elim X; intros. + destruct (Rsqrt_exists (nonneg x) (cond_nonneg x)) as (x0 & H1 & H2). cut (x0 = Rsqrt x). intros. - elim p; intros. - rewrite H in H0; assumption. + rewrite <- H; assumption. unfold Rsqrt. - case (Rsqrt_exists x (cond_nonneg x)). - intros. - elim p; elim a; intros. + case (Rsqrt_exists x (cond_nonneg x)) as (?,[]). apply Rsqr_inj. assumption. assumption. - rewrite <- H0; rewrite <- H2; reflexivity. + rewrite <- H0, <- H2; reflexivity. Qed. (**********) Lemma Rsqrt_Rsqrt : forall x:nonnegreal, Rsqrt x * Rsqrt x = x. Proof. intros. - assert (X := Rsqrt_exists (nonneg x) (cond_nonneg x)). - elim X; intros. + destruct (Rsqrt_exists (nonneg x) (cond_nonneg x)) as (x0 & H1 & H2). cut (x0 = Rsqrt x). intros. rewrite <- H. - elim p; intros. - rewrite H1; reflexivity. + rewrite H2; reflexivity. unfold Rsqrt. - case (Rsqrt_exists x (cond_nonneg x)). - intros. - elim p; elim a; intros. + case (Rsqrt_exists x (cond_nonneg x)) as (x1 & ? & ?). apply Rsqr_inj. assumption. assumption. - rewrite <- H0; rewrite <- H2; reflexivity. + rewrite <- H0, <- H2; reflexivity. Qed. diff --git a/theories/Reals/Rtopology.v b/theories/Reals/Rtopology.v index 9a345153..72e4142b 100644 --- a/theories/Reals/Rtopology.v +++ b/theories/Reals/Rtopology.v @@ -1,6 +1,6 @@ (************************************************************************) (* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *) +(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *) (* \VV/ **************************************************************) (* // * This file is distributed under the terms of the *) (* * GNU Lesser General Public License Version 2.1 *) @@ -84,7 +84,7 @@ Proof. apply H4. unfold del; rewrite <- (Rabs_Ropp (x - x1)); rewrite Ropp_minus_distr; ring. - unfold del; apply Rplus_lt_reg_r with (Rabs (x - x1)); + unfold del; apply Rplus_lt_reg_l with (Rabs (x - x1)); rewrite Rplus_0_r; replace (Rabs (x - x1) + (x0 - Rabs (x - x1))) with (pos x0); [ idtac | ring ]. @@ -139,7 +139,7 @@ Proof. apply H10. unfold del; simpl; rewrite <- (Rabs_Ropp (x - x1)); rewrite Ropp_minus_distr; ring. - apply Rplus_lt_reg_r with (Rabs (x - x1)); rewrite Rplus_0_r; + apply Rplus_lt_reg_l with (Rabs (x - x1)); rewrite Rplus_0_r; replace (Rabs (x - x1) + (x0 - Rabs (x - x1))) with (pos x0); [ rewrite <- Rabs_Ropp; rewrite Ropp_minus_distr; apply H6 | ring ]. Qed. @@ -254,7 +254,7 @@ Proof. apply H4. unfold del2; simpl; rewrite <- (Rabs_Ropp (x - x0)); rewrite Ropp_minus_distr; ring. - apply Rplus_lt_reg_r with (Rabs (x - x0)); rewrite Rplus_0_r; + apply Rplus_lt_reg_l with (Rabs (x - x0)); rewrite Rplus_0_r; replace (Rabs (x - x0) + (del - Rabs (x - x0))) with (pos del); [ rewrite <- Rabs_Ropp; rewrite Ropp_minus_distr; apply H2 | ring ]. apply interior_P1. @@ -623,87 +623,79 @@ Qed. (** Borel-Lebesgue's lemma *) Lemma compact_P3 : forall a b:R, compact (fun c:R => a <= c <= b). Proof. - intros; case (Rle_dec a b); intro. - unfold compact; intros; + intros a b; destruct (Rle_dec a b) as [Hle|Hnle]. + unfold compact; intros f0 (H,H5); set (A := fun x:R => a <= x <= b /\ (exists D : R -> Prop, - covering_finite (fun c:R => a <= c <= x) (subfamily f0 D))); - cut (A a). - intro; cut (bound A). - intro; cut (exists a0 : R, A a0). - intro; assert (H3 := completeness A H1 H2); elim H3; clear H3; intros m H3; - unfold is_lub in H3; cut (a <= m <= b). - intro; unfold covering_open_set in H; elim H; clear H; intros; - unfold covering in H; assert (H6 := H m H4); elim H6; - clear H6; intros y0 H6; unfold family_open_set in H5; - assert (H7 := H5 y0); unfold open_set in H7; assert (H8 := H7 m H6); - unfold neighbourhood in H8; elim H8; clear H8; intros eps H8; - cut (exists x : R, A x /\ m - eps < x <= m). - intro; elim H9; clear H9; intros x H9; elim H9; clear H9; intros; - case (Req_dec m b); intro. - rewrite H11 in H10; rewrite H11 in H8; unfold A in H9; elim H9; clear H9; - intros; elim H12; clear H12; intros Dx H12; - set (Db := fun x:R => Dx x \/ x = y0); exists Db; - unfold covering_finite; split. - unfold covering; unfold covering_finite in H12; elim H12; clear H12; - intros; unfold covering in H12; case (Rle_dec x0 x); - intro. - cut (a <= x0 <= x). - intro; assert (H16 := H12 x0 H15); elim H16; clear H16; intros; exists x1; - simpl in H16; simpl; unfold Db; elim H16; - clear H16; intros; split; [ apply H16 | left; apply H17 ]. - split. - elim H14; intros; assumption. - assumption. + covering_finite (fun c:R => a <= c <= x) (subfamily f0 D))). + cut (A a); [intro H0|]. + cut (bound A); [intro H1|]. + cut (exists a0 : R, A a0); [intro H2|]. + pose proof (completeness A H1 H2) as (m,H3); unfold is_lub in H3. + cut (a <= m <= b); [intro H4|]. + unfold covering in H; pose proof (H m H4) as (y0,H6). + unfold family_open_set in H5; pose proof (H5 y0 m H6) as (eps,H8). + cut (exists x : R, A x /\ m - eps < x <= m); + [intros (x,((H9 & Dx & H12 & H13),(Hltx,_)))|]. + destruct (Req_dec m b) as [->|H11]. + set (Db := fun x:R => Dx x \/ x = y0); exists Db; + unfold covering_finite; split. + unfold covering; intros x0 (H14,H18); + unfold covering in H12; destruct (Rle_dec x0 x) as [Hle'|Hnle']. + cut (a <= x0 <= x); [intro H15|]. + pose proof (H12 x0 H15) as (x1 & H16 & H17); exists x1; + simpl; unfold Db; split; [ apply H16 | left; apply H17 ]. + split; assumption. exists y0; simpl; split. - apply H8; unfold disc; rewrite <- Rabs_Ropp; rewrite Ropp_minus_distr; - rewrite Rabs_right. + apply H8; unfold disc; + rewrite <- Rabs_Ropp, Ropp_minus_distr, Rabs_right. apply Rlt_trans with (b - x). - unfold Rminus; apply Rplus_lt_compat_l; apply Ropp_lt_gt_contravar; + unfold Rminus; apply Rplus_lt_compat_l, Ropp_lt_gt_contravar; auto with real. - elim H10; intros H15 _; apply Rplus_lt_reg_r with (x - eps); + apply Rplus_lt_reg_l with (x - eps); replace (x - eps + (b - x)) with (b - eps); - [ replace (x - eps + eps) with x; [ apply H15 | ring ] | ring ]. - apply Rge_minus; apply Rle_ge; elim H14; intros _ H15; apply H15. + [ replace (x - eps + eps) with x; [ apply Hltx | ring ] | ring ]. + apply Rge_minus, Rle_ge, H18. unfold Db; right; reflexivity. - unfold family_finite; unfold domain_finite; - unfold covering_finite in H12; elim H12; clear H12; + unfold family_finite, domain_finite. intros; unfold family_finite in H13; unfold domain_finite in H13; - elim H13; clear H13; intros l H13; exists (cons y0 l); + destruct H13 as (l,H13); exists (cons y0 l); intro; split. - intro; simpl in H14; unfold intersection_domain in H14; elim (H13 x0); - clear H13; intros; case (Req_dec x0 y0); intro. + intro H14; simpl in H14; unfold intersection_domain in H14; + specialize H13 with x0; destruct H13 as (H13,H15); + destruct (Req_dec x0 y0) as [H16|H16]. simpl; left; apply H16. simpl; right; apply H13. simpl; unfold intersection_domain; unfold Db in H14; decompose [and or] H14. split; assumption. elim H16; assumption. - intro; simpl in H14; elim H14; intro; simpl; + intro H14; simpl in H14; destruct H14 as [H15|H15]; simpl; unfold intersection_domain. split. - apply (cond_fam f0); rewrite H15; exists m; apply H6. + apply (cond_fam f0); rewrite H15; exists b; apply H6. unfold Db; right; assumption. simpl; unfold intersection_domain; elim (H13 x0). intros _ H16; assert (H17 := H16 H15); simpl in H17; unfold intersection_domain in H17; split. elim H17; intros; assumption. unfold Db; left; elim H17; intros; assumption. - set (m' := Rmin (m + eps / 2) b); cut (A m'). - intro; elim H3; intros; unfold is_upper_bound in H13; - assert (H15 := H13 m' H12); cut (m < m'). - intro; elim (Rlt_irrefl _ (Rle_lt_trans _ _ _ H15 H16)). - unfold m'; unfold Rmin; case (Rle_dec (m + eps / 2) b); intro. - pattern m at 1; rewrite <- Rplus_0_r; apply Rplus_lt_compat_l; - unfold Rdiv; apply Rmult_lt_0_compat; - [ apply (cond_pos eps) | apply Rinv_0_lt_compat; prove_sup0 ]. - elim H4; intros. - elim H17; intro. - assumption. - elim H11; assumption. + set (m' := Rmin (m + eps / 2) b). + cut (A m'); [intro H7|]. + destruct H3 as (H14,H15); unfold is_upper_bound in H14. + assert (H16 := H14 m' H7). + cut (m < m'); [intro H17|]. + elim (Rlt_irrefl _ (Rle_lt_trans _ _ _ H16 H17))... + unfold m', Rmin; destruct (Rle_dec (m + eps / 2) b) as [Hle'|Hnle']. + pattern m at 1; rewrite <- Rplus_0_r; apply Rplus_lt_compat_l; + unfold Rdiv; apply Rmult_lt_0_compat; + [ apply (cond_pos eps) | apply Rinv_0_lt_compat; prove_sup0 ]. + destruct H4 as (_,[]). + assumption. + elim H11; assumption. unfold A; split. split. apply Rle_trans with m. @@ -712,38 +704,32 @@ Proof. pattern m at 1; rewrite <- Rplus_0_r; apply Rplus_le_compat_l; left; unfold Rdiv; apply Rmult_lt_0_compat; [ apply (cond_pos eps) | apply Rinv_0_lt_compat; prove_sup0 ]. - elim H4; intros. - elim H13; intro. + destruct H4. assumption. - elim H11; assumption. unfold m'; apply Rmin_r. - unfold A in H9; elim H9; clear H9; intros; elim H12; clear H12; intros Dx H12; - set (Db := fun x:R => Dx x \/ x = y0); exists Db; - unfold covering_finite; split. - unfold covering; unfold covering_finite in H12; elim H12; clear H12; - intros; unfold covering in H12; case (Rle_dec x0 x); - intro. - cut (a <= x0 <= x). - intro; assert (H16 := H12 x0 H15); elim H16; clear H16; intros; exists x1; - simpl in H16; simpl; unfold Db. - elim H16; clear H16; intros; split; [ apply H16 | left; apply H17 ]. - elim H14; intros; split; assumption. + set (Db := fun x:R => Dx x \/ x = y0); exists Db; + unfold covering_finite; split. + unfold covering; intros x0 (H14,H18); + unfold covering in H12; destruct (Rle_dec x0 x) as [Hle'|Hnle']. + cut (a <= x0 <= x); [intro H15|]. + pose proof (H12 x0 H15) as (x1 & H16 & H17); exists x1; + simpl; unfold Db; split; [ apply H16 | left; apply H17 ]. + split; assumption. exists y0; simpl; split. - apply H8; unfold disc; unfold Rabs; case (Rcase_abs (x0 - m)); - intro. + apply H8; unfold disc, Rabs; destruct (Rcase_abs (x0 - m)) as [Hlt|Hge]. rewrite Ropp_minus_distr; apply Rlt_trans with (m - x). unfold Rminus; apply Rplus_lt_compat_l; apply Ropp_lt_gt_contravar; auto with real. - apply Rplus_lt_reg_r with (x - eps); + apply Rplus_lt_reg_l with (x - eps); replace (x - eps + (m - x)) with (m - eps). replace (x - eps + eps) with x. - elim H10; intros; assumption. + assumption. ring. ring. apply Rle_lt_trans with (m' - m). unfold Rminus; do 2 rewrite <- (Rplus_comm (- m)); apply Rplus_le_compat_l; elim H14; intros; assumption. - apply Rplus_lt_reg_r with m; replace (m + (m' - m)) with m'. + apply Rplus_lt_reg_l with m; replace (m + (m' - m)) with m'. apply Rle_lt_trans with (m + eps / 2). unfold m'; apply Rmin_l. apply Rplus_lt_compat_l; apply Rmult_lt_reg_l with 2. @@ -755,22 +741,20 @@ Proof. discrR. ring. unfold Db; right; reflexivity. - unfold family_finite; unfold domain_finite; - unfold covering_finite in H12; elim H12; clear H12; - intros; unfold family_finite in H13; unfold domain_finite in H13; - elim H13; clear H13; intros l H13; exists (cons y0 l); + unfold family_finite, domain_finite; + unfold family_finite, domain_finite in H13; + destruct H13 as (l,H13); exists (cons y0 l); intro; split. - intro; simpl in H14; unfold intersection_domain in H14; elim (H13 x0); - clear H13; intros; case (Req_dec x0 y0); intro. - simpl; left; apply H16. + intro H14; simpl in H14; unfold intersection_domain in H14; + specialize (H13 x0); destruct H13 as (H13,H15); + destruct (Req_dec x0 y0) as [Heq|Hneq]. + simpl; left; apply Heq. simpl; right; apply H13; simpl; unfold intersection_domain; unfold Db in H14; decompose [and or] H14. split; assumption. - elim H16; assumption. - intro; simpl in H14; elim H14; intro; simpl; - unfold intersection_domain. - split. + elim Hneq; assumption. + intros [H15|H15]. split. apply (cond_fam f0); rewrite H15; exists m; apply H6. unfold Db; right; assumption. elim (H13 x0); intros _ H16. @@ -780,22 +764,22 @@ Proof. split. elim H17; intros; assumption. unfold Db; left; elim H17; intros; assumption. - elim (classic (exists x : R, A x /\ m - eps < x <= m)); intro. + elim (classic (exists x : R, A x /\ m - eps < x <= m)); intro H9. assumption. - elim H3; intros; cut (is_upper_bound A (m - eps)). - intro; assert (H13 := H11 _ H12); cut (m - eps < m). - intro; elim (Rlt_irrefl _ (Rle_lt_trans _ _ _ H13 H14)). + elim H3; intros H10 H11; cut (is_upper_bound A (m - eps)). + intro H12; assert (H13 := H11 _ H12); cut (m - eps < m). + intro H14; elim (Rlt_irrefl _ (Rle_lt_trans _ _ _ H13 H14)). pattern m at 2; rewrite <- Rplus_0_r; unfold Rminus; apply Rplus_lt_compat_l; apply Ropp_lt_cancel; rewrite Ropp_involutive; rewrite Ropp_0; apply (cond_pos eps). set (P := fun n:R => A n /\ m - eps < n <= m); assert (H12 := not_ex_all_not _ P H9); unfold P in H12; - unfold is_upper_bound; intros; + unfold is_upper_bound; intros x H13; assert (H14 := not_and_or _ _ (H12 x)); elim H14; - intro. + intro H15. elim H15; apply H13. - elim (not_and_or _ _ H15); intro. - case (Rle_dec x (m - eps)); intro. + destruct (not_and_or _ _ H15) as [H16|H16]. + destruct (Rle_dec x (m - eps)) as [H17|H17]. assumption. elim H16; auto with real. unfold is_upper_bound in H10; assert (H17 := H10 x H13); elim H16; apply H17. @@ -803,7 +787,8 @@ Proof. unfold is_upper_bound in H3. split. apply (H3 _ H0). - apply (H4 b); unfold is_upper_bound; intros; unfold A in H5; elim H5; + clear H5. + apply (H4 b); unfold is_upper_bound; intros x H5; unfold A in H5; elim H5; clear H5; intros H5 _; elim H5; clear H5; intros _ H5; apply H5. exists a; apply H0. @@ -811,30 +796,28 @@ Proof. unfold A in H1; elim H1; clear H1; intros H1 _; elim H1; clear H1; intros _ H1; apply H1. unfold A; split. - split; [ right; reflexivity | apply r ]. - unfold covering_open_set in H; elim H; clear H; intros; unfold covering in H; - cut (a <= a <= b). - intro; elim (H _ H1); intros y0 H2; set (D' := fun x:R => x = y0); exists D'; + split; [ right; reflexivity | apply Hle ]. + unfold covering in H; cut (a <= a <= b). + intro H1; elim (H _ H1); intros y0 H2; set (D' := fun x:R => x = y0); exists D'; unfold covering_finite; split. - unfold covering; simpl; intros; cut (x = a). - intro; exists y0; split. + unfold covering; simpl; intros x H3; cut (x = a). + intro H4; exists y0; split. rewrite H4; apply H2. unfold D'; reflexivity. elim H3; intros; apply Rle_antisym; assumption. unfold family_finite; unfold domain_finite; exists (cons y0 nil); intro; split. - simpl; unfold intersection_domain; intro; elim H3; clear H3; - intros; unfold D' in H4; left; apply H4. - simpl; unfold intersection_domain; intro; elim H3; intro. + simpl; unfold intersection_domain; intros (H3,H4). + unfold D' in H4; left; apply H4. + simpl; unfold intersection_domain; intros [H4|[]]. split; [ rewrite H4; apply (cond_fam f0); exists a; apply H2 | apply H4 ]. - elim H4. - split; [ right; reflexivity | apply r ]. + split; [ right; reflexivity | apply Hle ]. apply compact_eqDom with (fun c:R => False). apply compact_EMP. unfold eq_Dom; split. unfold included; intros; elim H. unfold included; intros; elim H; clear H; intros; - assert (H1 := Rle_trans _ _ _ H H0); elim n; apply H1. + assert (H1 := Rle_trans _ _ _ H H0); elim Hnle; apply H1. Qed. Lemma compact_P4 : @@ -982,12 +965,6 @@ Proof. intros; exists (f0 x0); apply H4. Qed. -Lemma Rlt_Rminus : forall a b:R, a < b -> 0 < b - a. -Proof. - intros; apply Rplus_lt_reg_r with a; rewrite Rplus_0_r; - replace (a + (b - a)) with b; [ assumption | ring ]. -Qed. - Lemma prolongement_C0 : forall (f:R -> R) (a b:R), a <= b -> @@ -1017,14 +994,14 @@ Proof. split. change (0 < a - x); apply Rlt_Rminus; assumption. intros; elim H5; clear H5; intros _ H5; unfold h. - case (Rle_dec x a); intro. - case (Rle_dec x0 a); intro. - unfold Rminus; rewrite Rplus_opp_r; rewrite Rabs_R0; assumption. - elim n; left; apply Rplus_lt_reg_r with (- x); + case (Rle_dec x a) as [|[]]. + case (Rle_dec x0 a) as [|[]]. + unfold Rminus; rewrite Rplus_opp_r, Rabs_R0; assumption. + left; apply Rplus_lt_reg_l with (- x); do 2 rewrite (Rplus_comm (- x)); apply Rle_lt_trans with (Rabs (x0 - x)). apply RRle_abs. assumption. - elim n; left; assumption. + left; assumption. elim H3; intro. assert (H5 : a <= a <= b). split; [ right; reflexivity | left; assumption ]. @@ -1039,20 +1016,20 @@ Proof. elim H8; intros; assumption. change (0 < b - a); apply Rlt_Rminus; assumption. intros; elim H9; clear H9; intros _ H9; cut (x1 < b). - intro; unfold h; case (Rle_dec x a); intro. - case (Rle_dec x1 a); intro. + intro; unfold h; case (Rle_dec x a) as [|[]]. + case (Rle_dec x1 a) as [Hlta|Hnlea]. unfold Rminus; rewrite Rplus_opp_r; rewrite Rabs_R0; assumption. - case (Rle_dec x1 b); intro. + case (Rle_dec x1 b) as [Hleb|[]]. elim H8; intros; apply H12; split. unfold D_x, no_cond; split. trivial. - red; intro; elim n; right; symmetry ; assumption. + red; intro; elim Hnlea; right; symmetry ; assumption. apply Rlt_le_trans with (Rmin x0 (b - a)). rewrite H4 in H9; apply H9. apply Rmin_l. - elim n0; left; assumption. - elim n; right; assumption. - apply Rplus_lt_reg_r with (- a); do 2 rewrite (Rplus_comm (- a)); + left; assumption. + right; assumption. + apply Rplus_lt_reg_l with (- a); do 2 rewrite (Rplus_comm (- a)); rewrite H4 in H9; apply Rle_lt_trans with (Rabs (x1 - a)). apply RRle_abs. apply Rlt_le_trans with (Rmin x0 (b - a)). @@ -1073,30 +1050,29 @@ Proof. assert (H12 : 0 < b - x). apply Rlt_Rminus; assumption. exists (Rmin x0 (Rmin (x - a) (b - x))); split. - unfold Rmin; case (Rle_dec (x - a) (b - x)); intro. - case (Rle_dec x0 (x - a)); intro. + unfold Rmin; case (Rle_dec (x - a) (b - x)) as [Hle|Hnle]. + case (Rle_dec x0 (x - a)) as [Hlea|Hnlea]. assumption. assumption. - case (Rle_dec x0 (b - x)); intro. + case (Rle_dec x0 (b - x)) as [Hleb|Hnleb]. assumption. assumption. - intros; elim H13; clear H13; intros; cut (a < x1 < b). - intro; elim H15; clear H15; intros; unfold h; case (Rle_dec x a); - intro. - elim (Rlt_irrefl _ (Rle_lt_trans _ _ _ r H4)). - case (Rle_dec x b); intro. - case (Rle_dec x1 a); intro. - elim (Rlt_irrefl _ (Rle_lt_trans _ _ _ r0 H15)). - case (Rle_dec x1 b); intro. + intros x1 (H13,H14); cut (a < x1 < b). + intro; elim H15; clear H15; intros; unfold h; case (Rle_dec x a) as [Hle|Hnle]. + elim (Rlt_irrefl _ (Rle_lt_trans _ _ _ Hle H4)). + case (Rle_dec x b) as [|[]]. + case (Rle_dec x1 a) as [Hle0|]. + elim (Rlt_irrefl _ (Rle_lt_trans _ _ _ Hle0 H15)). + case (Rle_dec x1 b) as [|[]]. apply H10; split. assumption. apply Rlt_le_trans with (Rmin x0 (Rmin (x - a) (b - x))). assumption. apply Rmin_l. - elim n1; left; assumption. - elim n0; left; assumption. + left; assumption. + left; assumption. split. - apply Ropp_lt_cancel; apply Rplus_lt_reg_r with x; + apply Ropp_lt_cancel; apply Rplus_lt_reg_l with x; apply Rle_lt_trans with (Rabs (x1 - x)). rewrite <- Rabs_Ropp; rewrite Ropp_minus_distr; apply RRle_abs. apply Rlt_le_trans with (Rmin x0 (Rmin (x - a) (b - x))). @@ -1104,7 +1080,7 @@ Proof. apply Rle_trans with (Rmin (x - a) (b - x)). apply Rmin_r. apply Rmin_l. - apply Rplus_lt_reg_r with (- x); do 2 rewrite (Rplus_comm (- x)); + apply Rplus_lt_reg_l with (- x); do 2 rewrite (Rplus_comm (- x)); apply Rle_lt_trans with (Rabs (x1 - x)). apply RRle_abs. apply Rlt_le_trans with (Rmin x0 (Rmin (x - a) (b - x))). @@ -1124,13 +1100,13 @@ Proof. elim H10; intros; assumption. change (0 < b - a); apply Rlt_Rminus; assumption. intros; elim H11; clear H11; intros _ H11; cut (a < x1). - intro; unfold h; case (Rle_dec x a); intro. - elim (Rlt_irrefl _ (Rle_lt_trans _ _ _ r H4)). - case (Rle_dec x1 a); intro. - elim (Rlt_irrefl _ (Rle_lt_trans _ _ _ r H12)). - case (Rle_dec x b); intro. - case (Rle_dec x1 b); intro. - rewrite H6; elim H10; intros; elim r0; intro. + intro; unfold h; case (Rle_dec x a) as [Hlea|Hnlea]. + elim (Rlt_irrefl _ (Rle_lt_trans _ _ _ Hlea H4)). + case (Rle_dec x1 a) as [Hlea'|Hnlea']. + elim (Rlt_irrefl _ (Rle_lt_trans _ _ _ Hlea' H12)). + case (Rle_dec x b) as [Hleb|Hnleb]. + case (Rle_dec x1 b) as [Hleb'|Hnleb']. + rewrite H6; elim H10; intros; destruct Hleb'. apply H14; split. unfold D_x, no_cond; split. trivial. @@ -1142,8 +1118,8 @@ Proof. assumption. rewrite H6; unfold Rminus; rewrite Rplus_opp_r; rewrite Rabs_R0; assumption. - elim n1; right; assumption. - rewrite H6 in H11; apply Ropp_lt_cancel; apply Rplus_lt_reg_r with b; + elim Hnleb; right; assumption. + rewrite H6 in H11; apply Ropp_lt_cancel; apply Rplus_lt_reg_l with b; apply Rle_lt_trans with (Rabs (x1 - b)). rewrite <- Rabs_Ropp; rewrite Ropp_minus_distr; apply RRle_abs. apply Rlt_le_trans with (Rmin x0 (b - a)). @@ -1156,26 +1132,25 @@ Proof. change (0 < x - b); apply Rlt_Rminus; assumption. intros; elim H8; clear H8; intros. assert (H10 : b < x0). - apply Ropp_lt_cancel; apply Rplus_lt_reg_r with x; + apply Ropp_lt_cancel; apply Rplus_lt_reg_l with x; apply Rle_lt_trans with (Rabs (x0 - x)). rewrite <- Rabs_Ropp; rewrite Ropp_minus_distr; apply RRle_abs. assumption. - unfold h; case (Rle_dec x a); intro. - elim (Rlt_irrefl _ (Rle_lt_trans _ _ _ r H4)). - case (Rle_dec x b); intro. - elim (Rlt_irrefl _ (Rle_lt_trans _ _ _ r H6)). - case (Rle_dec x0 a); intro. - elim (Rlt_irrefl _ (Rlt_trans _ _ _ H1 (Rlt_le_trans _ _ _ H10 r))). - case (Rle_dec x0 b); intro. - elim (Rlt_irrefl _ (Rle_lt_trans _ _ _ r H10)). + unfold h; case (Rle_dec x a) as [Hle|Hnle]. + elim (Rlt_irrefl _ (Rle_lt_trans _ _ _ Hle H4)). + case (Rle_dec x b) as [Hleb|Hnleb]. + elim (Rlt_irrefl _ (Rle_lt_trans _ _ _ Hleb H6)). + case (Rle_dec x0 a) as [Hlea'|Hnlea']. + elim (Rlt_irrefl _ (Rlt_trans _ _ _ H1 (Rlt_le_trans _ _ _ H10 Hlea'))). + case (Rle_dec x0 b) as [Hleb'|Hnleb']. + elim (Rlt_irrefl _ (Rle_lt_trans _ _ _ Hleb' H10)). unfold Rminus; rewrite Rplus_opp_r; rewrite Rabs_R0; assumption. - intros; elim H3; intros; unfold h; case (Rle_dec c a); intro. - elim r; intro. + intros; elim H3; intros; unfold h; case (Rle_dec c a) as [[|]|]. elim (Rlt_irrefl _ (Rle_lt_trans _ _ _ H4 H6)). rewrite H6; reflexivity. - case (Rle_dec c b); intro. + case (Rle_dec c b) as [|[]]. reflexivity. - elim n0; assumption. + assumption. exists (fun _:R => f0 a); split. apply derivable_continuous; apply (derivable_const (f0 a)). intros; elim H2; intros; rewrite H1 in H3; cut (b = c). @@ -1229,11 +1204,11 @@ Proof. apply Rplus_lt_compat_l; apply Ropp_lt_cancel; rewrite Ropp_0; rewrite Ropp_involutive; apply (cond_pos eps). unfold is_upper_bound, image_dir; intros; cut (x <= M). - intro; case (Rle_dec x (M - eps)); intro. - apply r. + intro; destruct (Rle_dec x (M - eps)) as [H13|]. + apply H13. elim (H9 x); unfold intersection_domain, disc, image_dir; split. rewrite <- Rabs_Ropp; rewrite Ropp_minus_distr; rewrite Rabs_right. - apply Rplus_lt_reg_r with (x - eps); + apply Rplus_lt_reg_l with (x - eps); replace (x - eps + (M - x)) with (M - eps). replace (x - eps + eps) with x. auto with real. @@ -1615,13 +1590,12 @@ Proof. apply H3. elim Hyp; intros; elim H4; intros; decompose [and] H5; assert (H10 := H3 _ H6); assert (H11 := H3 _ H8); - elim H10; intros; elim H11; intros; case (total_order_T x x0); - intro. - elim s; intro. + elim H10; intros; elim H11; intros; + destruct (total_order_T x x0) as [[|H15]|H15]. assumption. - rewrite b in H13; rewrite b in H7; elim H9; apply Rle_antisym; + rewrite H15 in H13, H7; elim H9; apply Rle_antisym; apply Rle_trans with x0; assumption. - elim (Rlt_irrefl _ (Rle_lt_trans _ _ _ (Rle_trans _ _ _ H13 H14) r)). + elim (Rlt_irrefl _ (Rle_lt_trans _ _ _ (Rle_trans _ _ _ H13 H14) H15)). elim X_enc; clear X_enc; intros m X_enc; elim X_enc; clear X_enc; intros M X_enc; elim X_enc; clear X_enc Hyp; intros X_enc Hyp; unfold uniform_continuity; intro; @@ -1675,9 +1649,9 @@ Proof. apply H7; split. unfold D_x, no_cond; split; [ trivial | assumption ]. apply Rlt_le_trans with (Rmin x0 (M - m)); [ apply H8 | apply Rmin_l ]. - assert (H8 := completeness _ H6 H7); elim H8; clear H8; intros; + destruct (completeness _ H6 H7) as (x1,p). cut (0 < x1 <= M - m). - intro; elim H8; clear H8; intros; exists (mkposreal _ H8); split. + intros (H8,H9); exists (mkposreal _ H8); split. intros; cut (exists alp : R, Rabs (z - x) < alp <= x1 /\ E alp). intros; elim H11; intros; elim H12; clear H12; intros; unfold E in H13; elim H13; intros; apply H15. @@ -1831,7 +1805,7 @@ Proof. apply H14; split; [ unfold D_x, no_cond; split; [ trivial | assumption ] | apply Rlt_le_trans with (Rmin x0 (M - m)); [ apply H15 | apply Rmin_l ] ]. - assert (H13 := completeness _ H11 H12); elim H13; clear H13; intros; + destruct (completeness _ H11 H12) as (x0,p). cut (0 < x0 <= M - m). intro; elim H13; clear H13; intros; exists x0; split. assumption. diff --git a/theories/Reals/Rtrigo.v b/theories/Reals/Rtrigo.v index 6818e9a1..44058358 100644 --- a/theories/Reals/Rtrigo.v +++ b/theories/Reals/Rtrigo.v @@ -1,6 +1,6 @@ (************************************************************************) (* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *) +(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *) (* \VV/ **************************************************************) (* // * This file is distributed under the terms of the *) (* * GNU Lesser General Public License Version 2.1 *) @@ -16,11 +16,10 @@ Require Export Cos_rel. Require Export Cos_plus. Require Import ZArith_base. Require Import Zcomplements. -Require Import Classical_Prop. Require Import Fourier. Require Import Ranalysis1. Require Import Rsqrt_def. Require Import PSeries_reg. Require Export Rtrigo1. Require Export Ratan. -Require Export Machin.
\ No newline at end of file +Require Export Machin. diff --git a/theories/Reals/Rtrigo1.v b/theories/Reals/Rtrigo1.v index b940455f..9e485ec5 100644 --- a/theories/Reals/Rtrigo1.v +++ b/theories/Reals/Rtrigo1.v @@ -1,6 +1,6 @@ (************************************************************************) (* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *) +(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *) (* \VV/ **************************************************************) (* // * This file is distributed under the terms of the *) (* * GNU Lesser General Public License Version 2.1 *) @@ -16,7 +16,6 @@ Require Export Cos_rel. Require Export Cos_plus. Require Import ZArith_base. Require Import Zcomplements. -Require Import Classical_Prop. Require Import Fourier. Require Import Ranalysis1. Require Import Rsqrt_def. @@ -40,7 +39,7 @@ Proof. (fun n:nat => sum_f_R0 (fun k:nat => Rabs (/ INR (fact (2 * k)) * r ^ (2 * k))) n) l }. - intro X; elim X; intros. + intros (x,p). exists x. split. apply p. @@ -148,11 +147,11 @@ Proof. apply H4. intros; rewrite (H0 x); rewrite (H0 x1); apply H5; apply H6. intro; unfold cos, SFL in |- *. - case (cv x); case (exist_cos (Rsqr x)); intros. - symmetry in |- *; eapply UL_sequence. - apply u. - unfold cos_in in c; unfold infinite_sum in c; unfold Un_cv in |- *; intros. - elim (c _ H0); intros N0 H1. + case (cv x) as (x1,HUn); case (exist_cos (Rsqr x)) as (x0,Hcos); intros. + symmetry; eapply UL_sequence. + apply HUn. + unfold cos_in, infinite_sum in Hcos; unfold Un_cv in |- *; intros. + elim (Hcos _ H0); intros N0 H1. exists N0; intros. unfold R_dist in H1; unfold R_dist, SP in |- *. replace (sum_f_R0 (fun k:nat => fn k x) n) with @@ -586,8 +585,8 @@ Qed. Lemma SIN_bound : forall x:R, -1 <= sin x <= 1. Proof. - intro; case (Rle_dec (-1) (sin x)); intro. - case (Rle_dec (sin x) 1); intro. + intro; destruct (Rle_dec (-1) (sin x)) as [Hle|Hnle]. + destruct (Rle_dec (sin x) 1) as [Hle'|Hnle']. split; assumption. cut (1 < sin x). intro; @@ -670,11 +669,11 @@ Proof. replace (Un 0%nat + - Un 1%nat + Un 2%nat + - Un 3%nat) with (Un 0%nat - Un 1%nat + (Un 2%nat - Un 3%nat)); [ idtac | ring ]. apply Rplus_lt_0_compat. - unfold Rminus in |- *; apply Rplus_lt_reg_r with (Un 1%nat); + unfold Rminus in |- *; apply Rplus_lt_reg_l with (Un 1%nat); rewrite Rplus_0_r; rewrite (Rplus_comm (Un 1%nat)); rewrite Rplus_assoc; rewrite Rplus_opp_l; rewrite Rplus_0_r; apply H1. - unfold Rminus in |- *; apply Rplus_lt_reg_r with (Un 3%nat); + unfold Rminus in |- *; apply Rplus_lt_reg_l with (Un 3%nat); rewrite Rplus_0_r; rewrite (Rplus_comm (Un 3%nat)); rewrite Rplus_assoc; rewrite Rplus_opp_l; rewrite Rplus_0_r; apply H1. @@ -722,7 +721,7 @@ Proof. unfold INR in |- *. replace ((2 * x + 1 + 1 + 1) * (2 * x + 1 + 1)) with (4 * x * x + 10 * x + 6); [ idtac | ring ]. - apply Rplus_lt_reg_r with (-4); rewrite Rplus_opp_l; + 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. @@ -1201,7 +1200,7 @@ Proof. replace (- (PI - x)) with (x - PI). replace (- (PI - y)) with (y - PI). intros; change (sin (y - PI) < sin (x - PI)) in H8; - apply Rplus_lt_reg_r with (- PI); rewrite Rplus_comm; + 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 (sin_increasing_0 (y - PI) (x - PI) H4 H5 H6 H7 H8). @@ -1273,7 +1272,7 @@ Proof. replace (-3 * (PI / 2) + 2 * PI) with (PI / 2). replace (-3 * (PI / 2) + PI) with (- (PI / 2)). clear H1 H2 H3 H4; intros H1 H2 H3 H4; - apply Rplus_lt_reg_r with (-3 * (PI / 2)); + 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)). apply (sin_increasing_0 (x - 3 * (PI / 2)) (y - 3 * (PI / 2)) H4 H3 H2 H1 H5). @@ -1352,7 +1351,7 @@ Proof. generalize (Rplus_le_compat_l PI 0 y H1); generalize (Rplus_le_compat_l PI y PI H2); rewrite Rplus_0_r. rewrite <- double. - clear H H0 H1 H2 H3; intros; apply Rplus_lt_reg_r with PI; + clear H H0 H1 H2 H3; intros; apply Rplus_lt_reg_l with PI; apply (cos_increasing_0 (PI + y) (PI + x) H0 H H2 H1 H4). Qed. @@ -1919,7 +1918,7 @@ Proof. apply (Rmult_lt_reg_r PI); [apply PI_RGT_0|rewrite Rmult_1_l]. replace (3*(PI/2)) with (PI/2 + PI) in GT by field. rewrite Rplus_comm in GT. - now apply Rplus_lt_reg_r in GT. } + now apply Rplus_lt_reg_l in GT. } omega. Qed. diff --git a/theories/Reals/Rtrigo_alt.v b/theories/Reals/Rtrigo_alt.v index cdc96f98..3d36cb34 100644 --- a/theories/Reals/Rtrigo_alt.v +++ b/theories/Reals/Rtrigo_alt.v @@ -1,6 +1,6 @@ (************************************************************************) (* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *) +(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *) (* \VV/ **************************************************************) (* // * This file is distributed under the terms of the *) (* * GNU Lesser General Public License Version 2.1 *) @@ -134,13 +134,13 @@ Proof. apply (fun m n p:nat => mult_le_compat_l p n m); apply le_n_S; assumption. apply le_n_Sn. ring. - assert (X := exist_sin (Rsqr a)); elim X; intros. - cut (x = sin a / a). - intro; rewrite H3 in p; unfold sin_in in p; unfold infinite_sum in p; - unfold R_dist in p; unfold Un_cv; unfold R_dist; + unfold sin. + destruct (exist_sin (Rsqr a)) as (x,p). + unfold sin_in, infinite_sum, R_dist in p; + unfold Un_cv, R_dist; intros. cut (0 < eps / Rabs a). - intro; elim (p _ H5); intros N H6. + intro H4; destruct (p _ H4) as (N,H6). exists N; intros. replace (sum_f_R0 (tg_alt Un) n0) with (a * (1 - sum_f_R0 (fun i:nat => sin_n i * Rsqr a ^ i) (S n0))). @@ -151,12 +151,12 @@ Proof. rewrite Rplus_opp_l; rewrite Rplus_0_r; apply Rmult_lt_reg_l with (/ Rabs a). apply Rinv_0_lt_compat; apply Rabs_pos_lt; assumption. pattern (/ Rabs a) at 1; rewrite <- (Rabs_Rinv a Hyp_a). - rewrite <- Rabs_mult; rewrite Rmult_plus_distr_l; rewrite <- Rmult_assoc; - rewrite <- Rinv_l_sym; [ rewrite Rmult_1_l | assumption ]; - rewrite (Rmult_comm (/ a)); rewrite (Rmult_comm (/ Rabs a)); - rewrite <- Rabs_Ropp; rewrite Ropp_plus_distr; rewrite Ropp_involutive; - unfold Rminus, Rdiv in H6; apply H6; unfold ge; - apply le_trans with n0; [ exact H7 | apply le_n_Sn ]. + rewrite <- Rabs_mult, Rmult_plus_distr_l, <- 2!Rmult_assoc, <- Rinv_l_sym; + [ rewrite Rmult_1_l | assumption ]; + rewrite (Rmult_comm (/ Rabs a)), + <- Rabs_Ropp, Ropp_plus_distr, Ropp_involutive, Rmult_1_l. + unfold Rminus, Rdiv in H6. apply H6; unfold ge; + apply le_trans with n0; [ exact H5 | apply le_n_Sn ]. rewrite (decomp_sum (fun i:nat => sin_n i * Rsqr a ^ i) (S n0)). replace (sin_n 0) with 1. simpl; rewrite Rmult_1_r; unfold Rminus; @@ -176,13 +176,6 @@ Proof. unfold Rdiv; apply Rmult_lt_0_compat. assumption. apply Rinv_0_lt_compat; apply Rabs_pos_lt; assumption. - unfold sin; case (exist_sin (Rsqr a)). - intros; cut (x = x0). - intro; rewrite H3; unfold Rdiv. - symmetry ; apply Rinv_r_simpl_m; assumption. - unfold sin_in in p; unfold sin_in in s; eapply uniqueness_sum. - apply p. - apply s. intros; elim H2; intros. replace (sin a - a) with (- (a - sin a)); [ idtac | ring ]. split; apply Ropp_le_contravar; assumption. @@ -318,12 +311,10 @@ Proof. apply le_n_2n. apply (fun m n p:nat => mult_le_compat_l p n m); apply le_n_Sn. apply (fun m n p:nat => mult_le_compat_l p n m); apply le_n_S; assumption. - assert (X := exist_cos (Rsqr a0)); elim X; intros. - cut (x = cos a0). - intro; rewrite H4 in p; unfold cos_in in p; unfold infinite_sum in p; - unfold R_dist in p; unfold Un_cv; unfold R_dist; - intros. - elim (p _ H5); intros N H6. + unfold cos. destruct (exist_cos (Rsqr a0)) as (x,p). + unfold cos_in, infinite_sum, R_dist in p; + unfold Un_cv, R_dist; intros. + destruct (p _ H4) as (N,H6). exists N; intros. replace (sum_f_R0 (tg_alt Un) n1) with (1 - sum_f_R0 (fun i:nat => cos_n i * Rsqr a0 ^ i) (S n1)). @@ -334,7 +325,7 @@ Proof. rewrite Ropp_plus_distr; rewrite Ropp_involutive; unfold Rminus in H6; apply H6. unfold ge; apply le_trans with n1. - exact H7. + exact H5. apply le_n_Sn. rewrite (decomp_sum (fun i:nat => cos_n i * Rsqr a0 ^ i) (S n1)). replace (cos_n 0) with 1. @@ -354,10 +345,6 @@ Proof. unfold cos_n; unfold Rdiv; simpl; rewrite Rinv_1; rewrite Rmult_1_r; reflexivity. apply lt_O_Sn. - unfold cos; case (exist_cos (Rsqr a0)); intros; unfold cos_in in p; - unfold cos_in in c; eapply uniqueness_sum. - apply p. - apply c. intros; elim H3; intros; replace (cos a0 - 1) with (- (1 - cos a0)); [ idtac | ring ]. split; apply Ropp_le_contravar; assumption. @@ -394,8 +381,7 @@ Proof. replace (2 * n0 + 1)%nat with (S (2 * n0)). apply lt_O_Sn. ring. - intros; case (total_order_T 0 a); intro. - elim s; intro. + intros; destruct (total_order_T 0 a) as [[Hlt|Heq]|Hgt]. apply H; [ left; assumption | assumption ]. apply H; [ right; assumption | assumption ]. cut (0 < - a). diff --git a/theories/Reals/Rtrigo_calc.v b/theories/Reals/Rtrigo_calc.v index 2ad65a92..281c152b 100644 --- a/theories/Reals/Rtrigo_calc.v +++ b/theories/Reals/Rtrigo_calc.v @@ -1,6 +1,6 @@ (************************************************************************) (* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *) +(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *) (* \VV/ **************************************************************) (* // * This file is distributed under the terms of the *) (* * GNU Lesser General Public License Version 2.1 *) diff --git a/theories/Reals/Rtrigo_def.v b/theories/Reals/Rtrigo_def.v index 60df6f78..ef3e31f1 100644 --- a/theories/Reals/Rtrigo_def.v +++ b/theories/Reals/Rtrigo_def.v @@ -1,6 +1,6 @@ (************************************************************************) (* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *) +(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *) (* \VV/ **************************************************************) (* // * This file is distributed under the terms of the *) (* * GNU Lesser General Public License Version 2.1 *) @@ -221,6 +221,7 @@ Proof. Qed. Lemma cosn_no_R0 : forall n:nat, cos_n n <> 0. +Proof. intro; unfold cos_n; unfold Rdiv; apply prod_neq_R0. apply pow_nonzero; discrR. apply Rinv_neq_0_compat. @@ -233,6 +234,7 @@ Definition cos_in (x l:R) : Prop := (**********) Lemma exist_cos : forall x:R, { l:R | cos_in x l }. +Proof. intro; generalize (Alembert_C3 cos_n x cosn_no_R0 Alembert_cos). unfold Pser, cos_in; trivial. Qed. @@ -338,7 +340,7 @@ Proof. apply INR_eq; repeat rewrite S_INR; rewrite plus_INR; repeat rewrite mult_INR; rewrite plus_INR; rewrite mult_INR; repeat rewrite S_INR; replace (INR 0) with 0; [ ring | reflexivity ]. -Defined. +Qed. Lemma sin_no_R0 : forall n:nat, sin_n n <> 0. Proof. diff --git a/theories/Reals/Rtrigo_fun.v b/theories/Reals/Rtrigo_fun.v index bc2f62a8..b921ee7b 100644 --- a/theories/Reals/Rtrigo_fun.v +++ b/theories/Reals/Rtrigo_fun.v @@ -1,6 +1,6 @@ (************************************************************************) (* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *) +(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *) (* \VV/ **************************************************************) (* // * This file is distributed under the terms of the *) (* * GNU Lesser General Public License Version 2.1 *) @@ -20,80 +20,79 @@ Local Open Scope R_scope. Lemma Alembert_exp : Un_cv (fun n:nat => Rabs (/ INR (fact (S n)) * / / INR (fact n))) 0. Proof. - unfold Un_cv; intros; elim (Rgt_dec eps 1); intro. - split with 0%nat; intros; rewrite (simpl_fact n); unfold R_dist; - rewrite (Rminus_0_r (Rabs (/ INR (S n)))); - rewrite (Rabs_Rabsolu (/ INR (S n))); cut (/ INR (S n) > 0). - intro; rewrite (Rabs_pos_eq (/ INR (S n))). - cut (/ eps - 1 < 0). - intro; generalize (Rlt_le_trans (/ eps - 1) 0 (INR n) H2 (pos_INR n)); - clear H2; intro; unfold Rminus in H2; - generalize (Rplus_lt_compat_l 1 (/ eps + -1) (INR n) H2); - replace (1 + (/ eps + -1)) with (/ eps); [ clear H2; intro | ring ]. - rewrite (Rplus_comm 1 (INR n)) in H2; rewrite <- (S_INR n) in H2; - generalize (Rmult_gt_0_compat (/ INR (S n)) eps H1 H); - intro; unfold Rgt in H3; - generalize (Rmult_lt_compat_l (/ INR (S n) * eps) (/ eps) (INR (S n)) H3 H2); - intro; rewrite (Rmult_assoc (/ INR (S n)) eps (/ eps)) in H4; - rewrite (Rinv_r eps (Rlt_dichotomy_converse eps 0 (or_intror (eps < 0) H))) - in H4; rewrite (let (H1, H2) := Rmult_ne (/ INR (S n)) in H1) in H4; - rewrite (Rmult_comm (/ INR (S n))) in H4; - rewrite (Rmult_assoc eps (/ INR (S n)) (INR (S n))) in H4; - rewrite (Rinv_l (INR (S n)) (not_O_INR (S n) (not_eq_sym (O_S n)))) in H4; - rewrite (let (H1, H2) := Rmult_ne eps in H1) in H4; - assumption. - apply Rlt_minus; unfold Rgt in a; rewrite <- Rinv_1; - apply (Rinv_lt_contravar 1 eps); auto; - rewrite (let (H1, H2) := Rmult_ne eps in H2); unfold Rgt in H; - assumption. - unfold Rgt in H1; apply Rlt_le; assumption. - unfold Rgt; apply Rinv_0_lt_compat; apply lt_INR_0; apply lt_O_Sn. -(**) - cut (0 <= up (/ eps - 1))%Z. - intro; elim (IZN (up (/ eps - 1)) H0); intros; split with x; intros; - rewrite (simpl_fact n); unfold R_dist; + unfold Un_cv; intros; destruct (Rgt_dec eps 1) as [Hgt|Hnotgt]. + - split with 0%nat; intros; rewrite (simpl_fact n); unfold R_dist; rewrite (Rminus_0_r (Rabs (/ INR (S n)))); rewrite (Rabs_Rabsolu (/ INR (S n))); cut (/ INR (S n) > 0). - intro; rewrite (Rabs_pos_eq (/ INR (S n))). - cut (/ eps - 1 < INR x). - intro ; - generalize - (Rlt_le_trans (/ eps - 1) (INR x) (INR n) H4 - (le_INR x n H2)); - clear H4; intro; unfold Rminus in H4; - generalize (Rplus_lt_compat_l 1 (/ eps + -1) (INR n) H4); - replace (1 + (/ eps + -1)) with (/ eps); [ clear H4; intro | ring ]. - rewrite (Rplus_comm 1 (INR n)) in H4; rewrite <- (S_INR n) in H4; - generalize (Rmult_gt_0_compat (/ INR (S n)) eps H3 H); - intro; unfold Rgt in H5; - generalize (Rmult_lt_compat_l (/ INR (S n) * eps) (/ eps) (INR (S n)) H5 H4); - intro; rewrite (Rmult_assoc (/ INR (S n)) eps (/ eps)) in H6; - rewrite (Rinv_r eps (Rlt_dichotomy_converse eps 0 (or_intror (eps < 0) H))) - in H6; rewrite (let (H1, H2) := Rmult_ne (/ INR (S n)) in H1) in H6; - rewrite (Rmult_comm (/ INR (S n))) in H6; - rewrite (Rmult_assoc eps (/ INR (S n)) (INR (S n))) in H6; - rewrite (Rinv_l (INR (S n)) (not_O_INR (S n) (not_eq_sym (O_S n)))) in H6; - rewrite (let (H1, H2) := Rmult_ne eps in H1) in H6; - assumption. - cut (IZR (up (/ eps - 1)) = IZR (Z.of_nat x)); - [ intro | rewrite H1; trivial ]. - elim (archimed (/ eps - 1)); intros; clear H6; unfold Rgt in H5; - rewrite H4 in H5; rewrite INR_IZR_INZ; assumption. - unfold Rgt in H1; apply Rlt_le; assumption. - unfold Rgt; apply Rinv_0_lt_compat; apply lt_INR_0; apply lt_O_Sn. - apply (le_O_IZR (up (/ eps - 1))); - apply (Rle_trans 0 (/ eps - 1) (IZR (up (/ eps - 1)))). - generalize (Rnot_gt_le eps 1 b); clear b; unfold Rle; intro; elim H0; - clear H0; intro. - left; unfold Rgt in H; - generalize (Rmult_lt_compat_l (/ eps) eps 1 (Rinv_0_lt_compat eps H) H0); - rewrite - (Rinv_l eps - (not_eq_sym (Rlt_dichotomy_converse 0 eps (or_introl (0 > eps) H)))) - ; rewrite (let (H1, H2) := Rmult_ne (/ eps) in H1); - intro; fold (/ eps - 1 > 0); apply Rgt_minus; - unfold Rgt; assumption. - right; rewrite H0; rewrite Rinv_1; symmetry; apply Rminus_diag_eq; auto. - elim (archimed (/ eps - 1)); intros; clear H1; unfold Rgt in H0; apply Rlt_le; - assumption. + intro; rewrite (Rabs_pos_eq (/ INR (S n))). + cut (/ eps - 1 < 0). + intro H2; generalize (Rlt_le_trans (/ eps - 1) 0 (INR n) H2 (pos_INR n)); + clear H2; intro; unfold Rminus in H2; + generalize (Rplus_lt_compat_l 1 (/ eps + -1) (INR n) H2); + replace (1 + (/ eps + -1)) with (/ eps); [ clear H2; intro | ring ]. + rewrite (Rplus_comm 1 (INR n)) in H2; rewrite <- (S_INR n) in H2; + generalize (Rmult_gt_0_compat (/ INR (S n)) eps H1 H); + intro; unfold Rgt in H3; + generalize (Rmult_lt_compat_l (/ INR (S n) * eps) (/ eps) (INR (S n)) H3 H2); + intro; rewrite (Rmult_assoc (/ INR (S n)) eps (/ eps)) in H4; + rewrite (Rinv_r eps (Rlt_dichotomy_converse eps 0 (or_intror (eps < 0) H))) + in H4; rewrite (let (H1, H2) := Rmult_ne (/ INR (S n)) in H1) in H4; + rewrite (Rmult_comm (/ INR (S n))) in H4; + rewrite (Rmult_assoc eps (/ INR (S n)) (INR (S n))) in H4; + rewrite (Rinv_l (INR (S n)) (not_O_INR (S n) (not_eq_sym (O_S n)))) in H4; + rewrite (let (H1, H2) := Rmult_ne eps in H1) in H4; + assumption. + apply Rlt_minus; unfold Rgt in Hgt; rewrite <- Rinv_1; + apply (Rinv_lt_contravar 1 eps); auto; + rewrite (let (H1, H2) := Rmult_ne eps in H2); unfold Rgt in H; + assumption. + unfold Rgt in H1; apply Rlt_le; assumption. + unfold Rgt; apply Rinv_0_lt_compat; apply lt_INR_0; apply lt_O_Sn. + - cut (0 <= up (/ eps - 1))%Z. + intro; elim (IZN (up (/ eps - 1)) H0); intros; split with x; intros; + rewrite (simpl_fact n); unfold R_dist; + rewrite (Rminus_0_r (Rabs (/ INR (S n)))); + rewrite (Rabs_Rabsolu (/ INR (S n))); cut (/ INR (S n) > 0). + intro; rewrite (Rabs_pos_eq (/ INR (S n))). + cut (/ eps - 1 < INR x). + intro ; + generalize + (Rlt_le_trans (/ eps - 1) (INR x) (INR n) H4 + (le_INR x n H2)); + clear H4; intro; unfold Rminus in H4; + generalize (Rplus_lt_compat_l 1 (/ eps + -1) (INR n) H4); + replace (1 + (/ eps + -1)) with (/ eps); [ clear H4; intro | ring ]. + rewrite (Rplus_comm 1 (INR n)) in H4; rewrite <- (S_INR n) in H4; + generalize (Rmult_gt_0_compat (/ INR (S n)) eps H3 H); + intro; unfold Rgt in H5; + generalize (Rmult_lt_compat_l (/ INR (S n) * eps) (/ eps) (INR (S n)) H5 H4); + intro; rewrite (Rmult_assoc (/ INR (S n)) eps (/ eps)) in H6; + rewrite (Rinv_r eps (Rlt_dichotomy_converse eps 0 (or_intror (eps < 0) H))) + in H6; rewrite (let (H1, H2) := Rmult_ne (/ INR (S n)) in H1) in H6; + rewrite (Rmult_comm (/ INR (S n))) in H6; + rewrite (Rmult_assoc eps (/ INR (S n)) (INR (S n))) in H6; + rewrite (Rinv_l (INR (S n)) (not_O_INR (S n) (not_eq_sym (O_S n)))) in H6; + rewrite (let (H1, H2) := Rmult_ne eps in H1) in H6; + assumption. + cut (IZR (up (/ eps - 1)) = IZR (Z.of_nat x)); + [ intro | rewrite H1; trivial ]. + elim (archimed (/ eps - 1)); intros; clear H6; unfold Rgt in H5; + rewrite H4 in H5; rewrite INR_IZR_INZ; assumption. + unfold Rgt in H1; apply Rlt_le; assumption. + unfold Rgt; apply Rinv_0_lt_compat; apply lt_INR_0; apply lt_O_Sn. + apply (le_O_IZR (up (/ eps - 1))); + apply (Rle_trans 0 (/ eps - 1) (IZR (up (/ eps - 1)))). + generalize (Rnot_gt_le eps 1 Hnotgt); clear Hnotgt; unfold Rle; intro; elim H0; + clear H0; intro. + left; unfold Rgt in H; + generalize (Rmult_lt_compat_l (/ eps) eps 1 (Rinv_0_lt_compat eps H) H0); + rewrite + (Rinv_l eps + (not_eq_sym (Rlt_dichotomy_converse 0 eps (or_introl (0 > eps) H)))) + ; rewrite (let (H1, H2) := Rmult_ne (/ eps) in H1); + intro; fold (/ eps - 1 > 0); apply Rgt_minus; + unfold Rgt; assumption. + right; rewrite H0; rewrite Rinv_1; symmetry; apply Rminus_diag_eq; auto. + elim (archimed (/ eps - 1)); intros; clear H1; unfold Rgt in H0; apply Rlt_le; + assumption. Qed. diff --git a/theories/Reals/Rtrigo_reg.v b/theories/Reals/Rtrigo_reg.v index 4e3d41e3..7845e6c4 100644 --- a/theories/Reals/Rtrigo_reg.v +++ b/theories/Reals/Rtrigo_reg.v @@ -1,6 +1,6 @@ (************************************************************************) (* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *) +(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *) (* \VV/ **************************************************************) (* // * This file is distributed under the terms of the *) (* * GNU Lesser General Public License Version 2.1 *) @@ -59,7 +59,7 @@ Proof. sum_f_R0 (fun k:nat => Rabs (/ INR (fact (2 * k + 1)) * r ^ (2 * k))) n) l }. - intro X; elim X; intros. + intros (x,p). exists x. split. apply p. @@ -176,14 +176,14 @@ Proof. intro; rewrite H9 in H8; rewrite H10 in H8. apply H8. unfold SFL, sin. - case (cv h); intros. - case (exist_sin (Rsqr h)); intros. + case (cv h) as (x,HUn). + case (exist_sin (Rsqr h)) as (x0,Hsin). unfold Rdiv; rewrite (Rinv_r_simpl_m h x0 H6). eapply UL_sequence. - apply u. - unfold sin_in in s; unfold sin_n, infinite_sum in s; + apply HUn. + unfold sin_in in Hsin; unfold sin_n, infinite_sum in Hsin; unfold SP, fn, Un_cv; intros. - elim (s _ H10); intros N0 H11. + elim (Hsin _ H10); intros N0 H11. exists N0; intros. unfold R_dist; unfold R_dist in H11. replace @@ -194,9 +194,9 @@ Proof. apply sum_eq; intros; apply Rmult_eq_compat_l; unfold Rsqr; rewrite pow_sqr; reflexivity. unfold SFL, sin. - case (cv 0); intros. + case (cv 0) as (?,HUn). eapply UL_sequence. - apply u. + apply HUn. unfold SP, fn; unfold Un_cv; intros; exists 1%nat; intros. unfold R_dist; replace diff --git a/theories/Reals/SeqProp.v b/theories/Reals/SeqProp.v index fb2eacee..9a6fb945 100644 --- a/theories/Reals/SeqProp.v +++ b/theories/Reals/SeqProp.v @@ -1,6 +1,6 @@ (************************************************************************) (* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *) +(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *) (* \VV/ **************************************************************) (* // * This file is distributed under the terms of the *) (* * GNU Lesser General Public License Version 2.1 *) @@ -10,6 +10,7 @@ Require Import Rbase. Require Import Rfunctions. Require Import Rseries. Require Import Max. +Require Import Omega. Local Open Scope R_scope. (*****************************************************************) @@ -27,7 +28,7 @@ Lemma growing_cv : forall Un:nat -> R, Un_growing Un -> has_ub Un -> { l:R | Un_cv Un l }. Proof. intros Un Hug Heub. - exists (projT1 (completeness (EUn Un) Heub (EUn_noempty Un))). + exists (proj1_sig (completeness (EUn Un) Heub (EUn_noempty Un))). destruct (completeness _ Heub (EUn_noempty Un)) as (l, H). now apply Un_cv_crit_lub. Qed. @@ -52,8 +53,7 @@ Proof. apply growing_cv. apply decreasing_growing; assumption. exact H0. - intro X. - elim X; intros. + intros (x,p). exists (- x). unfold Un_cv in p. unfold R_dist in p. @@ -150,7 +150,7 @@ Definition sequence_lb (Un:nat -> R) (pr:has_lb Un) (* Compatibility *) Notation sequence_majorant := sequence_ub (only parsing). Notation sequence_minorant := sequence_lb (only parsing). - +Unset Standard Proposition Elimination Names. Lemma Wn_decreasing : forall (Un:nat -> R) (pr:has_ub Un), Un_decreasing (sequence_ub Un pr). Proof. @@ -158,21 +158,15 @@ Proof. unfold Un_decreasing. intro. unfold sequence_ub. - assert (H := ub_to_lub (fun k:nat => Un (S n + k)%nat) (maj_ss Un (S n) pr)). - assert (H0 := ub_to_lub (fun k:nat => Un (n + k)%nat) (maj_ss Un n pr)). - elim H; intros. - elim H0; intros. + pose proof (ub_to_lub (fun k:nat => Un (S n + k)%nat) (maj_ss Un (S n) pr)) as (x,(H1,H2)). + pose proof (ub_to_lub (fun k:nat => Un (n + k)%nat) (maj_ss Un n pr)) as (x0,(H3,H4)). cut (lub (fun k:nat => Un (S n + k)%nat) (maj_ss Un (S n) pr) = x); [ intro Maj1; rewrite Maj1 | idtac ]. cut (lub (fun k:nat => Un (n + k)%nat) (maj_ss Un n pr) = x0); [ intro Maj2; rewrite Maj2 | idtac ]. - unfold is_lub in p. - unfold is_lub in p0. - elim p; intros. apply H2. - elim p0; intros. unfold is_upper_bound. - intros. + intros x1 H5. unfold is_upper_bound in H3. apply H3. elim H5; intros. @@ -183,12 +177,10 @@ Proof. cut (is_lub (EUn (fun k:nat => Un (n + k)%nat)) (lub (fun k:nat => Un (n + k)%nat) (maj_ss Un n pr))). - intro. - unfold is_lub in p0; unfold is_lub in H1. - elim p0; intros; elim H1; intros. - assert (H6 := H5 x0 H2). + intros (H5,H6). + assert (H7 := H6 x0 H3). assert - (H7 := H3 (lub (fun k:nat => Un (n + k)%nat) (maj_ss Un n pr)) H4). + (H8 := H4 (lub (fun k:nat => Un (n + k)%nat) (maj_ss Un n pr)) H5). apply Rle_antisym; assumption. unfold lub. case (ub_to_lub (fun k:nat => Un (n + k)%nat) (maj_ss Un n pr)). @@ -196,13 +188,11 @@ Proof. cut (is_lub (EUn (fun k:nat => Un (S n + k)%nat)) (lub (fun k:nat => Un (S n + k)%nat) (maj_ss Un (S n) pr))). - intro. - unfold is_lub in p; unfold is_lub in H1. - elim p; intros; elim H1; intros. - assert (H6 := H5 x H2). + intros (H5,H6). + assert (H7 := H6 x H1). assert - (H7 := - H3 (lub (fun k:nat => Un (S n + k)%nat) (maj_ss Un (S n) pr)) H4). + (H8 := + H2 (lub (fun k:nat => Un (S n + k)%nat) (maj_ss Un (S n) pr)) H5). apply Rle_antisym; assumption. unfold lub. case (ub_to_lub (fun k:nat => Un (S n + k)%nat) (maj_ss Un (S n) pr)). @@ -460,8 +450,7 @@ Lemma cond_eq : forall x y:R, (forall eps:R, 0 < eps -> Rabs (x - y) < eps) -> x = y. Proof. intros. - case (total_order_T x y); intro. - elim s; intro. + destruct (total_order_T x y) as [[Hlt|Heq]|Hgt]. cut (0 < y - x). intro. assert (H1 := H (y - x) H0). @@ -470,7 +459,7 @@ Proof. rewrite Rabs_right in H1. elim (Rlt_irrefl _ H1). left; assumption. - apply Rplus_lt_reg_r with x. + apply Rplus_lt_reg_l with x. rewrite Rplus_0_r; replace (x + (y - x)) with y; [ assumption | ring ]. assumption. cut (0 < x - y). @@ -479,7 +468,7 @@ Proof. rewrite Rabs_right in H1. elim (Rlt_irrefl _ H1). left; assumption. - apply Rplus_lt_reg_r with y. + apply Rplus_lt_reg_l with y. rewrite Rplus_0_r; replace (y + (x - y)) with x; [ assumption | ring ]. Qed. @@ -860,7 +849,7 @@ Proof. split. pattern k at 1; rewrite <- Rplus_0_r; apply Rplus_lt_compat_l. unfold Rdiv; apply Rmult_lt_0_compat. - apply Rplus_lt_reg_r with k; rewrite Rplus_0_r; replace (k + (1 - k)) with 1; + apply Rplus_lt_reg_l with k; rewrite Rplus_0_r; replace (k + (1 - k)) with 1; [ elim H; intros; assumption | ring ]. apply Rinv_0_lt_compat; prove_sup0. apply Rmult_lt_reg_l with 2. @@ -881,12 +870,12 @@ Proof. apply Rle_lt_trans with (Rabs (Rabs (An (S n) / An n) - k) + Rabs k). apply Rabs_triang. rewrite (Rabs_right k). - apply Rplus_lt_reg_r with (- k); rewrite <- (Rplus_comm k); + apply Rplus_lt_reg_l with (- k); rewrite <- (Rplus_comm k); repeat rewrite <- Rplus_assoc; rewrite Rplus_opp_l; repeat rewrite Rplus_0_l; apply H4. apply Rle_ge; elim H; intros; assumption. unfold Rdiv; apply Rmult_lt_0_compat. - apply Rplus_lt_reg_r with k; rewrite Rplus_0_r; elim H; intros; + apply Rplus_lt_reg_l with k; rewrite Rplus_0_r; elim H; intros; replace (k + (1 - k)) with 1; [ assumption | ring ]. apply Rinv_0_lt_compat; prove_sup0. Qed. @@ -896,8 +885,7 @@ Lemma growing_ineq : forall (Un:nat -> R) (l:R), Un_growing Un -> Un_cv Un l -> forall n:nat, Un n <= l. Proof. - intros; case (total_order_T (Un n) l); intro. - elim s; intro. + intros; destruct (total_order_T (Un n) l) as [[Hlt|Heq]|Hgt]. left; assumption. right; assumption. cut (0 < Un n - l). @@ -916,7 +904,7 @@ Proof. apply tech9. assumption. unfold N; apply le_max_l. - apply Rplus_lt_reg_r with l. + apply Rplus_lt_reg_l with l. rewrite Rplus_0_r. replace (l + (Un n - l)) with (Un n); [ assumption | ring ]. Qed. @@ -1102,11 +1090,11 @@ Proof. apply (cv_infty_cv_R0 (fun n:nat => INR (S n))). intro; apply not_O_INR; discriminate. assumption. - unfold cv_infty; intro; case (total_order_T M0 0); intro. - elim s; intro. + unfold cv_infty; intro; + destruct (total_order_T M0 0) as [[Hlt|Heq]|Hgt]. exists 0%nat; intros. apply Rlt_trans with 0; [ assumption | apply lt_INR_0; apply lt_O_Sn ]. - exists 0%nat; intros; rewrite b; apply lt_INR_0; apply lt_O_Sn. + exists 0%nat; intros; rewrite Heq; apply lt_INR_0; apply lt_O_Sn. set (M0_z := up M0). assert (H10 := archimed M0). cut (0 <= M0_z)%Z. diff --git a/theories/Reals/SeqSeries.v b/theories/Reals/SeqSeries.v index 5f2173c7..25fe4848 100644 --- a/theories/Reals/SeqSeries.v +++ b/theories/Reals/SeqSeries.v @@ -1,6 +1,6 @@ (************************************************************************) (* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *) +(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *) (* \VV/ **************************************************************) (* // * This file is distributed under the terms of the *) (* * GNU Lesser General Public License Version 2.1 *) @@ -222,39 +222,37 @@ Proof. intro; apply Rle_lt_trans with (R_dist (sum_f_R0 Bn n) (sum_f_R0 Bn m)). assumption. apply H2; assumption. - assert (H5 := lt_eq_lt_dec n m). - elim H5; intro. - elim a; intro. - rewrite (tech2 An n m); [ idtac | assumption ]. - rewrite (tech2 Bn n m); [ idtac | assumption ]. - unfold R_dist; unfold Rminus; do 2 rewrite Ropp_plus_distr; - do 2 rewrite <- Rplus_assoc; do 2 rewrite Rplus_opp_r; - do 2 rewrite Rplus_0_l; do 2 rewrite Rabs_Ropp; repeat rewrite Rabs_right. - apply sum_Rle; intros. - elim (H (S n + n0)%nat); intros. - apply H8. - apply Rle_ge; apply cond_pos_sum; intro. - elim (H (S n + n0)%nat); intros. - apply Rle_trans with (An (S n + n0)%nat); assumption. - apply Rle_ge; apply cond_pos_sum; intro. - elim (H (S n + n0)%nat); intros; assumption. - rewrite b; unfold R_dist; unfold Rminus; + destruct (lt_eq_lt_dec n m) as [[| -> ]|]. + - rewrite (tech2 An n m); [ idtac | assumption ]. + rewrite (tech2 Bn n m); [ idtac | assumption ]. + unfold R_dist; unfold Rminus; do 2 rewrite Ropp_plus_distr; + do 2 rewrite <- Rplus_assoc; do 2 rewrite Rplus_opp_r; + do 2 rewrite Rplus_0_l; do 2 rewrite Rabs_Ropp; repeat rewrite Rabs_right. + apply sum_Rle; intros. + elim (H (S n + n0)%nat); intros H7 H8. + apply H8. + apply Rle_ge; apply cond_pos_sum; intro. + elim (H (S n + n0)%nat); intros. + apply Rle_trans with (An (S n + n0)%nat); assumption. + apply Rle_ge; apply cond_pos_sum; intro. + elim (H (S n + n0)%nat); intros; assumption. + - unfold R_dist; unfold Rminus; do 2 rewrite Rplus_opp_r; rewrite Rabs_R0; right; reflexivity. - rewrite (tech2 An m n); [ idtac | assumption ]. - rewrite (tech2 Bn m n); [ idtac | assumption ]. - unfold R_dist; unfold Rminus; do 2 rewrite Rplus_assoc; - rewrite (Rplus_comm (sum_f_R0 An m)); rewrite (Rplus_comm (sum_f_R0 Bn m)); - do 2 rewrite Rplus_assoc; do 2 rewrite Rplus_opp_l; - do 2 rewrite Rplus_0_r; repeat rewrite Rabs_right. - apply sum_Rle; intros. - elim (H (S m + n0)%nat); intros; apply H8. - apply Rle_ge; apply cond_pos_sum; intro. - elim (H (S m + n0)%nat); intros. - apply Rle_trans with (An (S m + n0)%nat); assumption. - apply Rle_ge. - apply cond_pos_sum; intro. - elim (H (S m + n0)%nat); intros; assumption. + - rewrite (tech2 An m n); [ idtac | assumption ]. + rewrite (tech2 Bn m n); [ idtac | assumption ]. + unfold R_dist; unfold Rminus; do 2 rewrite Rplus_assoc; + rewrite (Rplus_comm (sum_f_R0 An m)); rewrite (Rplus_comm (sum_f_R0 Bn m)); + do 2 rewrite Rplus_assoc; do 2 rewrite Rplus_opp_l; + do 2 rewrite Rplus_0_r; repeat rewrite Rabs_right. + apply sum_Rle; intros. + elim (H (S m + n0)%nat); intros H7 H8; apply H8. + apply Rle_ge; apply cond_pos_sum; intro. + elim (H (S m + n0)%nat); intros. + apply Rle_trans with (An (S m + n0)%nat); assumption. + apply Rle_ge. + apply cond_pos_sum; intro. + elim (H (S m + n0)%nat); intros; assumption. Qed. (** Cesaro's theorem *) @@ -361,7 +359,7 @@ Proof with trivial. replace (sum_f_R0 (fun k:nat => An k * (Bn k - l)) n) with (sum_f_R0 (fun k:nat => An k * Bn k) n + sum_f_R0 (fun k:nat => An k * - l) n)... - rewrite <- (scal_sum An n (- l)); field... + rewrite <- (scal_sum An n (- l)); field... rewrite <- plus_sum; apply sum_eq; intros; ring... Qed. @@ -375,11 +373,11 @@ Proof with trivial. assert (H1 : forall n:nat, 0 < sum_f_R0 An n)... intro; apply tech1... assert (H2 : cv_infty (fun n:nat => sum_f_R0 An n))... - unfold cv_infty; intro; case (Rle_dec M 0); intro... + unfold cv_infty; intro; destruct (Rle_dec M 0) as [Hle|Hnle]... exists 0%nat; intros; apply Rle_lt_trans with 0... assert (H2 : 0 < M)... auto with real... - clear n; set (m := up M); elim (archimed M); intros; + clear Hnle; set (m := up M); elim (archimed M); intros; assert (H5 : (0 <= m)%Z)... apply le_IZR; unfold m; simpl; left; apply Rlt_trans with M... elim (IZN _ H5); intros; exists x; intros; unfold An; rewrite sum_cte; diff --git a/theories/Reals/SplitAbsolu.v b/theories/Reals/SplitAbsolu.v index 3557e2e9..64f4f1c9 100644 --- a/theories/Reals/SplitAbsolu.v +++ b/theories/Reals/SplitAbsolu.v @@ -1,6 +1,6 @@ (************************************************************************) (* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *) +(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *) (* \VV/ **************************************************************) (* // * This file is distributed under the terms of the *) (* * GNU Lesser General Public License Version 2.1 *) @@ -11,7 +11,7 @@ Require Import Rbasic_fun. Ltac split_case_Rabs := match goal with | |- context [(Rcase_abs ?X1)] => - case (Rcase_abs X1); try split_case_Rabs + destruct (Rcase_abs X1) as [?Hlt|?Hge]; try split_case_Rabs end. diff --git a/theories/Reals/SplitRmult.v b/theories/Reals/SplitRmult.v index 7380f8ad..fec28518 100644 --- a/theories/Reals/SplitRmult.v +++ b/theories/Reals/SplitRmult.v @@ -1,6 +1,6 @@ (************************************************************************) (* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *) +(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *) (* \VV/ **************************************************************) (* // * This file is distributed under the terms of the *) (* * GNU Lesser General Public License Version 2.1 *) diff --git a/theories/Reals/Sqrt_reg.v b/theories/Reals/Sqrt_reg.v index a74aeef2..dd8738e1 100644 --- a/theories/Reals/Sqrt_reg.v +++ b/theories/Reals/Sqrt_reg.v @@ -1,6 +1,6 @@ (************************************************************************) (* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *) +(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *) (* \VV/ **************************************************************) (* // * This file is distributed under the terms of the *) (* * GNU Lesser General Public License Version 2.1 *) @@ -18,8 +18,7 @@ Lemma sqrt_var_maj : Proof. intros; cut (0 <= 1 + h). intro; apply Rle_trans with (Rabs (sqrt (Rsqr (1 + h)) - 1)). - case (total_order_T h 0); intro. - elim s; intro. + destruct (total_order_T h 0) as [[Hlt|Heq]|Hgt]. repeat rewrite Rabs_left. unfold Rminus; do 2 rewrite <- (Rplus_comm (-1)). do 2 rewrite Ropp_plus_distr; rewrite Ropp_involutive; @@ -32,7 +31,7 @@ Proof. apply H0. pattern 1 at 2; rewrite <- Rplus_0_r; apply Rplus_le_compat_l; left; assumption. - apply Rplus_lt_reg_r with 1; rewrite Rplus_0_r; rewrite Rplus_comm; + apply Rplus_lt_reg_l with 1; rewrite Rplus_0_r; rewrite Rplus_comm; unfold Rminus; rewrite Rplus_assoc; rewrite Rplus_opp_l; rewrite Rplus_0_r. pattern 1 at 2; rewrite <- sqrt_1; apply sqrt_lt_1. @@ -43,7 +42,7 @@ Proof. assumption. apply H0. left; apply Rlt_0_1. - apply Rplus_lt_reg_r with 1; rewrite Rplus_0_r; rewrite Rplus_comm; + apply Rplus_lt_reg_l with 1; rewrite Rplus_0_r; rewrite Rplus_comm; unfold Rminus; rewrite Rplus_assoc; rewrite Rplus_opp_l; rewrite Rplus_0_r. pattern 1 at 2; rewrite <- sqrt_1; apply sqrt_lt_1. @@ -51,7 +50,7 @@ Proof. left; apply Rlt_0_1. pattern 1 at 2; rewrite <- Rplus_0_r; apply Rplus_lt_compat_l; assumption. - rewrite b; rewrite Rplus_0_r; rewrite Rsqr_1; rewrite sqrt_1; right; + rewrite Heq; rewrite Rplus_0_r; rewrite Rsqr_1; rewrite sqrt_1; right; reflexivity. repeat rewrite Rabs_right. unfold Rminus; do 2 rewrite <- (Rplus_comm (-1)); @@ -75,7 +74,7 @@ Proof. assumption. left; apply Rlt_0_1. apply H0. - apply Rle_ge; left; apply Rplus_lt_reg_r with 1. + apply Rle_ge; left; apply Rplus_lt_reg_l with 1. rewrite Rplus_0_r; rewrite Rplus_comm; unfold Rminus; rewrite Rplus_assoc; rewrite Rplus_opp_l; rewrite Rplus_0_r. pattern 1 at 1; rewrite <- sqrt_1; apply sqrt_lt_1. @@ -86,16 +85,15 @@ Proof. rewrite sqrt_Rsqr. replace (1 + h - 1) with h; [ right; reflexivity | ring ]. apply H0. - case (total_order_T h 0); intro. - elim s; intro. - rewrite (Rabs_left h a) in H. + destruct (total_order_T h 0) as [[Hlt|Heq]|Hgt]. + rewrite (Rabs_left h Hlt) in H. apply Rplus_le_reg_l with (- h). rewrite Rplus_0_r; rewrite Rplus_comm; rewrite Rplus_assoc; rewrite Rplus_opp_r; rewrite Rplus_0_r; exact H. - left; rewrite b; rewrite Rplus_0_r; apply Rlt_0_1. + left; rewrite Heq; rewrite Rplus_0_r; apply Rlt_0_1. left; apply Rplus_lt_0_compat. apply Rlt_0_1. - apply r. + apply Hgt. Qed. (** sqrt is continuous in 1 *) @@ -203,8 +201,8 @@ Proof. left; apply Rlt_0_1. left; apply H. elim H6; intros. - case (Rcase_abs (x0 - x)); intro. - rewrite (Rabs_left (x0 - x) r) in H8. + destruct (Rcase_abs (x0 - x)) as [Hlt|Hgt]. + rewrite (Rabs_left (x0 - x) Hlt) in H8. rewrite Rplus_comm. apply Rplus_le_reg_l with (- ((x0 - x) / x)). rewrite Rplus_0_r; rewrite <- Rplus_assoc; rewrite Rplus_opp_l; @@ -220,7 +218,7 @@ Proof. apply Rplus_le_le_0_compat. left; apply Rlt_0_1. unfold Rdiv; apply Rmult_le_pos. - apply Rge_le; exact r. + apply Rge_le; exact Hgt. left; apply Rinv_0_lt_compat; apply H. unfold Rdiv; apply Rmult_lt_0_compat. apply H1. @@ -273,8 +271,8 @@ Proof. apply Rplus_lt_le_0_compat. apply sqrt_lt_R0; apply H. apply sqrt_positivity; apply H10. - case (Rcase_abs h); intro. - rewrite (Rabs_left h r) in H9. + destruct (Rcase_abs h) as [Hlt|Hgt]. + rewrite (Rabs_left h Hlt) in H9. apply Rplus_le_reg_l with (- h). rewrite Rplus_0_r; rewrite Rplus_comm; rewrite Rplus_assoc; rewrite Rplus_opp_r; rewrite Rplus_0_r; left; apply Rlt_le_trans with alpha1. @@ -282,7 +280,7 @@ Proof. unfold alpha1; apply Rmin_r. apply Rplus_le_le_0_compat. left; assumption. - apply Rge_le; apply r. + apply Rge_le; apply Hgt. unfold alpha1; unfold Rmin; case (Rle_dec alpha x); intro. apply H5. apply H. @@ -341,17 +339,16 @@ Proof. rewrite <- H1; rewrite sqrt_0; unfold Rminus; rewrite Ropp_0; rewrite Rplus_0_r; rewrite <- H1 in H5; unfold Rminus in H5; rewrite Ropp_0 in H5; rewrite Rplus_0_r in H5. - case (Rcase_abs x0); intro. - unfold sqrt; case (Rcase_abs x0); intro. + destruct (Rcase_abs x0) as [Hlt|Hgt]_eqn:Heqs. + unfold sqrt. rewrite Heqs. rewrite Rabs_R0; apply H2. - assert (H6 := Rge_le _ _ r0); elim (Rlt_irrefl _ (Rle_lt_trans _ _ _ H6 r)). rewrite Rabs_right. apply Rsqr_incrst_0. rewrite Rsqr_sqrt. - rewrite (Rabs_right x0 r) in H5; apply H5. - apply Rge_le; exact r. - apply sqrt_positivity; apply Rge_le; exact r. + rewrite (Rabs_right x0 Hgt) in H5; apply H5. + apply Rge_le; exact Hgt. + apply sqrt_positivity; apply Rge_le; exact Hgt. left; exact H2. - apply Rle_ge; apply sqrt_positivity; apply Rge_le; exact r. + apply Rle_ge; apply sqrt_positivity; apply Rge_le; exact Hgt. elim (Rlt_irrefl _ (Rlt_le_trans _ _ _ H1 H)). Qed. diff --git a/theories/Reals/vo.itarget b/theories/Reals/vo.itarget index 36dd0f56..0c8f0b97 100644 --- a/theories/Reals/vo.itarget +++ b/theories/Reals/vo.itarget @@ -8,7 +8,6 @@ Cos_rel.vo DiscrR.vo Exp_prop.vo Integration.vo -LegacyRfield.vo Machin.vo MVT.vo NewtonInt.vo |