Arithmetic-geometric sequences #
An arithmetic-geometric sequence is a sequence defined by the recurrence relation
u (n + 1) = a * u n + b
.
Main definitions #
arithGeom a b u₀
: arithmetic-geometric sequence with starting valueu₀
and recurrence relationu (n + 1) = a * u n + b
.
Main statements #
arithGeom_eq
: fora ≠ 1
,arithGeom a b u₀ n = a ^ n * (u₀ - (b / (1 - a))) + b / (1 - a)
tendsto_arithGeom_atTop_of_one_lt
: if1 < a
andb / (1 - a) < u₀
, thenarithGeom a b u₀ n
tends to+∞
asn
tends to+∞
.tendsto_arithGeom_nhds_of_lt_one
: if0 ≤ a < 1
, thenarithGeom a b u₀ n
tends tob / (1 - a)
asn
tends to+∞
.arithGeom_strictMono
: if1 < a
andb / (1 - a) < u₀
, thenarithGeom a b u₀
is strictly monotone.
theorem
arithGeom_strictMono
{R : Type u_1}
{a b u₀ : R}
[Field R]
[LinearOrder R]
[IsStrictOrderedRing R]
(ha : 1 < a)
(h0 : b / (1 - a) < u₀)
:
StrictMono (arithGeom a b u₀)
theorem
tendsto_arithGeom_atTop_of_one_lt
{R : Type u_1}
{a b u₀ : R}
[Field R]
[LinearOrder R]
[IsStrictOrderedRing R]
[Archimedean R]
(ha : 1 < a)
(h0 : b / (1 - a) < u₀)
:
Filter.Tendsto (arithGeom a b u₀) Filter.atTop Filter.atTop
theorem
tendsto_arithGeom_nhds_of_lt_one
{R : Type u_1}
{a b u₀ : R}
[Field R]
[LinearOrder R]
[IsStrictOrderedRing R]
[Archimedean R]
[TopologicalSpace R]
[OrderTopology R]
(ha_pos : 0 ≤ a)
(ha : a < 1)
:
Filter.Tendsto (arithGeom a b u₀) Filter.atTop (nhds (b / (1 - a)))