Theory SEQ

Up to index of Isabelle/HOL/HOL-Complex

theory SEQ
imports Real
begin

(*  Title       : SEQ.thy
    Author      : Jacques D. Fleuriot
    Copyright   : 1998  University of Cambridge
    Description : Convergence of sequences and series
    Conversion to Isar and new proofs by Lawrence C Paulson, 2004
    Additional contributions by Jeremy Avigad and Brian Huffman
*)

header {* Sequences and Convergence *}

theory SEQ
imports "../Real/Real"
begin

definition
  Zseq :: "[nat => 'a::real_normed_vector] => bool" where
    --{*Standard definition of sequence converging to zero*}
  "Zseq X = (∀r>0. ∃no. ∀n≥no. norm (X n) < r)"

definition
  LIMSEQ :: "[nat => 'a::real_normed_vector, 'a] => bool"
    ("((_)/ ----> (_))" [60, 60] 60) where
    --{*Standard definition of convergence of sequence*}
  "X ----> L = (∀r. 0 < r --> (∃no. ∀n. no ≤ n --> norm (X n - L) < r))"

definition
  lim :: "(nat => 'a::real_normed_vector) => 'a" where
    --{*Standard definition of limit using choice operator*}
  "lim X = (THE L. X ----> L)"

definition
  convergent :: "(nat => 'a::real_normed_vector) => bool" where
    --{*Standard definition of convergence*}
  "convergent X = (∃L. X ----> L)"

definition
  Bseq :: "(nat => 'a::real_normed_vector) => bool" where
    --{*Standard definition for bounded sequence*}
  "Bseq X = (∃K>0.∀n. norm (X n) ≤ K)"

definition
  monoseq :: "(nat=>real)=>bool" where
    --{*Definition for monotonicity*}
  "monoseq X = ((∀m. ∀n≥m. X m ≤ X n) | (∀m. ∀n≥m. X n ≤ X m))"

definition
  subseq :: "(nat => nat) => bool" where
    --{*Definition of subsequence*}
  "subseq f = (∀m. ∀n>m. (f m) < (f n))"

definition
  Cauchy :: "(nat => 'a::real_normed_vector) => bool" where
    --{*Standard definition of the Cauchy condition*}
  "Cauchy X = (∀e>0. ∃M. ∀m ≥ M. ∀n ≥ M. norm (X m - X n) < e)"


subsection {* Bounded Sequences *}

lemma BseqI': assumes K: "!!n. norm (X n) ≤ K" shows "Bseq X"
unfolding Bseq_def
proof (intro exI conjI allI)
  show "0 < max K 1" by simp
next
  fix n::nat
  have "norm (X n) ≤ K" by (rule K)
  thus "norm (X n) ≤ max K 1" by simp
qed

lemma BseqE: "[|Bseq X; !!K. [|0 < K; ∀n. norm (X n) ≤ K|] ==> Q|] ==> Q"
unfolding Bseq_def by auto

lemma BseqI2': assumes K: "∀n≥N. norm (X n) ≤ K" shows "Bseq X"
proof (rule BseqI')
  let ?A = "norm ` X ` {..N}"
  have 1: "finite ?A" by simp
  fix n::nat
  show "norm (X n) ≤ max K (Max ?A)"
  proof (cases rule: linorder_le_cases)
    assume "n ≥ N"
    hence "norm (X n) ≤ K" using K by simp
    thus "norm (X n) ≤ max K (Max ?A)" by simp
  next
    assume "n ≤ N"
    hence "norm (X n) ∈ ?A" by simp
    with 1 have "norm (X n) ≤ Max ?A" by (rule Max_ge)
    thus "norm (X n) ≤ max K (Max ?A)" by simp
  qed
qed

lemma Bseq_ignore_initial_segment: "Bseq X ==> Bseq (λn. X (n + k))"
unfolding Bseq_def by auto

lemma Bseq_offset: "Bseq (λn. X (n + k)) ==> Bseq X"
apply (erule BseqE)
apply (rule_tac N="k" and K="K" in BseqI2')
apply clarify
apply (drule_tac x="n - k" in spec, simp)
done


subsection {* Sequences That Converge to Zero *}

lemma ZseqI:
  "(!!r. 0 < r ==> ∃no. ∀n≥no. norm (X n) < r) ==> Zseq X"
unfolding Zseq_def by simp

lemma ZseqD:
  "[|Zseq X; 0 < r|] ==> ∃no. ∀n≥no. norm (X n) < r"
unfolding Zseq_def by simp

lemma Zseq_zero: "Zseq (λn. 0)"
unfolding Zseq_def by simp

lemma Zseq_const_iff: "Zseq (λn. k) = (k = 0)"
unfolding Zseq_def by force

lemma Zseq_norm_iff: "Zseq (λn. norm (X n)) = Zseq (λn. X n)"
unfolding Zseq_def by simp

lemma Zseq_imp_Zseq:
  assumes X: "Zseq X"
  assumes Y: "!!n. norm (Y n) ≤ norm (X n) * K"
  shows "Zseq (λn. Y n)"
proof (cases)
  assume K: "0 < K"
  show ?thesis
  proof (rule ZseqI)
    fix r::real assume "0 < r"
    hence "0 < r / K"
      using K by (rule divide_pos_pos)
    then obtain N where "∀n≥N. norm (X n) < r / K"
      using ZseqD [OF X] by fast
    hence "∀n≥N. norm (X n) * K < r"
      by (simp add: pos_less_divide_eq K)
    hence "∀n≥N. norm (Y n) < r"
      by (simp add: order_le_less_trans [OF Y])
    thus "∃N. ∀n≥N. norm (Y n) < r" ..
  qed
next
  assume "¬ 0 < K"
  hence K: "K ≤ 0" by (simp only: linorder_not_less)
  {
    fix n::nat
    have "norm (Y n) ≤ norm (X n) * K" by (rule Y)
    also have "… ≤ norm (X n) * 0"
      using K norm_ge_zero by (rule mult_left_mono)
    finally have "norm (Y n) = 0" by simp
  }
  thus ?thesis by (simp add: Zseq_zero)
qed

lemma Zseq_le: "[|Zseq Y; ∀n. norm (X n) ≤ norm (Y n)|] ==> Zseq X"
by (erule_tac K="1" in Zseq_imp_Zseq, simp)

lemma Zseq_add:
  assumes X: "Zseq X"
  assumes Y: "Zseq Y"
  shows "Zseq (λn. X n + Y n)"
proof (rule ZseqI)
  fix r::real assume "0 < r"
  hence r: "0 < r / 2" by simp
  obtain M where M: "∀n≥M. norm (X n) < r/2"
    using ZseqD [OF X r] by fast
  obtain N where N: "∀n≥N. norm (Y n) < r/2"
    using ZseqD [OF Y r] by fast
  show "∃N. ∀n≥N. norm (X n + Y n) < r"
  proof (intro exI allI impI)
    fix n assume n: "max M N ≤ n"
    have "norm (X n + Y n) ≤ norm (X n) + norm (Y n)"
      by (rule norm_triangle_ineq)
    also have "… < r/2 + r/2"
    proof (rule add_strict_mono)
      from M n show "norm (X n) < r/2" by simp
      from N n show "norm (Y n) < r/2" by simp
    qed
    finally show "norm (X n + Y n) < r" by simp
  qed
qed

lemma Zseq_minus: "Zseq X ==> Zseq (λn. - X n)"
unfolding Zseq_def by simp

lemma Zseq_diff: "[|Zseq X; Zseq Y|] ==> Zseq (λn. X n - Y n)"
by (simp only: diff_minus Zseq_add Zseq_minus)

lemma (in bounded_linear) Zseq:
  assumes X: "Zseq X"
  shows "Zseq (λn. f (X n))"
proof -
  obtain K where "!!x. norm (f x) ≤ norm x * K"
    using bounded by fast
  with X show ?thesis
    by (rule Zseq_imp_Zseq)
qed

lemma (in bounded_bilinear) Zseq:
  assumes X: "Zseq X"
  assumes Y: "Zseq Y"
  shows "Zseq (λn. X n ** Y n)"
proof (rule ZseqI)
  fix r::real assume r: "0 < r"
  obtain K where K: "0 < K"
    and norm_le: "!!x y. norm (x ** y) ≤ norm x * norm y * K"
    using pos_bounded by fast
  from K have K': "0 < inverse K"
    by (rule positive_imp_inverse_positive)
  obtain M where M: "∀n≥M. norm (X n) < r"
    using ZseqD [OF X r] by fast
  obtain N where N: "∀n≥N. norm (Y n) < inverse K"
    using ZseqD [OF Y K'] by fast
  show "∃N. ∀n≥N. norm (X n ** Y n) < r"
  proof (intro exI allI impI)
    fix n assume n: "max M N ≤ n"
    have "norm (X n ** Y n) ≤ norm (X n) * norm (Y n) * K"
      by (rule norm_le)
    also have "norm (X n) * norm (Y n) * K < r * inverse K * K"
    proof (intro mult_strict_right_mono mult_strict_mono' norm_ge_zero K)
      from M n show Xn: "norm (X n) < r" by simp
      from N n show Yn: "norm (Y n) < inverse K" by simp
    qed
    also from K have "r * inverse K * K = r" by simp
    finally show "norm (X n ** Y n) < r" .
  qed
qed

lemma (in bounded_bilinear) Zseq_prod_Bseq:
  assumes X: "Zseq X"
  assumes Y: "Bseq Y"
  shows "Zseq (λn. X n ** Y n)"
proof -
  obtain K where K: "0 ≤ K"
    and norm_le: "!!x y. norm (x ** y) ≤ norm x * norm y * K"
    using nonneg_bounded by fast
  obtain B where B: "0 < B"
    and norm_Y: "!!n. norm (Y n) ≤ B"
    using Y [unfolded Bseq_def] by fast
  from X show ?thesis
  proof (rule Zseq_imp_Zseq)
    fix n::nat
    have "norm (X n ** Y n) ≤ norm (X n) * norm (Y n) * K"
      by (rule norm_le)
    also have "… ≤ norm (X n) * B * K"
      by (intro mult_mono' order_refl norm_Y norm_ge_zero
                mult_nonneg_nonneg K)
    also have "… = norm (X n) * (B * K)"
      by (rule mult_assoc)
    finally show "norm (X n ** Y n) ≤ norm (X n) * (B * K)" .
  qed
qed

lemma (in bounded_bilinear) Bseq_prod_Zseq:
  assumes X: "Bseq X"
  assumes Y: "Zseq Y"
  shows "Zseq (λn. X n ** Y n)"
proof -
  obtain K where K: "0 ≤ K"
    and norm_le: "!!x y. norm (x ** y) ≤ norm x * norm y * K"
    using nonneg_bounded by fast
  obtain B where B: "0 < B"
    and norm_X: "!!n. norm (X n) ≤ B"
    using X [unfolded Bseq_def] by fast
  from Y show ?thesis
  proof (rule Zseq_imp_Zseq)
    fix n::nat
    have "norm (X n ** Y n) ≤ norm (X n) * norm (Y n) * K"
      by (rule norm_le)
    also have "… ≤ B * norm (Y n) * K"
      by (intro mult_mono' order_refl norm_X norm_ge_zero
                mult_nonneg_nonneg K)
    also have "… = norm (Y n) * (B * K)"
      by (simp only: mult_ac)
    finally show "norm (X n ** Y n) ≤ norm (Y n) * (B * K)" .
  qed
qed

lemma (in bounded_bilinear) Zseq_left:
  "Zseq X ==> Zseq (λn. X n ** a)"
by (rule bounded_linear_left [THEN bounded_linear.Zseq])

lemma (in bounded_bilinear) Zseq_right:
  "Zseq X ==> Zseq (λn. a ** X n)"
by (rule bounded_linear_right [THEN bounded_linear.Zseq])

lemmas Zseq_mult = mult.Zseq
lemmas Zseq_mult_right = mult.Zseq_right
lemmas Zseq_mult_left = mult.Zseq_left


subsection {* Limits of Sequences *}

lemma LIMSEQ_iff:
      "(X ----> L) = (∀r>0. ∃no. ∀n ≥ no. norm (X n - L) < r)"
by (rule LIMSEQ_def)

lemma LIMSEQ_Zseq_iff: "((λn. X n) ----> L) = Zseq (λn. X n - L)"
by (simp only: LIMSEQ_def Zseq_def)

lemma LIMSEQ_I:
  "(!!r. 0 < r ==> ∃no. ∀n≥no. norm (X n - L) < r) ==> X ----> L"
by (simp add: LIMSEQ_def)

lemma LIMSEQ_D:
  "[|X ----> L; 0 < r|] ==> ∃no. ∀n≥no. norm (X n - L) < r"
by (simp add: LIMSEQ_def)

lemma LIMSEQ_const: "(λn. k) ----> k"
by (simp add: LIMSEQ_def)

lemma LIMSEQ_const_iff: "(λn. k) ----> l = (k = l)"
by (simp add: LIMSEQ_Zseq_iff Zseq_const_iff)

lemma LIMSEQ_norm: "X ----> a ==> (λn. norm (X n)) ----> norm a"
apply (simp add: LIMSEQ_def, safe)
apply (drule_tac x="r" in spec, safe)
apply (rule_tac x="no" in exI, safe)
apply (drule_tac x="n" in spec, safe)
apply (erule order_le_less_trans [OF norm_triangle_ineq3])
done

lemma LIMSEQ_ignore_initial_segment:
  "f ----> a ==> (λn. f (n + k)) ----> a"
apply (rule LIMSEQ_I)
apply (drule (1) LIMSEQ_D)
apply (erule exE, rename_tac N)
apply (rule_tac x=N in exI)
apply simp
done

lemma LIMSEQ_offset:
  "(λn. f (n + k)) ----> a ==> f ----> a"
apply (rule LIMSEQ_I)
apply (drule (1) LIMSEQ_D)
apply (erule exE, rename_tac N)
apply (rule_tac x="N + k" in exI)
apply clarify
apply (drule_tac x="n - k" in spec)
apply (simp add: le_diff_conv2)
done

lemma LIMSEQ_Suc: "f ----> l ==> (λn. f (Suc n)) ----> l"
by (drule_tac k="1" in LIMSEQ_ignore_initial_segment, simp)

lemma LIMSEQ_imp_Suc: "(λn. f (Suc n)) ----> l ==> f ----> l"
by (rule_tac k="1" in LIMSEQ_offset, simp)

lemma LIMSEQ_Suc_iff: "(λn. f (Suc n)) ----> l = f ----> l"
by (blast intro: LIMSEQ_imp_Suc LIMSEQ_Suc)

lemma add_diff_add:
  fixes a b c d :: "'a::ab_group_add"
  shows "(a + c) - (b + d) = (a - b) + (c - d)"
by simp

lemma minus_diff_minus:
  fixes a b :: "'a::ab_group_add"
  shows "(- a) - (- b) = - (a - b)"
by simp

lemma LIMSEQ_add: "[|X ----> a; Y ----> b|] ==> (λn. X n + Y n) ----> a + b"
by (simp only: LIMSEQ_Zseq_iff add_diff_add Zseq_add)

lemma LIMSEQ_minus: "X ----> a ==> (λn. - X n) ----> - a"
by (simp only: LIMSEQ_Zseq_iff minus_diff_minus Zseq_minus)

lemma LIMSEQ_minus_cancel: "(λn. - X n) ----> - a ==> X ----> a"
by (drule LIMSEQ_minus, simp)

lemma LIMSEQ_diff: "[|X ----> a; Y ----> b|] ==> (λn. X n - Y n) ----> a - b"
by (simp add: diff_minus LIMSEQ_add LIMSEQ_minus)

lemma LIMSEQ_unique: "[|X ----> a; X ----> b|] ==> a = b"
by (drule (1) LIMSEQ_diff, simp add: LIMSEQ_const_iff)

lemma (in bounded_linear) LIMSEQ:
  "X ----> a ==> (λn. f (X n)) ----> f a"
by (simp only: LIMSEQ_Zseq_iff diff [symmetric] Zseq)

lemma (in bounded_bilinear) LIMSEQ:
  "[|X ----> a; Y ----> b|] ==> (λn. X n ** Y n) ----> a ** b"
by (simp only: LIMSEQ_Zseq_iff prod_diff_prod
               Zseq_add Zseq Zseq_left Zseq_right)

lemma LIMSEQ_mult:
  fixes a b :: "'a::real_normed_algebra"
  shows "[| X ----> a; Y ----> b |] ==> (%n. X n * Y n) ----> a * b"
by (rule mult.LIMSEQ)

lemma inverse_diff_inverse:
  "[|(a::'a::division_ring) ≠ 0; b ≠ 0|]
   ==> inverse a - inverse b = - (inverse a * (a - b) * inverse b)"
by (simp add: ring_simps)

lemma Bseq_inverse_lemma:
  fixes x :: "'a::real_normed_div_algebra"
  shows "[|r ≤ norm x; 0 < r|] ==> norm (inverse x) ≤ inverse r"
apply (subst nonzero_norm_inverse, clarsimp)
apply (erule (1) le_imp_inverse_le)
done

lemma Bseq_inverse:
  fixes a :: "'a::real_normed_div_algebra"
  assumes X: "X ----> a"
  assumes a: "a ≠ 0"
  shows "Bseq (λn. inverse (X n))"
proof -
  from a have "0 < norm a" by simp
  hence "∃r>0. r < norm a" by (rule dense)
  then obtain r where r1: "0 < r" and r2: "r < norm a" by fast
  obtain N where N: "!!n. N ≤ n ==> norm (X n - a) < r"
    using LIMSEQ_D [OF X r1] by fast
  show ?thesis
  proof (rule BseqI2' [rule_format])
    fix n assume n: "N ≤ n"
    hence 1: "norm (X n - a) < r" by (rule N)
    hence 2: "X n ≠ 0" using r2 by auto
    hence "norm (inverse (X n)) = inverse (norm (X n))"
      by (rule nonzero_norm_inverse)
    also have "… ≤ inverse (norm a - r)"
    proof (rule le_imp_inverse_le)
      show "0 < norm a - r" using r2 by simp
    next
      have "norm a - norm (X n) ≤ norm (a - X n)"
        by (rule norm_triangle_ineq2)
      also have "… = norm (X n - a)"
        by (rule norm_minus_commute)
      also have "… < r" using 1 .
      finally show "norm a - r ≤ norm (X n)" by simp
    qed
    finally show "norm (inverse (X n)) ≤ inverse (norm a - r)" .
  qed
qed

lemma LIMSEQ_inverse_lemma:
  fixes a :: "'a::real_normed_div_algebra"
  shows "[|X ----> a; a ≠ 0; ∀n. X n ≠ 0|]
         ==> (λn. inverse (X n)) ----> inverse a"
apply (subst LIMSEQ_Zseq_iff)
apply (simp add: inverse_diff_inverse nonzero_imp_inverse_nonzero)
apply (rule Zseq_minus)
apply (rule Zseq_mult_left)
apply (rule mult.Bseq_prod_Zseq)
apply (erule (1) Bseq_inverse)
apply (simp add: LIMSEQ_Zseq_iff)
done

lemma LIMSEQ_inverse:
  fixes a :: "'a::real_normed_div_algebra"
  assumes X: "X ----> a"
  assumes a: "a ≠ 0"
  shows "(λn. inverse (X n)) ----> inverse a"
proof -
  from a have "0 < norm a" by simp
  then obtain k where "∀n≥k. norm (X n - a) < norm a"
    using LIMSEQ_D [OF X] by fast
  hence "∀n≥k. X n ≠ 0" by auto
  hence k: "∀n. X (n + k) ≠ 0" by simp

  from X have "(λn. X (n + k)) ----> a"
    by (rule LIMSEQ_ignore_initial_segment)
  hence "(λn. inverse (X (n + k))) ----> inverse a"
    using a k by (rule LIMSEQ_inverse_lemma)
  thus "(λn. inverse (X n)) ----> inverse a"
    by (rule LIMSEQ_offset)
qed

lemma LIMSEQ_divide:
  fixes a b :: "'a::real_normed_field"
  shows "[|X ----> a; Y ----> b; b ≠ 0|] ==> (λn. X n / Y n) ----> a / b"
by (simp add: LIMSEQ_mult LIMSEQ_inverse divide_inverse)

lemma LIMSEQ_pow:
  fixes a :: "'a::{real_normed_algebra,recpower}"
  shows "X ----> a ==> (λn. (X n) ^ m) ----> a ^ m"
by (induct m) (simp_all add: power_Suc LIMSEQ_const LIMSEQ_mult)

lemma LIMSEQ_setsum:
  assumes n: "!!n. n ∈ S ==> X n ----> L n"
  shows "(λm. ∑n∈S. X n m) ----> (∑n∈S. L n)"
proof (cases "finite S")
  case True
  thus ?thesis using n
  proof (induct)
    case empty
    show ?case
      by (simp add: LIMSEQ_const)
  next
    case insert
    thus ?case
      by (simp add: LIMSEQ_add)
  qed
next
  case False
  thus ?thesis
    by (simp add: LIMSEQ_const)
qed

lemma LIMSEQ_setprod:
  fixes L :: "'a => 'b::{real_normed_algebra,comm_ring_1}"
  assumes n: "!!n. n ∈ S ==> X n ----> L n"
  shows "(λm. ∏n∈S. X n m) ----> (∏n∈S. L n)"
proof (cases "finite S")
  case True
  thus ?thesis using n
  proof (induct)
    case empty
    show ?case
      by (simp add: LIMSEQ_const)
  next
    case insert
    thus ?case
      by (simp add: LIMSEQ_mult)
  qed
next
  case False
  thus ?thesis
    by (simp add: setprod_def LIMSEQ_const)
qed

lemma LIMSEQ_add_const: "f ----> a ==> (%n.(f n + b)) ----> a + b"
by (simp add: LIMSEQ_add LIMSEQ_const)

(* FIXME: delete *)
lemma LIMSEQ_add_minus:
     "[| X ----> a; Y ----> b |] ==> (%n. X n + -Y n) ----> a + -b"
by (simp only: LIMSEQ_add LIMSEQ_minus)

lemma LIMSEQ_diff_const: "f ----> a ==> (%n.(f n  - b)) ----> a - b"
by (simp add: LIMSEQ_diff LIMSEQ_const)

lemma LIMSEQ_diff_approach_zero: 
  "g ----> L ==> (%x. f x - g x) ----> 0  ==>
     f ----> L"
  apply (drule LIMSEQ_add)
  apply assumption
  apply simp
done

lemma LIMSEQ_diff_approach_zero2: 
  "f ----> L ==> (%x. f x - g x) ----> 0  ==>
     g ----> L";
  apply (drule LIMSEQ_diff)
  apply assumption
  apply simp
done

text{*A sequence tends to zero iff its abs does*}
lemma LIMSEQ_norm_zero: "((λn. norm (X n)) ----> 0) = (X ----> 0)"
by (simp add: LIMSEQ_def)

lemma LIMSEQ_rabs_zero: "((%n. ¦f n¦) ----> 0) = (f ----> (0::real))"
by (simp add: LIMSEQ_def)

lemma LIMSEQ_imp_rabs: "f ----> (l::real) ==> (%n. ¦f n¦) ----> ¦l¦"
by (drule LIMSEQ_norm, simp)

text{*An unbounded sequence's inverse tends to 0*}

lemma LIMSEQ_inverse_zero:
  "∀r::real. ∃N. ∀n≥N. r < X n ==> (λn. inverse (X n)) ----> 0"
apply (rule LIMSEQ_I)
apply (drule_tac x="inverse r" in spec, safe)
apply (rule_tac x="N" in exI, safe)
apply (drule_tac x="n" in spec, safe)
apply (frule positive_imp_inverse_positive)
apply (frule (1) less_imp_inverse_less)
apply (subgoal_tac "0 < X n", simp)
apply (erule (1) order_less_trans)
done

text{*The sequence @{term "1/n"} tends to 0 as @{term n} tends to infinity*}

lemma LIMSEQ_inverse_real_of_nat: "(%n. inverse(real(Suc n))) ----> 0"
apply (rule LIMSEQ_inverse_zero, safe)
apply (cut_tac x = r in reals_Archimedean2)
apply (safe, rule_tac x = n in exI)
apply (auto simp add: real_of_nat_Suc)
done

text{*The sequence @{term "r + 1/n"} tends to @{term r} as @{term n} tends to
infinity is now easily proved*}

lemma LIMSEQ_inverse_real_of_nat_add:
     "(%n. r + inverse(real(Suc n))) ----> r"
by (cut_tac LIMSEQ_add [OF LIMSEQ_const LIMSEQ_inverse_real_of_nat], auto)

lemma LIMSEQ_inverse_real_of_nat_add_minus:
     "(%n. r + -inverse(real(Suc n))) ----> r"
by (cut_tac LIMSEQ_add_minus [OF LIMSEQ_const LIMSEQ_inverse_real_of_nat], auto)

lemma LIMSEQ_inverse_real_of_nat_add_minus_mult:
     "(%n. r*( 1 + -inverse(real(Suc n)))) ----> r"
by (cut_tac b=1 in
        LIMSEQ_mult [OF LIMSEQ_const LIMSEQ_inverse_real_of_nat_add_minus], auto)

lemma LIMSEQ_le_const:
  "[|X ----> (x::real); ∃N. ∀n≥N. a ≤ X n|] ==> a ≤ x"
apply (rule ccontr, simp only: linorder_not_le)
apply (drule_tac r="a - x" in LIMSEQ_D, simp)
apply clarsimp
apply (drule_tac x="max N no" in spec, drule mp, rule le_maxI1)
apply (drule_tac x="max N no" in spec, drule mp, rule le_maxI2)
apply simp
done

lemma LIMSEQ_le_const2:
  "[|X ----> (x::real); ∃N. ∀n≥N. X n ≤ a|] ==> x ≤ a"
apply (subgoal_tac "- a ≤ - x", simp)
apply (rule LIMSEQ_le_const)
apply (erule LIMSEQ_minus)
apply simp
done

lemma LIMSEQ_le:
  "[|X ----> x; Y ----> y; ∃N. ∀n≥N. X n ≤ Y n|] ==> x ≤ (y::real)"
apply (subgoal_tac "0 ≤ y - x", simp)
apply (rule LIMSEQ_le_const)
apply (erule (1) LIMSEQ_diff)
apply (simp add: le_diff_eq)
done


subsection {* Convergence *}

lemma limI: "X ----> L ==> lim X = L"
apply (simp add: lim_def)
apply (blast intro: LIMSEQ_unique)
done

lemma convergentD: "convergent X ==> ∃L. (X ----> L)"
by (simp add: convergent_def)

lemma convergentI: "(X ----> L) ==> convergent X"
by (auto simp add: convergent_def)

lemma convergent_LIMSEQ_iff: "convergent X = (X ----> lim X)"
by (auto intro: theI LIMSEQ_unique simp add: convergent_def lim_def)

lemma convergent_minus_iff: "(convergent X) = (convergent (%n. -(X n)))"
apply (simp add: convergent_def)
apply (auto dest: LIMSEQ_minus)
apply (drule LIMSEQ_minus, auto)
done


subsection {* Bounded Monotonic Sequences *}

text{*Subsequence (alternative definition, (e.g. Hoskins)*}

lemma subseq_Suc_iff: "subseq f = (∀n. (f n) < (f (Suc n)))"
apply (simp add: subseq_def)
apply (auto dest!: less_imp_Suc_add)
apply (induct_tac k)
apply (auto intro: less_trans)
done

lemma monoseq_Suc:
   "monoseq X = ((∀n. X n ≤ X (Suc n))
                 | (∀n. X (Suc n) ≤ X n))"
apply (simp add: monoseq_def)
apply (auto dest!: le_imp_less_or_eq)
apply (auto intro!: lessI [THEN less_imp_le] dest!: less_imp_Suc_add)
apply (induct_tac "ka")
apply (auto intro: order_trans)
apply (erule contrapos_np)
apply (induct_tac "k")
apply (auto intro: order_trans)
done

lemma monoI1: "∀m. ∀ n ≥ m. X m ≤ X n ==> monoseq X"
by (simp add: monoseq_def)

lemma monoI2: "∀m. ∀ n ≥ m. X n ≤ X m ==> monoseq X"
by (simp add: monoseq_def)

lemma mono_SucI1: "∀n. X n ≤ X (Suc n) ==> monoseq X"
by (simp add: monoseq_Suc)

lemma mono_SucI2: "∀n. X (Suc n) ≤ X n ==> monoseq X"
by (simp add: monoseq_Suc)

text{*Bounded Sequence*}

lemma BseqD: "Bseq X ==> ∃K. 0 < K & (∀n. norm (X n) ≤ K)"
by (simp add: Bseq_def)

lemma BseqI: "[| 0 < K; ∀n. norm (X n) ≤ K |] ==> Bseq X"
by (auto simp add: Bseq_def)

lemma lemma_NBseq_def:
     "(∃K > 0. ∀n. norm (X n) ≤ K) =
      (∃N. ∀n. norm (X n) ≤ real(Suc N))"
apply auto
 prefer 2 apply force
apply (cut_tac x = K in reals_Archimedean2, clarify)
apply (rule_tac x = n in exI, clarify)
apply (drule_tac x = na in spec)
apply (auto simp add: real_of_nat_Suc)
done

text{* alternative definition for Bseq *}
lemma Bseq_iff: "Bseq X = (∃N. ∀n. norm (X n) ≤ real(Suc N))"
apply (simp add: Bseq_def)
apply (simp (no_asm) add: lemma_NBseq_def)
done

lemma lemma_NBseq_def2:
     "(∃K > 0. ∀n. norm (X n) ≤ K) = (∃N. ∀n. norm (X n) < real(Suc N))"
apply (subst lemma_NBseq_def, auto)
apply (rule_tac x = "Suc N" in exI)
apply (rule_tac [2] x = N in exI)
apply (auto simp add: real_of_nat_Suc)
 prefer 2 apply (blast intro: order_less_imp_le)
apply (drule_tac x = n in spec, simp)
done

(* yet another definition for Bseq *)
lemma Bseq_iff1a: "Bseq X = (∃N. ∀n. norm (X n) < real(Suc N))"
by (simp add: Bseq_def lemma_NBseq_def2)

subsubsection{*Upper Bounds and Lubs of Bounded Sequences*}

lemma Bseq_isUb:
  "!!(X::nat=>real). Bseq X ==> ∃U. isUb (UNIV::real set) {x. ∃n. X n = x} U"
by (auto intro: isUbI setleI simp add: Bseq_def abs_le_iff)


text{* Use completeness of reals (supremum property)
   to show that any bounded sequence has a least upper bound*}

lemma Bseq_isLub:
  "!!(X::nat=>real). Bseq X ==>
   ∃U. isLub (UNIV::real set) {x. ∃n. X n = x} U"
by (blast intro: reals_complete Bseq_isUb)

subsubsection{*A Bounded and Monotonic Sequence Converges*}

lemma lemma_converg1:
     "!!(X::nat=>real). [| ∀m. ∀ n ≥ m. X m ≤ X n;
                  isLub (UNIV::real set) {x. ∃n. X n = x} (X ma)
               |] ==> ∀n ≥ ma. X n = X ma"
apply safe
apply (drule_tac y = "X n" in isLubD2)
apply (blast dest: order_antisym)+
done

text{* The best of both worlds: Easier to prove this result as a standard
   theorem and then use equivalence to "transfer" it into the
   equivalent nonstandard form if needed!*}

lemma Bmonoseq_LIMSEQ: "∀n. m ≤ n --> X n = X m ==> ∃L. (X ----> L)"
apply (simp add: LIMSEQ_def)
apply (rule_tac x = "X m" in exI, safe)
apply (rule_tac x = m in exI, safe)
apply (drule spec, erule impE, auto)
done

lemma lemma_converg2:
   "!!(X::nat=>real).
    [| ∀m. X m ~= U;  isLub UNIV {x. ∃n. X n = x} U |] ==> ∀m. X m < U"
apply safe
apply (drule_tac y = "X m" in isLubD2)
apply (auto dest!: order_le_imp_less_or_eq)
done

lemma lemma_converg3: "!!(X ::nat=>real). ∀m. X m ≤ U ==> isUb UNIV {x. ∃n. X n = x} U"
by (rule setleI [THEN isUbI], auto)

text{* FIXME: @{term "U - T < U"} is redundant *}
lemma lemma_converg4: "!!(X::nat=> real).
               [| ∀m. X m ~= U;
                  isLub UNIV {x. ∃n. X n = x} U;
                  0 < T;
                  U + - T < U
               |] ==> ∃m. U + -T < X m & X m < U"
apply (drule lemma_converg2, assumption)
apply (rule ccontr, simp)
apply (simp add: linorder_not_less)
apply (drule lemma_converg3)
apply (drule isLub_le_isUb, assumption)
apply (auto dest: order_less_le_trans)
done

text{*A standard proof of the theorem for monotone increasing sequence*}

lemma Bseq_mono_convergent:
     "[| Bseq X; ∀m. ∀n ≥ m. X m ≤ X n |] ==> convergent (X::nat=>real)"
apply (simp add: convergent_def)
apply (frule Bseq_isLub, safe)
apply (case_tac "∃m. X m = U", auto)
apply (blast dest: lemma_converg1 Bmonoseq_LIMSEQ)
(* second case *)
apply (rule_tac x = U in exI)
apply (subst LIMSEQ_iff, safe)
apply (frule lemma_converg2, assumption)
apply (drule lemma_converg4, auto)
apply (rule_tac x = m in exI, safe)
apply (subgoal_tac "X m ≤ X n")
 prefer 2 apply blast
apply (drule_tac x=n and P="%m. X m < U" in spec, arith)
done

lemma Bseq_minus_iff: "Bseq (%n. -(X n)) = Bseq X"
by (simp add: Bseq_def)

text{*Main monotonicity theorem*}
lemma Bseq_monoseq_convergent: "[| Bseq X; monoseq X |] ==> convergent X"
apply (simp add: monoseq_def, safe)
apply (rule_tac [2] convergent_minus_iff [THEN ssubst])
apply (drule_tac [2] Bseq_minus_iff [THEN ssubst])
apply (auto intro!: Bseq_mono_convergent)
done

subsubsection{*A Few More Equivalence Theorems for Boundedness*}

text{*alternative formulation for boundedness*}
lemma Bseq_iff2: "Bseq X = (∃k > 0. ∃x. ∀n. norm (X(n) + -x) ≤ k)"
apply (unfold Bseq_def, safe)
apply (rule_tac [2] x = "k + norm x" in exI)
apply (rule_tac x = K in exI, simp)
apply (rule exI [where x = 0], auto)
apply (erule order_less_le_trans, simp)
apply (drule_tac x=n in spec, fold diff_def)
apply (drule order_trans [OF norm_triangle_ineq2])
apply simp
done

text{*alternative formulation for boundedness*}
lemma Bseq_iff3: "Bseq X = (∃k > 0. ∃N. ∀n. norm(X(n) + -X(N)) ≤ k)"
apply safe
apply (simp add: Bseq_def, safe)
apply (rule_tac x = "K + norm (X N)" in exI)
apply auto
apply (erule order_less_le_trans, simp)
apply (rule_tac x = N in exI, safe)
apply (drule_tac x = n in spec)
apply (rule order_trans [OF norm_triangle_ineq], simp)
apply (auto simp add: Bseq_iff2)
done

lemma BseqI2: "(∀n. k ≤ f n & f n ≤ (K::real)) ==> Bseq f"
apply (simp add: Bseq_def)
apply (rule_tac x = " (¦k¦ + ¦K¦) + 1" in exI, auto)
apply (drule_tac x = n in spec, arith)
done


subsection {* Cauchy Sequences *}

lemma CauchyI:
  "(!!e. 0 < e ==> ∃M. ∀m≥M. ∀n≥M. norm (X m - X n) < e) ==> Cauchy X"
by (simp add: Cauchy_def)

lemma CauchyD:
  "[|Cauchy X; 0 < e|] ==> ∃M. ∀m≥M. ∀n≥M. norm (X m - X n) < e"
by (simp add: Cauchy_def)

subsubsection {* Cauchy Sequences are Bounded *}

text{*A Cauchy sequence is bounded -- this is the standard
  proof mechanization rather than the nonstandard proof*}

lemma lemmaCauchy: "∀n ≥ M. norm (X M - X n) < (1::real)
          ==>  ∀n ≥ M. norm (X n :: 'a::real_normed_vector) < 1 + norm (X M)"
apply (clarify, drule spec, drule (1) mp)
apply (simp only: norm_minus_commute)
apply (drule order_le_less_trans [OF norm_triangle_ineq2])
apply simp
done

lemma Cauchy_Bseq: "Cauchy X ==> Bseq X"
apply (simp add: Cauchy_def)
apply (drule spec, drule mp, rule zero_less_one, safe)
apply (drule_tac x="M" in spec, simp)
apply (drule lemmaCauchy)
apply (rule_tac k="M" in Bseq_offset)
apply (simp add: Bseq_def)
apply (rule_tac x="1 + norm (X M)" in exI)
apply (rule conjI, rule order_less_le_trans [OF zero_less_one], simp)
apply (simp add: order_less_imp_le)
done

subsubsection {* Cauchy Sequences are Convergent *}

axclass banach ⊆ real_normed_vector
  Cauchy_convergent: "Cauchy X ==> convergent X"

theorem LIMSEQ_imp_Cauchy:
  assumes X: "X ----> a" shows "Cauchy X"
proof (rule CauchyI)
  fix e::real assume "0 < e"
  hence "0 < e/2" by simp
  with X have "∃N. ∀n≥N. norm (X n - a) < e/2" by (rule LIMSEQ_D)
  then obtain N where N: "∀n≥N. norm (X n - a) < e/2" ..
  show "∃N. ∀m≥N. ∀n≥N. norm (X m - X n) < e"
  proof (intro exI allI impI)
    fix m assume "N ≤ m"
    hence m: "norm (X m - a) < e/2" using N by fast
    fix n assume "N ≤ n"
    hence n: "norm (X n - a) < e/2" using N by fast
    have "norm (X m - X n) = norm ((X m - a) - (X n - a))" by simp
    also have "… ≤ norm (X m - a) + norm (X n - a)"
      by (rule norm_triangle_ineq4)
    also from m n have "… < e" by(simp add:field_simps)
    finally show "norm (X m - X n) < e" .
  qed
qed

lemma convergent_Cauchy: "convergent X ==> Cauchy X"
unfolding convergent_def
by (erule exE, erule LIMSEQ_imp_Cauchy)

text {*
Proof that Cauchy sequences converge based on the one from
http://pirate.shu.edu/~wachsmut/ira/numseq/proofs/cauconv.html
*}

text {*
  If sequence @{term "X"} is Cauchy, then its limit is the lub of
  @{term "{r::real. ∃N. ∀n≥N. r < X n}"}
*}

lemma isUb_UNIV_I: "(!!y. y ∈ S ==> y ≤ u) ==> isUb UNIV S u"
by (simp add: isUbI setleI)

lemma real_abs_diff_less_iff:
  "(¦x - a¦ < (r::real)) = (a - r < x ∧ x < a + r)"
by auto

locale (open) real_Cauchy =
  fixes X :: "nat => real"
  assumes X: "Cauchy X"
  fixes S :: "real set"
  defines S_def: "S ≡ {x::real. ∃N. ∀n≥N. x < X n}"

lemma (in real_Cauchy) mem_S: "∀n≥N. x < X n ==> x ∈ S"
by (unfold S_def, auto)

lemma (in real_Cauchy) bound_isUb:
  assumes N: "∀n≥N. X n < x"
  shows "isUb UNIV S x"
proof (rule isUb_UNIV_I)
  fix y::real assume "y ∈ S"
  hence "∃M. ∀n≥M. y < X n"
    by (simp add: S_def)
  then obtain M where "∀n≥M. y < X n" ..
  hence "y < X (max M N)" by simp
  also have "… < x" using N by simp
  finally show "y ≤ x"
    by (rule order_less_imp_le)
qed

lemma (in real_Cauchy) isLub_ex: "∃u. isLub UNIV S u"
proof (rule reals_complete)
  obtain N where "∀m≥N. ∀n≥N. norm (X m - X n) < 1"
    using CauchyD [OF X zero_less_one] by fast
  hence N: "∀n≥N. norm (X n - X N) < 1" by simp
  show "∃x. x ∈ S"
  proof
    from N have "∀n≥N. X N - 1 < X n"
      by (simp add: real_abs_diff_less_iff)
    thus "X N - 1 ∈ S" by (rule mem_S)
  qed
  show "∃u. isUb UNIV S u"
  proof
    from N have "∀n≥N. X n < X N + 1"
      by (simp add: real_abs_diff_less_iff)
    thus "isUb UNIV S (X N + 1)"
      by (rule bound_isUb)
  qed
qed

lemma (in real_Cauchy) isLub_imp_LIMSEQ:
  assumes x: "isLub UNIV S x"
  shows "X ----> x"
proof (rule LIMSEQ_I)
  fix r::real assume "0 < r"
  hence r: "0 < r/2" by simp
  obtain N where "∀n≥N. ∀m≥N. norm (X n - X m) < r/2"
    using CauchyD [OF X r] by fast
  hence "∀n≥N. norm (X n - X N) < r/2" by simp
  hence N: "∀n≥N. X N - r/2 < X n ∧ X n < X N + r/2"
    by (simp only: real_norm_def real_abs_diff_less_iff)

  from N have "∀n≥N. X N - r/2 < X n" by fast
  hence "X N - r/2 ∈ S" by (rule mem_S)
  hence 1: "X N - r/2 ≤ x" using x isLub_isUb isUbD by fast

  from N have "∀n≥N. X n < X N + r/2" by fast
  hence "isUb UNIV S (X N + r/2)" by (rule bound_isUb)
  hence 2: "x ≤ X N + r/2" using x isLub_le_isUb by fast

  show "∃N. ∀n≥N. norm (X n - x) < r"
  proof (intro exI allI impI)
    fix n assume n: "N ≤ n"
    from N n have "X n < X N + r/2" and "X N - r/2 < X n" by simp+
    thus "norm (X n - x) < r" using 1 2
      by (simp add: real_abs_diff_less_iff)
  qed
qed

lemma (in real_Cauchy) LIMSEQ_ex: "∃x. X ----> x"
proof -
  obtain x where "isLub UNIV S x"
    using isLub_ex by fast
  hence "X ----> x"
    by (rule isLub_imp_LIMSEQ)
  thus ?thesis ..
qed

lemma real_Cauchy_convergent:
  fixes X :: "nat => real"
  shows "Cauchy X ==> convergent X"
unfolding convergent_def by (rule real_Cauchy.LIMSEQ_ex)

instance real :: banach
by intro_classes (rule real_Cauchy_convergent)

lemma Cauchy_convergent_iff:
  fixes X :: "nat => 'a::banach"
  shows "Cauchy X = convergent X"
by (fast intro: Cauchy_convergent convergent_Cauchy)


subsection {* Power Sequences *}

text{*The sequence @{term "x^n"} tends to 0 if @{term "0≤x"} and @{term
"x<1"}.  Proof will use (NS) Cauchy equivalence for convergence and
  also fact that bounded and monotonic sequence converges.*}

lemma Bseq_realpow: "[| 0 ≤ (x::real); x ≤ 1 |] ==> Bseq (%n. x ^ n)"
apply (simp add: Bseq_def)
apply (rule_tac x = 1 in exI)
apply (simp add: power_abs)
apply (auto dest: power_mono)
done

lemma monoseq_realpow: "[| 0 ≤ x; x ≤ 1 |] ==> monoseq (%n. x ^ n)"
apply (clarify intro!: mono_SucI2)
apply (cut_tac n = n and N = "Suc n" and a = x in power_decreasing, auto)
done

lemma convergent_realpow:
  "[| 0 ≤ (x::real); x ≤ 1 |] ==> convergent (%n. x ^ n)"
by (blast intro!: Bseq_monoseq_convergent Bseq_realpow monoseq_realpow)

lemma LIMSEQ_inverse_realpow_zero_lemma:
  fixes x :: real
  assumes x: "0 ≤ x"
  shows "real n * x + 1 ≤ (x + 1) ^ n"
apply (induct n)
apply simp
apply simp
apply (rule order_trans)
prefer 2
apply (erule mult_left_mono)
apply (rule add_increasing [OF x], simp)
apply (simp add: real_of_nat_Suc)
apply (simp add: ring_distribs)
apply (simp add: mult_nonneg_nonneg x)
done

lemma LIMSEQ_inverse_realpow_zero:
  "1 < (x::real) ==> (λn. inverse (x ^ n)) ----> 0"
proof (rule LIMSEQ_inverse_zero [rule_format])
  fix y :: real
  assume x: "1 < x"
  hence "0 < x - 1" by simp
  hence "∀y. ∃N::nat. y < real N * (x - 1)"
    by (rule reals_Archimedean3)
  hence "∃N::nat. y < real N * (x - 1)" ..
  then obtain N::nat where "y < real N * (x - 1)" ..
  also have "… ≤ real N * (x - 1) + 1" by simp
  also have "… ≤ (x - 1 + 1) ^ N"
    by (rule LIMSEQ_inverse_realpow_zero_lemma, cut_tac x, simp)
  also have "… = x ^ N" by simp
  finally have "y < x ^ N" .
  hence "∀n≥N. y < x ^ n"
    apply clarify
    apply (erule order_less_le_trans)
    apply (erule power_increasing)
    apply (rule order_less_imp_le [OF x])
    done
  thus "∃N. ∀n≥N. y < x ^ n" ..
qed

lemma LIMSEQ_realpow_zero:
  "[|0 ≤ (x::real); x < 1|] ==> (λn. x ^ n) ----> 0"
proof (cases)
  assume "x = 0"
  hence "(λn. x ^ Suc n) ----> 0" by (simp add: LIMSEQ_const)
  thus ?thesis by (rule LIMSEQ_imp_Suc)
next
  assume "0 ≤ x" and "x ≠ 0"
  hence x0: "0 < x" by simp
  assume x1: "x < 1"
  from x0 x1 have "1 < inverse x"
    by (rule real_inverse_gt_one)
  hence "(λn. inverse (inverse x ^ n)) ----> 0"
    by (rule LIMSEQ_inverse_realpow_zero)
  thus ?thesis by (simp add: power_inverse)
qed

lemma LIMSEQ_power_zero:
  fixes x :: "'a::{real_normed_algebra_1,recpower}"
  shows "norm x < 1 ==> (λn. x ^ n) ----> 0"
apply (drule LIMSEQ_realpow_zero [OF norm_ge_zero])
apply (simp only: LIMSEQ_Zseq_iff, erule Zseq_le)
apply (simp add: power_abs norm_power_ineq)
done

lemma LIMSEQ_divide_realpow_zero:
  "1 < (x::real) ==> (%n. a / (x ^ n)) ----> 0"
apply (cut_tac a = a and x1 = "inverse x" in
        LIMSEQ_mult [OF LIMSEQ_const LIMSEQ_realpow_zero])
apply (auto simp add: divide_inverse power_inverse)
apply (simp add: inverse_eq_divide pos_divide_less_eq)
done

text{*Limit of @{term "c^n"} for @{term"¦c¦ < 1"}*}

lemma LIMSEQ_rabs_realpow_zero: "¦c¦ < (1::real) ==> (%n. ¦c¦ ^ n) ----> 0"
by (rule LIMSEQ_realpow_zero [OF abs_ge_zero])

lemma LIMSEQ_rabs_realpow_zero2: "¦c¦ < (1::real) ==> (%n. c ^ n) ----> 0"
apply (rule LIMSEQ_rabs_zero [THEN iffD1])
apply (auto intro: LIMSEQ_rabs_realpow_zero simp add: power_abs)
done

end

Bounded Sequences

lemma BseqI':

  (!!n. norm (X n)  K) ==> Bseq X

lemma BseqE:

  [| Bseq X; !!K. [| 0 < K; ∀n. norm (X n)  K |] ==> Q |] ==> Q

lemma BseqI2':

  nN. norm (X n)  K ==> Bseq X

lemma Bseq_ignore_initial_segment:

  Bseq X ==> Bseqn. X (n + k))

lemma Bseq_offset:

  Bseqn. X (n + k)) ==> Bseq X

Sequences That Converge to Zero

lemma ZseqI:

  (!!r. 0 < r ==> ∃no. ∀nno. norm (X n) < r) ==> Zseq X

lemma ZseqD:

  [| Zseq X; 0 < r |] ==> ∃no. ∀nno. norm (X n) < r

lemma Zseq_zero:

  Zseqn. 0::'a)

lemma Zseq_const_iff:

  Zseqn. k) = (k = (0::'a))

lemma Zseq_norm_iff:

  Zseqn. norm (X n)) = Zseq X

lemma Zseq_imp_Zseq:

  [| Zseq X; !!n. norm (Y n)  norm (X n) * K |] ==> Zseq Y

lemma Zseq_le:

  [| Zseq Y; ∀n. norm (X n)  norm (Y n) |] ==> Zseq X

lemma Zseq_add:

  [| Zseq X; Zseq Y |] ==> Zseqn. X n + Y n)

lemma Zseq_minus:

  Zseq X ==> Zseqn. - X n)

lemma Zseq_diff:

  [| Zseq X; Zseq Y |] ==> Zseqn. X n - Y n)

lemma Zseq:

  Zseq X ==> Zseqn. f (X n))

lemma Zseq:

  [| Zseq X; Zseq Y |] ==> Zseqn. X n ** Y n)

lemma Zseq_prod_Bseq:

  [| Zseq X; Bseq Y |] ==> Zseqn. X n ** Y n)

lemma Bseq_prod_Zseq:

  [| Bseq X; Zseq Y |] ==> Zseqn. X n ** Y n)

lemma Zseq_left:

  Zseq X ==> Zseqn. X n ** a)

lemma Zseq_right:

  Zseq X ==> Zseqn. a ** X n)

lemma Zseq_mult:

  [| Zseq X; Zseq Y |] ==> Zseqn. X n * Y n)

lemma Zseq_mult_right:

  Zseq X ==> Zseqn. a * X n)

lemma Zseq_mult_left:

  Zseq X ==> Zseqn. X n * a)

Limits of Sequences

lemma LIMSEQ_iff:

  X ----> L = (∀r>0. ∃no. ∀nno. norm (X n - L) < r)

lemma LIMSEQ_Zseq_iff:

  X ----> L = Zseqn. X n - L)

lemma LIMSEQ_I:

  (!!r. 0 < r ==> ∃no. ∀nno. norm (X n - L) < r) ==> X ----> L

lemma LIMSEQ_D:

  [| X ----> L; 0 < r |] ==> ∃no. ∀nno. norm (X n - L) < r

lemma LIMSEQ_const:

  n. k) ----> k

lemma LIMSEQ_const_iff:

  n. k) ----> l = (k = l)

lemma LIMSEQ_norm:

  X ----> a ==> (λn. norm (X n)) ----> norm a

lemma LIMSEQ_ignore_initial_segment:

  f ----> a ==> (λn. f (n + k)) ----> a

lemma LIMSEQ_offset:

  n. f (n + k)) ----> a ==> f ----> a

lemma LIMSEQ_Suc:

  f ----> l ==> (λn. f (Suc n)) ----> l

lemma LIMSEQ_imp_Suc:

  n. f (Suc n)) ----> l ==> f ----> l

lemma LIMSEQ_Suc_iff:

  n. f (Suc n)) ----> l = f ----> l

lemma add_diff_add:

  a + c - (b + d) = a - b + (c - d)

lemma minus_diff_minus:

  - a - - b = - (a - b)

lemma LIMSEQ_add:

  [| X ----> a; Y ----> b |] ==> (λn. X n + Y n) ----> a + b

lemma LIMSEQ_minus:

  X ----> a ==> (λn. - X n) ----> - a

lemma LIMSEQ_minus_cancel:

  n. - X n) ----> - a ==> X ----> a

lemma LIMSEQ_diff:

  [| X ----> a; Y ----> b |] ==> (λn. X n - Y n) ----> a - b

lemma LIMSEQ_unique:

  [| X ----> a; X ----> b |] ==> a = b

lemma LIMSEQ:

  X ----> a ==> (λn. f (X n)) ----> f a

lemma LIMSEQ:

  [| X ----> a; Y ----> b |] ==> (λn. X n ** Y n) ----> a ** b

lemma LIMSEQ_mult:

  [| X ----> a; Y ----> b |] ==> (λn. X n * Y n) ----> a * b

lemma inverse_diff_inverse:

  [| a  (0::'a); b  (0::'a) |]
  ==> inverse a - inverse b = - (inverse a * (a - b) * inverse b)

lemma Bseq_inverse_lemma:

  [| r  norm x; 0 < r |] ==> norm (inverse x)  inverse r

lemma Bseq_inverse:

  [| X ----> a; a  (0::'a) |] ==> Bseqn. inverse (X n))

lemma LIMSEQ_inverse_lemma:

  [| X ----> a; a  (0::'a); ∀n. X n  (0::'a) |]
  ==> (λn. inverse (X n)) ----> inverse a

lemma LIMSEQ_inverse:

  [| X ----> a; a  (0::'a) |] ==> (λn. inverse (X n)) ----> inverse a

lemma LIMSEQ_divide:

  [| X ----> a; Y ----> b; b  (0::'a) |] ==> (λn. X n / Y n) ----> a / b

lemma LIMSEQ_pow:

  X ----> a ==> (λn. X n ^ m) ----> a ^ m

lemma LIMSEQ_setsum:

  (!!n. nS ==> X n ----> L n) ==> (λm. ∑nS. X n m) ----> setsum L S

lemma LIMSEQ_setprod:

  (!!n. nS ==> X n ----> L n) ==> (λm. ∏nS. X n m) ----> setprod L S

lemma LIMSEQ_add_const:

  f ----> a ==> (λn. f n + b) ----> a + b

lemma LIMSEQ_add_minus:

  [| X ----> a; Y ----> b |] ==> (λn. X n + - Y n) ----> a + - b

lemma LIMSEQ_diff_const:

  f ----> a ==> (λn. f n - b) ----> a - b

lemma LIMSEQ_diff_approach_zero:

  [| g ----> L; (λx. f x - g x) ----> (0::'a) |] ==> f ----> L

lemma LIMSEQ_diff_approach_zero2:

  [| f ----> L; (λx. f x - g x) ----> (0::'a) |] ==> g ----> L

lemma LIMSEQ_norm_zero:

  n. norm (X n)) ----> 0 = X ----> (0::'a)

lemma LIMSEQ_rabs_zero:

  n. ¦f n¦) ----> 0 = f ----> 0

lemma LIMSEQ_imp_rabs:

  f ----> l ==> (λn. ¦f n¦) ----> ¦l¦

lemma LIMSEQ_inverse_zero:

  r. ∃N. ∀nN. r < X n ==> (λn. inverse (X n)) ----> 0

lemma LIMSEQ_inverse_real_of_nat:

  n. inverse (real (Suc n))) ----> 0

lemma LIMSEQ_inverse_real_of_nat_add:

  n. r + inverse (real (Suc n))) ----> r

lemma LIMSEQ_inverse_real_of_nat_add_minus:

  n. r + - inverse (real (Suc n))) ----> r

lemma LIMSEQ_inverse_real_of_nat_add_minus_mult:

  n. r * (1 + - inverse (real (Suc n)))) ----> r

lemma LIMSEQ_le_const:

  [| X ----> x; ∃N. ∀nN. a  X n |] ==> a  x

lemma LIMSEQ_le_const2:

  [| X ----> x; ∃N. ∀nN. X n  a |] ==> x  a

lemma LIMSEQ_le:

  [| X ----> x; Y ----> y; ∃N. ∀nN. X n  Y n |] ==> x  y

Convergence

lemma limI:

  X ----> L ==> lim X = L

lemma convergentD:

  convergent X ==> ∃L. X ----> L

lemma convergentI:

  X ----> L ==> convergent X

lemma convergent_LIMSEQ_iff:

  convergent X = X ----> lim X

lemma convergent_minus_iff:

  convergent X = convergentn. - X n)

Bounded Monotonic Sequences

lemma subseq_Suc_iff:

  subseq f = (∀n. f n < f (Suc n))

lemma monoseq_Suc:

  monoseq X = ((∀n. X n  X (Suc n)) ∨ (∀n. X (Suc n)  X n))

lemma monoI1:

  m n. m  n --> X m  X n ==> monoseq X

lemma monoI2:

  m n. m  n --> X n  X m ==> monoseq X

lemma mono_SucI1:

  n. X n  X (Suc n) ==> monoseq X

lemma mono_SucI2:

  n. X (Suc n)  X n ==> monoseq X

lemma BseqD:

  Bseq X ==> ∃K>0. ∀n. norm (X n)  K

lemma BseqI:

  [| 0 < K; ∀n. norm (X n)  K |] ==> Bseq X

lemma lemma_NBseq_def:

  (∃K>0. ∀n. norm (X n)  K) = (∃N. ∀n. norm (X n)  real (Suc N))

lemma Bseq_iff:

  Bseq X = (∃N. ∀n. norm (X n)  real (Suc N))

lemma lemma_NBseq_def2:

  (∃K>0. ∀n. norm (X n)  K) = (∃N. ∀n. norm (X n) < real (Suc N))

lemma Bseq_iff1a:

  Bseq X = (∃N. ∀n. norm (X n) < real (Suc N))

Upper Bounds and Lubs of Bounded Sequences

lemma Bseq_isUb:

  Bseq X ==> ∃U. isUb UNIV {x. ∃n. X n = x} U

lemma Bseq_isLub:

  Bseq X ==> ∃U. isLub UNIV {x. ∃n. X n = x} U

A Bounded and Monotonic Sequence Converges

lemma lemma_converg1:

  [| ∀m n. m  n --> X m  X n; isLub UNIV {x. ∃n. X n = x} (X ma) |]
  ==> ∀nma. X n = X ma

lemma Bmonoseq_LIMSEQ:

  nm. X n = X m ==> ∃L. X ----> L

lemma lemma_converg2:

  [| ∀m. X m  U; isLub UNIV {x. ∃n. X n = x} U |] ==> ∀m. X m < U

lemma lemma_converg3:

  m. X m  U ==> isUb UNIV {x. ∃n. X n = x} U

lemma lemma_converg4:

  [| ∀m. X m  U; isLub UNIV {x. ∃n. X n = x} U; 0 < T; U + - T < U |]
  ==> ∃m. U + - T < X mX m < U

lemma Bseq_mono_convergent:

  [| Bseq X; ∀m n. m  n --> X m  X n |] ==> convergent X

lemma Bseq_minus_iff:

  Bseqn. - X n) = Bseq X

lemma Bseq_monoseq_convergent:

  [| Bseq X; monoseq X |] ==> convergent X

A Few More Equivalence Theorems for Boundedness

lemma Bseq_iff2:

  Bseq X = (∃k>0. ∃x. ∀n. norm (X n + - x)  k)

lemma Bseq_iff3:

  Bseq X = (∃k>0. ∃N. ∀n. norm (X n + - X N)  k)

lemma BseqI2:

  n. k  f nf n  K ==> Bseq f

Cauchy Sequences

lemma CauchyI:

  (!!e. 0 < e ==> ∃M. ∀mM. ∀nM. norm (X m - X n) < e) ==> Cauchy X

lemma CauchyD:

  [| Cauchy X; 0 < e |] ==> ∃M. ∀mM. ∀nM. norm (X m - X n) < e

Cauchy Sequences are Bounded

lemma lemmaCauchy:

  nM. norm (X M - X n) < 1 ==> ∀nM. norm (X n) < 1 + norm (X M)

lemma Cauchy_Bseq:

  Cauchy X ==> Bseq X

Cauchy Sequences are Convergent

theorem LIMSEQ_imp_Cauchy:

  X ----> a ==> Cauchy X

lemma convergent_Cauchy:

  convergent X ==> Cauchy X

lemma isUb_UNIV_I:

  (!!y. yS ==> y  u) ==> isUb UNIV S u

lemma real_abs_diff_less_iff:

  (¦x - a¦ < r) = (a - r < xx < a + r)

lemma mem_S:

  nN. x < X n ==> xS

lemma bound_isUb:

  nN. X n < x ==> isUb UNIV S x

lemma isLub_ex:

  u. isLub UNIV S u

lemma isLub_imp_LIMSEQ:

  isLub UNIV S x ==> X ----> x

lemma LIMSEQ_ex:

  x. X ----> x

lemma real_Cauchy_convergent:

  Cauchy X ==> convergent X

lemma Cauchy_convergent_iff:

  Cauchy X = convergent X

Power Sequences

lemma Bseq_realpow:

  [| 0  x; x  1 |] ==> Bseq (op ^ x)

lemma monoseq_realpow:

  [| 0  x; x  1 |] ==> monoseq (op ^ x)

lemma convergent_realpow:

  [| 0  x; x  1 |] ==> convergent (op ^ x)

lemma LIMSEQ_inverse_realpow_zero_lemma:

  0  x ==> real n * x + 1  (x + 1) ^ n

lemma LIMSEQ_inverse_realpow_zero:

  1 < x ==> (λn. inverse (x ^ n)) ----> 0

lemma LIMSEQ_realpow_zero:

  [| 0  x; x < 1 |] ==> op ^ x ----> 0

lemma LIMSEQ_power_zero:

  norm x < 1 ==> op ^ x ----> (0::'a)

lemma LIMSEQ_divide_realpow_zero:

  1 < x ==> (λn. a / x ^ n) ----> 0

lemma LIMSEQ_rabs_realpow_zero:

  ¦c¦ < 1 ==> op ^ ¦c¦ ----> 0

lemma LIMSEQ_rabs_realpow_zero2:

  ¦c¦ < 1 ==> op ^ c ----> 0