![]() |
Intuitionistic Logic Explorer |
< Previous
Next >
Nearby theorems |
|
Mirrors > Home > ILE Home > Th. List > cvgratnn | GIF version |
Description: Ratio test for convergence of a complex infinite series. If the ratio 𝐴 of the absolute values of successive terms in an infinite sequence 𝐹 is less than 1 for all terms, then the infinite sum of the terms of 𝐹 converges to a complex number. Although this theorem is similar to cvgratz 11678 and cvgratgt0 11679, the decision to index starting at one is not merely cosmetic, as proving convergence using climcvg1n 11496 is sensitive to how a sequence is indexed. (Contributed by NM, 26-Apr-2005.) (Revised by Jim Kingdon, 12-Nov-2022.) |
Ref | Expression |
---|---|
cvgratnn.3 | ⊢ (𝜑 → 𝐴 ∈ ℝ) |
cvgratnn.4 | ⊢ (𝜑 → 𝐴 < 1) |
cvgratnn.gt0 | ⊢ (𝜑 → 0 < 𝐴) |
cvgratnn.6 | ⊢ ((𝜑 ∧ 𝑘 ∈ ℕ) → (𝐹‘𝑘) ∈ ℂ) |
cvgratnn.7 | ⊢ ((𝜑 ∧ 𝑘 ∈ ℕ) → (abs‘(𝐹‘(𝑘 + 1))) ≤ (𝐴 · (abs‘(𝐹‘𝑘)))) |
Ref | Expression |
---|---|
cvgratnn | ⊢ (𝜑 → seq1( + , 𝐹) ∈ dom ⇝ ) |
Step | Hyp | Ref | Expression |
---|---|---|---|
1 | nnuz 9631 | . . 3 ⊢ ℕ = (ℤ≥‘1) | |
2 | 1zzd 9347 | . . 3 ⊢ (𝜑 → 1 ∈ ℤ) | |
3 | cvgratnn.6 | . . 3 ⊢ ((𝜑 ∧ 𝑘 ∈ ℕ) → (𝐹‘𝑘) ∈ ℂ) | |
4 | 1, 2, 3 | serf 10557 | . 2 ⊢ (𝜑 → seq1( + , 𝐹):ℕ⟶ℂ) |
5 | cvgratnn.3 | . . . . . . . . . 10 ⊢ (𝜑 → 𝐴 ∈ ℝ) | |
6 | cvgratnn.gt0 | . . . . . . . . . 10 ⊢ (𝜑 → 0 < 𝐴) | |
7 | 5, 6 | elrpd 9762 | . . . . . . . . 9 ⊢ (𝜑 → 𝐴 ∈ ℝ+) |
8 | 7 | rprecred 9777 | . . . . . . . 8 ⊢ (𝜑 → (1 / 𝐴) ∈ ℝ) |
9 | 1red 8036 | . . . . . . . 8 ⊢ (𝜑 → 1 ∈ ℝ) | |
10 | 8, 9 | resubcld 8402 | . . . . . . 7 ⊢ (𝜑 → ((1 / 𝐴) − 1) ∈ ℝ) |
11 | cvgratnn.4 | . . . . . . . . 9 ⊢ (𝜑 → 𝐴 < 1) | |
12 | 7 | reclt1d 9779 | . . . . . . . . 9 ⊢ (𝜑 → (𝐴 < 1 ↔ 1 < (1 / 𝐴))) |
13 | 11, 12 | mpbid 147 | . . . . . . . 8 ⊢ (𝜑 → 1 < (1 / 𝐴)) |
14 | 9, 8 | posdifd 8553 | . . . . . . . 8 ⊢ (𝜑 → (1 < (1 / 𝐴) ↔ 0 < ((1 / 𝐴) − 1))) |
15 | 13, 14 | mpbid 147 | . . . . . . 7 ⊢ (𝜑 → 0 < ((1 / 𝐴) − 1)) |
16 | 10, 15 | elrpd 9762 | . . . . . 6 ⊢ (𝜑 → ((1 / 𝐴) − 1) ∈ ℝ+) |
17 | 16 | rpreccld 9776 | . . . . 5 ⊢ (𝜑 → (1 / ((1 / 𝐴) − 1)) ∈ ℝ+) |
18 | 17, 7 | rpdivcld 9783 | . . . 4 ⊢ (𝜑 → ((1 / ((1 / 𝐴) − 1)) / 𝐴) ∈ ℝ+) |
19 | fveq2 5555 | . . . . . . . 8 ⊢ (𝑘 = 1 → (𝐹‘𝑘) = (𝐹‘1)) | |
20 | 19 | eleq1d 2262 | . . . . . . 7 ⊢ (𝑘 = 1 → ((𝐹‘𝑘) ∈ ℂ ↔ (𝐹‘1) ∈ ℂ)) |
21 | 3 | ralrimiva 2567 | . . . . . . 7 ⊢ (𝜑 → ∀𝑘 ∈ ℕ (𝐹‘𝑘) ∈ ℂ) |
22 | 1nn 8995 | . . . . . . . 8 ⊢ 1 ∈ ℕ | |
23 | 22 | a1i 9 | . . . . . . 7 ⊢ (𝜑 → 1 ∈ ℕ) |
24 | 20, 21, 23 | rspcdva 2870 | . . . . . 6 ⊢ (𝜑 → (𝐹‘1) ∈ ℂ) |
25 | 24 | abscld 11328 | . . . . 5 ⊢ (𝜑 → (abs‘(𝐹‘1)) ∈ ℝ) |
26 | 24 | absge0d 11331 | . . . . 5 ⊢ (𝜑 → 0 ≤ (abs‘(𝐹‘1))) |
27 | 25, 26 | ge0p1rpd 9796 | . . . 4 ⊢ (𝜑 → ((abs‘(𝐹‘1)) + 1) ∈ ℝ+) |
28 | 18, 27 | rpmulcld 9782 | . . 3 ⊢ (𝜑 → (((1 / ((1 / 𝐴) − 1)) / 𝐴) · ((abs‘(𝐹‘1)) + 1)) ∈ ℝ+) |
29 | 9, 5 | resubcld 8402 | . . . . 5 ⊢ (𝜑 → (1 − 𝐴) ∈ ℝ) |
30 | 5, 9 | posdifd 8553 | . . . . . 6 ⊢ (𝜑 → (𝐴 < 1 ↔ 0 < (1 − 𝐴))) |
31 | 11, 30 | mpbid 147 | . . . . 5 ⊢ (𝜑 → 0 < (1 − 𝐴)) |
32 | 29, 31 | elrpd 9762 | . . . 4 ⊢ (𝜑 → (1 − 𝐴) ∈ ℝ+) |
33 | 7, 32 | rpdivcld 9783 | . . 3 ⊢ (𝜑 → (𝐴 / (1 − 𝐴)) ∈ ℝ+) |
34 | 28, 33 | rpmulcld 9782 | . 2 ⊢ (𝜑 → ((((1 / ((1 / 𝐴) − 1)) / 𝐴) · ((abs‘(𝐹‘1)) + 1)) · (𝐴 / (1 − 𝐴))) ∈ ℝ+) |
35 | 5 | adantr 276 | . . . 4 ⊢ ((𝜑 ∧ (𝑚 ∈ ℕ ∧ 𝑛 ∈ (ℤ≥‘𝑚))) → 𝐴 ∈ ℝ) |
36 | 11 | adantr 276 | . . . 4 ⊢ ((𝜑 ∧ (𝑚 ∈ ℕ ∧ 𝑛 ∈ (ℤ≥‘𝑚))) → 𝐴 < 1) |
37 | 6 | adantr 276 | . . . 4 ⊢ ((𝜑 ∧ (𝑚 ∈ ℕ ∧ 𝑛 ∈ (ℤ≥‘𝑚))) → 0 < 𝐴) |
38 | 3 | adantlr 477 | . . . 4 ⊢ (((𝜑 ∧ (𝑚 ∈ ℕ ∧ 𝑛 ∈ (ℤ≥‘𝑚))) ∧ 𝑘 ∈ ℕ) → (𝐹‘𝑘) ∈ ℂ) |
39 | cvgratnn.7 | . . . . 5 ⊢ ((𝜑 ∧ 𝑘 ∈ ℕ) → (abs‘(𝐹‘(𝑘 + 1))) ≤ (𝐴 · (abs‘(𝐹‘𝑘)))) | |
40 | 39 | adantlr 477 | . . . 4 ⊢ (((𝜑 ∧ (𝑚 ∈ ℕ ∧ 𝑛 ∈ (ℤ≥‘𝑚))) ∧ 𝑘 ∈ ℕ) → (abs‘(𝐹‘(𝑘 + 1))) ≤ (𝐴 · (abs‘(𝐹‘𝑘)))) |
41 | simprl 529 | . . . 4 ⊢ ((𝜑 ∧ (𝑚 ∈ ℕ ∧ 𝑛 ∈ (ℤ≥‘𝑚))) → 𝑚 ∈ ℕ) | |
42 | simprr 531 | . . . 4 ⊢ ((𝜑 ∧ (𝑚 ∈ ℕ ∧ 𝑛 ∈ (ℤ≥‘𝑚))) → 𝑛 ∈ (ℤ≥‘𝑚)) | |
43 | 35, 36, 37, 38, 40, 41, 42 | cvgratnnlemrate 11676 | . . 3 ⊢ ((𝜑 ∧ (𝑚 ∈ ℕ ∧ 𝑛 ∈ (ℤ≥‘𝑚))) → (abs‘((seq1( + , 𝐹)‘𝑛) − (seq1( + , 𝐹)‘𝑚))) < (((((1 / ((1 / 𝐴) − 1)) / 𝐴) · ((abs‘(𝐹‘1)) + 1)) · (𝐴 / (1 − 𝐴))) / 𝑚)) |
44 | 43 | ralrimivva 2576 | . 2 ⊢ (𝜑 → ∀𝑚 ∈ ℕ ∀𝑛 ∈ (ℤ≥‘𝑚)(abs‘((seq1( + , 𝐹)‘𝑛) − (seq1( + , 𝐹)‘𝑚))) < (((((1 / ((1 / 𝐴) − 1)) / 𝐴) · ((abs‘(𝐹‘1)) + 1)) · (𝐴 / (1 − 𝐴))) / 𝑚)) |
45 | 4, 34, 44 | climcvg1n 11496 | 1 ⊢ (𝜑 → seq1( + , 𝐹) ∈ dom ⇝ ) |
Colors of variables: wff set class |
Syntax hints: → wi 4 ∧ wa 104 = wceq 1364 ∈ wcel 2164 class class class wbr 4030 dom cdm 4660 ‘cfv 5255 (class class class)co 5919 ℂcc 7872 ℝcr 7873 0cc0 7874 1c1 7875 + caddc 7877 · cmul 7879 < clt 8056 ≤ cle 8057 − cmin 8192 / cdiv 8693 ℕcn 8984 ℤ≥cuz 9595 seqcseq 10521 abscabs 11144 ⇝ cli 11424 |
This theorem was proved from axioms: ax-mp 5 ax-1 6 ax-2 7 ax-ia1 106 ax-ia2 107 ax-ia3 108 ax-in1 615 ax-in2 616 ax-io 710 ax-5 1458 ax-7 1459 ax-gen 1460 ax-ie1 1504 ax-ie2 1505 ax-8 1515 ax-10 1516 ax-11 1517 ax-i12 1518 ax-bndl 1520 ax-4 1521 ax-17 1537 ax-i9 1541 ax-ial 1545 ax-i5r 1546 ax-13 2166 ax-14 2167 ax-ext 2175 ax-coll 4145 ax-sep 4148 ax-nul 4156 ax-pow 4204 ax-pr 4239 ax-un 4465 ax-setind 4570 ax-iinf 4621 ax-cnex 7965 ax-resscn 7966 ax-1cn 7967 ax-1re 7968 ax-icn 7969 ax-addcl 7970 ax-addrcl 7971 ax-mulcl 7972 ax-mulrcl 7973 ax-addcom 7974 ax-mulcom 7975 ax-addass 7976 ax-mulass 7977 ax-distr 7978 ax-i2m1 7979 ax-0lt1 7980 ax-1rid 7981 ax-0id 7982 ax-rnegex 7983 ax-precex 7984 ax-cnre 7985 ax-pre-ltirr 7986 ax-pre-ltwlin 7987 ax-pre-lttrn 7988 ax-pre-apti 7989 ax-pre-ltadd 7990 ax-pre-mulgt0 7991 ax-pre-mulext 7992 ax-arch 7993 ax-caucvg 7994 |
This theorem depends on definitions: df-bi 117 df-dc 836 df-3or 981 df-3an 982 df-tru 1367 df-fal 1370 df-nf 1472 df-sb 1774 df-eu 2045 df-mo 2046 df-clab 2180 df-cleq 2186 df-clel 2189 df-nfc 2325 df-ne 2365 df-nel 2460 df-ral 2477 df-rex 2478 df-reu 2479 df-rmo 2480 df-rab 2481 df-v 2762 df-sbc 2987 df-csb 3082 df-dif 3156 df-un 3158 df-in 3160 df-ss 3167 df-nul 3448 df-if 3559 df-pw 3604 df-sn 3625 df-pr 3626 df-op 3628 df-uni 3837 df-int 3872 df-iun 3915 df-br 4031 df-opab 4092 df-mpt 4093 df-tr 4129 df-id 4325 df-po 4328 df-iso 4329 df-iord 4398 df-on 4400 df-ilim 4401 df-suc 4403 df-iom 4624 df-xp 4666 df-rel 4667 df-cnv 4668 df-co 4669 df-dm 4670 df-rn 4671 df-res 4672 df-ima 4673 df-iota 5216 df-fun 5257 df-fn 5258 df-f 5259 df-f1 5260 df-fo 5261 df-f1o 5262 df-fv 5263 df-isom 5264 df-riota 5874 df-ov 5922 df-oprab 5923 df-mpo 5924 df-1st 6195 df-2nd 6196 df-recs 6360 df-irdg 6425 df-frec 6446 df-1o 6471 df-oadd 6475 df-er 6589 df-en 6797 df-dom 6798 df-fin 6799 df-pnf 8058 df-mnf 8059 df-xr 8060 df-ltxr 8061 df-le 8062 df-sub 8194 df-neg 8195 df-reap 8596 df-ap 8603 df-div 8694 df-inn 8985 df-2 9043 df-3 9044 df-4 9045 df-n0 9244 df-z 9321 df-uz 9596 df-q 9688 df-rp 9723 df-ico 9963 df-fz 10078 df-fzo 10212 df-seqfrec 10522 df-exp 10613 df-ihash 10850 df-cj 10989 df-re 10990 df-im 10991 df-rsqrt 11145 df-abs 11146 df-clim 11425 df-sumdc 11500 |
This theorem is referenced by: cvgratz 11678 |
Copyright terms: Public domain | W3C validator |