From 9058fb97426307536f56c3e7447be2f70798e081 Mon Sep 17 00:00:00 2001 From: herbelin Date: Sat, 29 Nov 2003 16:15:58 +0000 Subject: Deplacement des fichiers ancienne syntaxe dans theories7, contrib7 et states7 git-svn-id: svn+ssh://scm.gforge.inria.fr/svn/coq/trunk@5026 85f007b7-540e-0410-9357-904b9bb8a0f7 --- theories7/Reals/Alembert.v | 549 +++++++++++++++++++++++++++++++++++++++++++++ 1 file changed, 549 insertions(+) create mode 100644 theories7/Reals/Alembert.v (limited to 'theories7/Reals/Alembert.v') diff --git a/theories7/Reals/Alembert.v b/theories7/Reals/Alembert.v new file mode 100644 index 000000000..455803aa1 --- /dev/null +++ b/theories7/Reals/Alembert.v @@ -0,0 +1,549 @@ +(***********************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* R) ((n:nat)``0<(An n)``) -> (Un_cv [n:nat](Rabsolu ``(An (S n))/(An n)``) R0) -> (SigT R [l:R](Un_cv [N:nat](sum_f_R0 An N) l)). +Intros An H H0. +Cut (sigTT R [l:R](is_lub (EUn [N:nat](sum_f_R0 An N)) l)) -> (SigT R [l:R](Un_cv [N:nat](sum_f_R0 An N) l)). +Intro; Apply X. +Apply complet. +Unfold Un_cv in H0; Unfold bound; Cut ``0``(An (S n))R) ((n:nat)``(An n)<>0``) -> (Un_cv [n:nat](Rabsolu ``(An (S n))/(An n)``) R0) -> (SigT R [l:R](Un_cv [N:nat](sum_f_R0 An N) l)). +Intros. +Pose Vn := [i:nat]``(2*(Rabsolu (An i))+(An i))/2``. +Pose Wn := [i:nat]``(2*(Rabsolu (An i))-(An i))/2``. +Cut (n:nat)``0<(Vn n)``. +Intro; Cut (n:nat)``0<(Wn n)``. +Intro; Cut (Un_cv [n:nat](Rabsolu ``(Vn (S n))/(Vn n)``) ``0``). +Intro; Cut (Un_cv [n:nat](Rabsolu ``(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. +Apply Specif.existT with ``x-x0``; Unfold Un_cv; Unfold Un_cv in p; Unfold Un_cv in p0; Intros; Cut ``0R;x:R) ``x<>0`` -> ((n:nat)``(An n)<>0``) -> (Un_cv [n:nat](Rabsolu ``(An (S n))/(An n)``) ``0``) -> (SigT R [l:R](Pser An x l)). +Intros; Pose Bn := [i:nat]``(An i)*(pow x i)``. +Cut (n:nat)``(Bn n)<>0``. +Intro; Cut (Un_cv [n:nat](Rabsolu ``(Bn (S n))/(Bn n)``) ``0``). +Intro; Assert H4 := (Alembert_C2 Bn H2 H3). +Elim H4; Intros. +Apply Specif.existT with x0; Unfold Bn in p; Apply tech12; Assumption. +Unfold Un_cv; Intros; Unfold Un_cv in H1; Cut ``0R;x:R) ``x==0`` -> (SigT R [l:R](Pser An x l)). +Intros; Apply Specif.existT with (An O). +Unfold Pser; Unfold infinit_sum; Intros; Exists O; Intros; Replace (sum_f_R0 [n0:nat]``(An n0)*(pow x n0)`` n) with (An O). +Unfold R_dist; Unfold Rminus; Rewrite Rplus_Ropp_r; Rewrite Rabsolu_R0; Assumption. +Induction n. +Simpl; Ring. +Rewrite tech5; Rewrite Hrecn; [Rewrite H; Simpl; Ring | Unfold ge; Apply le_O_n]. +Qed. + +(* An useful criterion of convergence for power series *) +Theorem Alembert_C3 : (An:nat->R;x:R) ((n:nat)``(An n)<>0``) -> (Un_cv [n:nat](Rabsolu ``(An (S n))/(An n)``) ``0``) -> (SigT R [l:R](Pser An x l)). +Intros; Case (total_order_T x R0); Intro. +Elim s; Intro. +Cut ``x<>0``. +Intro; Apply AlembertC3_step1; Assumption. +Red; Intro; Rewrite H1 in a; Elim (Rlt_antirefl ? a). +Apply AlembertC3_step2; Assumption. +Cut ``x<>0``. +Intro; Apply AlembertC3_step1; Assumption. +Red; Intro; Rewrite H1 in r; Elim (Rlt_antirefl ? r). +Qed. + +Lemma Alembert_C4 : (An:nat->R;k:R) ``0<=k<1`` -> ((n:nat)``0<(An n)``) -> (Un_cv [n:nat](Rabsolu ``(An (S n))/(An n)``) k) -> (SigT R [l:R](Un_cv [N:nat](sum_f_R0 An N) l)). +Intros An k Hyp H H0. +Cut (sigTT R [l:R](is_lub (EUn [N:nat](sum_f_R0 An N)) l)) -> (SigT R [l:R](Un_cv [N:nat](sum_f_R0 An N) l)). +Intro; Apply X. +Apply complet. +Assert H1 := (tech13 ? ? Hyp H0). +Elim H1; Intros. +Elim H2; Intros. +Elim H4; Intros. +Unfold bound; Exists ``(sum_f_R0 An x0)+/(1-x)*(An (S x0))``. +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. +Replace (sum_f_R0 An x0) with (Rplus (sum_f_R0 An x2) (sum_f_R0 [i:nat](An (plus (S x2) i)) (minus x0 (S x2)))). +Pattern 1 (sum_f_R0 An x2); Rewrite <- Rplus_Or. +Rewrite Rplus_assoc; Apply Rle_compatibility. +Left; Apply gt0_plus_gt0_is_gt0. +Apply tech1. +Intros; Apply H. +Apply Rmult_lt_pos. +Apply Rlt_Rinv; Apply Rlt_anti_compatibility with x; Rewrite Rplus_Or; Replace ``x+(1-x)`` with R1; [Elim H3; Intros; Assumption | Ring]. +Apply H. +Symmetry; Apply tech2; Assumption. +Rewrite b; Pattern 1 (sum_f_R0 An x0); Rewrite <- Rplus_Or; Apply Rle_compatibility. +Left; Apply Rmult_lt_pos. +Apply Rlt_Rinv; Apply Rlt_anti_compatibility with x; Rewrite Rplus_Or; Replace ``x+(1-x)`` with R1; [Elim H3; Intros; Assumption | Ring]. +Apply H. +Replace (sum_f_R0 An x2) with (Rplus (sum_f_R0 An x0) (sum_f_R0 [i:nat](An (plus (S x0) i)) (minus x2 (S x0)))). +Apply Rle_compatibility. +Cut (Rle (sum_f_R0 [i:nat](An (plus (S x0) i)) (minus x2 (S x0))) (Rmult (An (S x0)) (sum_f_R0 [i:nat](pow x i) (minus x2 (S x0))))). +Intro; Apply Rle_trans with (Rmult (An (S x0)) (sum_f_R0 [i:nat](pow x i) (minus x2 (S x0)))). +Assumption. +Rewrite <- (Rmult_sym (An (S x0))); Apply Rle_monotony. +Left; Apply H. +Rewrite tech3. +Unfold Rdiv; Apply Rle_monotony_contra with ``1-x``. +Apply Rlt_anti_compatibility with x; Rewrite Rplus_Or. +Replace ``x+(1-x)`` with R1; [Elim H3; Intros; Assumption | Ring]. +Do 2 Rewrite (Rmult_sym ``1-x``). +Rewrite Rmult_assoc; Rewrite <- Rinv_l_sym. +Rewrite Rmult_1r; Apply Rle_anti_compatibility with ``(pow x (S (minus x2 (S x0))))``. +Replace ``(pow x (S (minus x2 (S x0))))+(1-(pow x (S (minus x2 (S x0)))))`` with R1; [Idtac | Ring]. +Rewrite <- (Rplus_sym R1); Pattern 1 R1; Rewrite <- Rplus_Or; Apply Rle_compatibility. +Left; Apply pow_lt. +Apply Rle_lt_trans with k. +Elim Hyp; Intros; Assumption. +Elim H3; Intros; Assumption. +Apply Rminus_eq_contra. +Red; Intro. +Elim H3; Intros. +Rewrite H10 in H12; Elim (Rlt_antirefl ? H12). +Red; Intro. +Elim H3; Intros. +Rewrite H10 in H12; Elim (Rlt_antirefl ? H12). +Replace (An (S x0)) with (An (plus (S x0) O)). +Apply (tech6 [i:nat](An (plus (S x0) i)) x). +Left; Apply Rle_lt_trans with k. +Elim Hyp; Intros; Assumption. +Elim H3; Intros; Assumption. +Intro. +Cut (n:nat)(ge n x0)->``(An (S n))R;k:R) ``0<=k<1`` -> ((n:nat)``(An n)<>0``) -> (Un_cv [n:nat](Rabsolu ``(An (S n))/(An n)``) k) -> (SigT R [l:R](Un_cv [N:nat](sum_f_R0 An N) l)). +Intros. +Cut (sigTT R [l:R](Un_cv [N:nat](sum_f_R0 An N) l)) -> (SigT R [l:R](Un_cv [N:nat](sum_f_R0 An N) l)). +Intro Hyp0; Apply Hyp0. +Apply cv_cauchy_2. +Apply cauchy_abs. +Apply cv_cauchy_1. +Cut (SigT R [l:R](Un_cv [N:nat](sum_f_R0 [i:nat](Rabsolu (An i)) N) l)) -> (sigTT R [l:R](Un_cv [N:nat](sum_f_R0 [i:nat](Rabsolu (An i)) N) l)). +Intro Hyp; Apply Hyp. +Apply (Alembert_C4 [i:nat](Rabsolu (An i)) k). +Assumption. +Intro; Apply Rabsolu_pos_lt; Apply H0. +Unfold Un_cv. +Unfold Un_cv in H1. +Unfold Rdiv. +Intros. +Elim (H1 eps H2); Intros. +Exists x; Intros. +Rewrite <- Rabsolu_Rinv. +Rewrite <- Rabsolu_mult. +Rewrite Rabsolu_Rabsolu. +Unfold Rdiv in H3; Apply H3; Assumption. +Apply H0. +Intro. +Elim X; Intros. +Apply existTT with x. +Assumption. +Intro. +Elim X; Intros. +Apply Specif.existT with x. +Assumption. +Qed. + +(* Convergence of power series in D(O,1/k) *) +(* k=0 is described in Alembert_C3 *) +Lemma Alembert_C6 : (An:nat->R;x,k:R) ``0 ((n:nat)``(An n)<>0``) -> (Un_cv [n:nat](Rabsolu ``(An (S n))/(An n)``) k) -> ``(Rabsolu x) (SigT R [l:R](Pser An x l)). +Intros. +Cut (SigT R [l:R](Un_cv [N:nat](sum_f_R0 [i:nat]``(An i)*(pow x i)`` N) l)). +Intro. +Elim X; Intros. +Apply Specif.existT with x0. +Apply tech12; Assumption. +Case (total_order_T x R0); Intro. +Elim s; Intro. +EApply Alembert_C5 with ``k*(Rabsolu x)``. +Split. +Unfold Rdiv; Apply Rmult_le_pos. +Left; Assumption. +Left; Apply Rabsolu_pos_lt. +Red; Intro; Rewrite H3 in a; Elim (Rlt_antirefl ? a). +Apply Rlt_monotony_contra with ``/k``. +Apply Rlt_Rinv; Assumption. +Rewrite <- Rmult_assoc. +Rewrite <- Rinv_l_sym. +Rewrite Rmult_1l. +Rewrite Rmult_1r; Assumption. +Red; Intro; Rewrite H3 in H; Elim (Rlt_antirefl ? H). +Intro; Apply prod_neq_R0. +Apply H0. +Apply pow_nonzero. +Red; Intro; Rewrite H3 in a; Elim (Rlt_antirefl ? a). +Unfold Un_cv; Unfold Un_cv in H1. +Intros. +Cut ``0