src/HOL/Nonstandard_Analysis/HSEQ.thy
author eberlm <eberlm@in.tum.de>
Thu, 30 Nov 2017 16:59:59 +0100
changeset 67107 cef76a19125e
parent 67006 b1278ed3cd46
child 67443 3abf6a722518
permissions -rw-r--r--
Existence of a holomorphic logarithm
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
62479
716336f19aa9 clarified session;
wenzelm
parents: 61975
diff changeset
     1
(*  Title:      HOL/Nonstandard_Analysis/HSEQ.thy
716336f19aa9 clarified session;
wenzelm
parents: 61975
diff changeset
     2
    Author:     Jacques D. Fleuriot
716336f19aa9 clarified session;
wenzelm
parents: 61975
diff changeset
     3
    Copyright:  1998  University of Cambridge
716336f19aa9 clarified session;
wenzelm
parents: 61975
diff changeset
     4
716336f19aa9 clarified session;
wenzelm
parents: 61975
diff changeset
     5
Convergence of sequences and series.
716336f19aa9 clarified session;
wenzelm
parents: 61975
diff changeset
     6
716336f19aa9 clarified session;
wenzelm
parents: 61975
diff changeset
     7
Conversion to Isar and new proofs by Lawrence C Paulson, 2004
716336f19aa9 clarified session;
wenzelm
parents: 61975
diff changeset
     8
Additional contributions by Jeremy Avigad and Brian Huffman.
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
     9
*)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    10
61975
b4b11391c676 isabelle update_cartouches -c -t;
wenzelm
parents: 61970
diff changeset
    11
section \<open>Sequences and Convergence (Nonstandard)\<close>
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    12
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    13
theory HSEQ
67006
b1278ed3cd46 prefer main entry points of HOL;
wenzelm
parents: 66453
diff changeset
    14
  imports Complex_Main NatStar
63579
73939a9b70a3 support 'abbrevs' within theory header;
wenzelm
parents: 62479
diff changeset
    15
  abbrevs "--->" = "\<longlonglongrightarrow>\<^sub>N\<^sub>S"
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    16
begin
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    17
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
    18
definition NSLIMSEQ :: "(nat \<Rightarrow> 'a::real_normed_vector) \<Rightarrow> 'a \<Rightarrow> bool"
61970
6226261144d7 more symbols;
wenzelm
parents: 61969
diff changeset
    19
    ("((_)/ \<longlonglongrightarrow>\<^sub>N\<^sub>S (_))" [60, 60] 60) where
61975
b4b11391c676 isabelle update_cartouches -c -t;
wenzelm
parents: 61970
diff changeset
    20
    \<comment>\<open>Nonstandard definition of convergence of sequence\<close>
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
    21
  "X \<longlonglongrightarrow>\<^sub>N\<^sub>S L \<longleftrightarrow> (\<forall>N \<in> HNatInfinite. ( *f* X) N \<approx> star_of L)"
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    22
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
    23
definition nslim :: "(nat \<Rightarrow> 'a::real_normed_vector) \<Rightarrow> 'a"
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
    24
  where "nslim X = (THE L. X \<longlonglongrightarrow>\<^sub>N\<^sub>S L)"
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
    25
  \<comment> \<open>Nonstandard definition of limit using choice operator\<close>
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
    26
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    27
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
    28
definition NSconvergent :: "(nat \<Rightarrow> 'a::real_normed_vector) \<Rightarrow> bool"
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
    29
  where "NSconvergent X \<longleftrightarrow> (\<exists>L. X \<longlonglongrightarrow>\<^sub>N\<^sub>S L)"
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
    30
  \<comment> \<open>Nonstandard definition of convergence\<close>
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    31
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
    32
definition NSBseq :: "(nat \<Rightarrow> 'a::real_normed_vector) \<Rightarrow> bool"
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
    33
  where "NSBseq X \<longleftrightarrow> (\<forall>N \<in> HNatInfinite. ( *f* X) N \<in> HFinite)"
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
    34
  \<comment> \<open>Nonstandard definition for bounded sequence\<close>
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
    35
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    36
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
    37
definition NSCauchy :: "(nat \<Rightarrow> 'a::real_normed_vector) \<Rightarrow> bool"
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
    38
  where "NSCauchy X \<longleftrightarrow> (\<forall>M \<in> HNatInfinite. \<forall>N \<in> HNatInfinite. ( *f* X) M \<approx> ( *f* X) N)"
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
    39
  \<comment> \<open>Nonstandard definition\<close>
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
    40
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    41
61975
b4b11391c676 isabelle update_cartouches -c -t;
wenzelm
parents: 61970
diff changeset
    42
subsection \<open>Limits of Sequences\<close>
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    43
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
    44
lemma NSLIMSEQ_iff: "(X \<longlonglongrightarrow>\<^sub>N\<^sub>S L) \<longleftrightarrow> (\<forall>N \<in> HNatInfinite. ( *f* X) N \<approx> star_of L)"
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
    45
  by (simp add: NSLIMSEQ_def)
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
    46
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
    47
lemma NSLIMSEQ_I: "(\<And>N. N \<in> HNatInfinite \<Longrightarrow> starfun X N \<approx> star_of L) \<Longrightarrow> X \<longlonglongrightarrow>\<^sub>N\<^sub>S L"
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
    48
  by (simp add: NSLIMSEQ_def)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    49
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
    50
lemma NSLIMSEQ_D: "X \<longlonglongrightarrow>\<^sub>N\<^sub>S L \<Longrightarrow> N \<in> HNatInfinite \<Longrightarrow> starfun X N \<approx> star_of L"
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
    51
  by (simp add: NSLIMSEQ_def)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    52
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
    53
lemma NSLIMSEQ_const: "(\<lambda>n. k) \<longlonglongrightarrow>\<^sub>N\<^sub>S k"
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
    54
  by (simp add: NSLIMSEQ_def)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    55
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
    56
lemma NSLIMSEQ_add: "X \<longlonglongrightarrow>\<^sub>N\<^sub>S a \<Longrightarrow> Y \<longlonglongrightarrow>\<^sub>N\<^sub>S b \<Longrightarrow> (\<lambda>n. X n + Y n) \<longlonglongrightarrow>\<^sub>N\<^sub>S a + b"
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
    57
  by (auto intro: approx_add simp add: NSLIMSEQ_def)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    58
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
    59
lemma NSLIMSEQ_add_const: "f \<longlonglongrightarrow>\<^sub>N\<^sub>S a \<Longrightarrow> (\<lambda>n. f n + b) \<longlonglongrightarrow>\<^sub>N\<^sub>S a + b"
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
    60
  by (simp only: NSLIMSEQ_add NSLIMSEQ_const)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    61
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
    62
lemma NSLIMSEQ_mult: "X \<longlonglongrightarrow>\<^sub>N\<^sub>S a \<Longrightarrow> Y \<longlonglongrightarrow>\<^sub>N\<^sub>S b \<Longrightarrow> (\<lambda>n. X n * Y n) \<longlonglongrightarrow>\<^sub>N\<^sub>S a * b"
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
    63
  for a b :: "'a::real_normed_algebra"
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
    64
  by (auto intro!: approx_mult_HFinite simp add: NSLIMSEQ_def)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    65
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
    66
lemma NSLIMSEQ_minus: "X \<longlonglongrightarrow>\<^sub>N\<^sub>S a \<Longrightarrow> (\<lambda>n. - X n) \<longlonglongrightarrow>\<^sub>N\<^sub>S - a"
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
    67
  by (auto simp add: NSLIMSEQ_def)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    68
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
    69
lemma NSLIMSEQ_minus_cancel: "(\<lambda>n. - X n) \<longlonglongrightarrow>\<^sub>N\<^sub>S -a \<Longrightarrow> X \<longlonglongrightarrow>\<^sub>N\<^sub>S a"
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
    70
  by (drule NSLIMSEQ_minus) simp
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    71
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
    72
lemma NSLIMSEQ_diff: "X \<longlonglongrightarrow>\<^sub>N\<^sub>S a \<Longrightarrow> Y \<longlonglongrightarrow>\<^sub>N\<^sub>S b \<Longrightarrow> (\<lambda>n. X n - Y n) \<longlonglongrightarrow>\<^sub>N\<^sub>S a - b"
54230
b1d955791529 more simplification rules on unary and binary minus
haftmann
parents: 51525
diff changeset
    73
  using NSLIMSEQ_add [of X a "- Y" "- b"] by (simp add: NSLIMSEQ_minus fun_Compl_def)
b1d955791529 more simplification rules on unary and binary minus
haftmann
parents: 51525
diff changeset
    74
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    75
(* FIXME: delete *)
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
    76
lemma NSLIMSEQ_add_minus: "X \<longlonglongrightarrow>\<^sub>N\<^sub>S a \<Longrightarrow> Y \<longlonglongrightarrow>\<^sub>N\<^sub>S b \<Longrightarrow> (\<lambda>n. X n + - Y n) \<longlonglongrightarrow>\<^sub>N\<^sub>S a + - b"
54230
b1d955791529 more simplification rules on unary and binary minus
haftmann
parents: 51525
diff changeset
    77
  by (simp add: NSLIMSEQ_diff)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    78
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
    79
lemma NSLIMSEQ_diff_const: "f \<longlonglongrightarrow>\<^sub>N\<^sub>S a \<Longrightarrow> (\<lambda>n. f n - b) \<longlonglongrightarrow>\<^sub>N\<^sub>S a - b"
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
    80
  by (simp add: NSLIMSEQ_diff NSLIMSEQ_const)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    81
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
    82
lemma NSLIMSEQ_inverse: "X \<longlonglongrightarrow>\<^sub>N\<^sub>S a \<Longrightarrow> a \<noteq> 0 \<Longrightarrow> (\<lambda>n. inverse (X n)) \<longlonglongrightarrow>\<^sub>N\<^sub>S inverse a"
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
    83
  for a :: "'a::real_normed_div_algebra"
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
    84
  by (simp add: NSLIMSEQ_def star_of_approx_inverse)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    85
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
    86
lemma NSLIMSEQ_mult_inverse: "X \<longlonglongrightarrow>\<^sub>N\<^sub>S a \<Longrightarrow> Y \<longlonglongrightarrow>\<^sub>N\<^sub>S b \<Longrightarrow> b \<noteq> 0 \<Longrightarrow> (\<lambda>n. X n / Y n) \<longlonglongrightarrow>\<^sub>N\<^sub>S a / b"
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
    87
  for a b :: "'a::real_normed_field"
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
    88
  by (simp add: NSLIMSEQ_mult NSLIMSEQ_inverse divide_inverse)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    89
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    90
lemma starfun_hnorm: "\<And>x. hnorm (( *f* f) x) = ( *f* (\<lambda>x. norm (f x))) x"
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
    91
  by transfer simp
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    92
61970
6226261144d7 more symbols;
wenzelm
parents: 61969
diff changeset
    93
lemma NSLIMSEQ_norm: "X \<longlonglongrightarrow>\<^sub>N\<^sub>S a \<Longrightarrow> (\<lambda>n. norm (X n)) \<longlonglongrightarrow>\<^sub>N\<^sub>S norm a"
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
    94
  by (simp add: NSLIMSEQ_def starfun_hnorm [symmetric] approx_hnorm)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    95
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
    96
text \<open>Uniqueness of limit.\<close>
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
    97
lemma NSLIMSEQ_unique: "X \<longlonglongrightarrow>\<^sub>N\<^sub>S a \<Longrightarrow> X \<longlonglongrightarrow>\<^sub>N\<^sub>S b \<Longrightarrow> a = b"
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
    98
  apply (simp add: NSLIMSEQ_def)
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
    99
  apply (drule HNatInfinite_whn [THEN [2] bspec])+
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   100
  apply (auto dest: approx_trans3)
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   101
  done
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   102
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   103
lemma NSLIMSEQ_pow [rule_format]: "(X \<longlonglongrightarrow>\<^sub>N\<^sub>S a) \<longrightarrow> ((\<lambda>n. (X n) ^ m) \<longlonglongrightarrow>\<^sub>N\<^sub>S a ^ m)"
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   104
  for a :: "'a::{real_normed_algebra,power}"
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   105
  by (induct m) (auto intro: NSLIMSEQ_mult NSLIMSEQ_const)
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   106
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   107
text \<open>We can now try and derive a few properties of sequences,
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   108
  starting with the limit comparison property for sequences.\<close>
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   109
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   110
lemma NSLIMSEQ_le: "f \<longlonglongrightarrow>\<^sub>N\<^sub>S l \<Longrightarrow> g \<longlonglongrightarrow>\<^sub>N\<^sub>S m \<Longrightarrow> \<exists>N. \<forall>n \<ge> N. f n \<le> g n \<Longrightarrow> l \<le> m"
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   111
  for l m :: real
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   112
  apply (simp add: NSLIMSEQ_def, safe)
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   113
  apply (drule starfun_le_mono)
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   114
  apply (drule HNatInfinite_whn [THEN [2] bspec])+
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   115
  apply (drule_tac x = whn in spec)
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   116
  apply (drule bex_Infinitesimal_iff2 [THEN iffD2])+
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   117
  apply clarify
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   118
  apply (auto intro: hypreal_of_real_le_add_Infininitesimal_cancel2)
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   119
  done
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   120
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   121
lemma NSLIMSEQ_le_const: "X \<longlonglongrightarrow>\<^sub>N\<^sub>S r \<Longrightarrow> \<forall>n. a \<le> X n \<Longrightarrow> a \<le> r"
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   122
  for a r :: real
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   123
  by (erule NSLIMSEQ_le [OF NSLIMSEQ_const]) auto
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   124
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   125
lemma NSLIMSEQ_le_const2: "X \<longlonglongrightarrow>\<^sub>N\<^sub>S r \<Longrightarrow> \<forall>n. X n \<le> a \<Longrightarrow> r \<le> a"
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   126
  for a r :: real
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   127
  by (erule NSLIMSEQ_le [OF _ NSLIMSEQ_const]) auto
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   128
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   129
text \<open>Shift a convergent series by 1:
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   130
  By the equivalence between Cauchiness and convergence and because
61975
b4b11391c676 isabelle update_cartouches -c -t;
wenzelm
parents: 61970
diff changeset
   131
  the successor of an infinite hypernatural is also infinite.\<close>
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   132
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   133
lemma NSLIMSEQ_Suc: "f \<longlonglongrightarrow>\<^sub>N\<^sub>S l \<Longrightarrow> (\<lambda>n. f(Suc n)) \<longlonglongrightarrow>\<^sub>N\<^sub>S l"
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   134
  apply (unfold NSLIMSEQ_def)
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   135
  apply safe
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   136
  apply (drule_tac x="N + 1" in bspec)
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   137
   apply (erule HNatInfinite_add)
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   138
  apply (simp add: starfun_shift_one)
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   139
  done
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   140
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   141
lemma NSLIMSEQ_imp_Suc: "(\<lambda>n. f(Suc n)) \<longlonglongrightarrow>\<^sub>N\<^sub>S l \<Longrightarrow> f \<longlonglongrightarrow>\<^sub>N\<^sub>S l"
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   142
  apply (unfold NSLIMSEQ_def)
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   143
  apply safe
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   144
  apply (drule_tac x="N - 1" in bspec)
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   145
   apply (erule Nats_1 [THEN [2] HNatInfinite_diff])
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   146
  apply (simp add: starfun_shift_one one_le_HNatInfinite)
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   147
  done
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   148
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   149
lemma NSLIMSEQ_Suc_iff: "(\<lambda>n. f (Suc n)) \<longlonglongrightarrow>\<^sub>N\<^sub>S l \<longleftrightarrow> f \<longlonglongrightarrow>\<^sub>N\<^sub>S l"
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   150
  by (blast intro: NSLIMSEQ_imp_Suc NSLIMSEQ_Suc)
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   151
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   152
61975
b4b11391c676 isabelle update_cartouches -c -t;
wenzelm
parents: 61970
diff changeset
   153
subsubsection \<open>Equivalence of @{term LIMSEQ} and @{term NSLIMSEQ}\<close>
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   154
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   155
lemma LIMSEQ_NSLIMSEQ:
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   156
  assumes X: "X \<longlonglongrightarrow> L"
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   157
  shows "X \<longlonglongrightarrow>\<^sub>N\<^sub>S L"
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   158
proof (rule NSLIMSEQ_I)
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   159
  fix N
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   160
  assume N: "N \<in> HNatInfinite"
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   161
  have "starfun X N - star_of L \<in> Infinitesimal"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   162
  proof (rule InfinitesimalI2)
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   163
    fix r :: real
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   164
    assume r: "0 < r"
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   165
    from LIMSEQ_D [OF X r] obtain no where "\<forall>n\<ge>no. norm (X n - L) < r" ..
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   166
    then have "\<forall>n\<ge>star_of no. hnorm (starfun X n - star_of L) < star_of r"
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   167
      by transfer
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   168
    then show "hnorm (starfun X N - star_of L) < star_of r"
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   169
      using N by (simp add: star_of_le_HNatInfinite)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   170
  qed
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   171
  then show "starfun X N \<approx> star_of L"
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   172
    by (simp only: approx_def)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   173
qed
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   174
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   175
lemma NSLIMSEQ_LIMSEQ:
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   176
  assumes X: "X \<longlonglongrightarrow>\<^sub>N\<^sub>S L"
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   177
  shows "X \<longlonglongrightarrow> L"
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   178
proof (rule LIMSEQ_I)
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   179
  fix r :: real
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   180
  assume r: "0 < r"
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   181
  have "\<exists>no. \<forall>n\<ge>no. hnorm (starfun X n - star_of L) < star_of r"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   182
  proof (intro exI allI impI)
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   183
    fix n
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   184
    assume "whn \<le> n"
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   185
    with HNatInfinite_whn have "n \<in> HNatInfinite"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   186
      by (rule HNatInfinite_upward_closed)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   187
    with X have "starfun X n \<approx> star_of L"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   188
      by (rule NSLIMSEQ_D)
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   189
    then have "starfun X n - star_of L \<in> Infinitesimal"
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   190
      by (simp only: approx_def)
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   191
    then show "hnorm (starfun X n - star_of L) < star_of r"
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   192
      using r by (rule InfinitesimalD2)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   193
  qed
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   194
  then show "\<exists>no. \<forall>n\<ge>no. norm (X n - L) < r"
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   195
    by transfer
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   196
qed
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   197
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   198
theorem LIMSEQ_NSLIMSEQ_iff: "f \<longlonglongrightarrow> L \<longleftrightarrow> f \<longlonglongrightarrow>\<^sub>N\<^sub>S L"
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   199
  by (blast intro: LIMSEQ_NSLIMSEQ NSLIMSEQ_LIMSEQ)
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   200
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   201
61975
b4b11391c676 isabelle update_cartouches -c -t;
wenzelm
parents: 61970
diff changeset
   202
subsubsection \<open>Derived theorems about @{term NSLIMSEQ}\<close>
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   203
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   204
text \<open>We prove the NS version from the standard one, since the NS proof
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   205
  seems more complicated than the standard one above!\<close>
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   206
lemma NSLIMSEQ_norm_zero: "(\<lambda>n. norm (X n)) \<longlonglongrightarrow>\<^sub>N\<^sub>S 0 \<longleftrightarrow> X \<longlonglongrightarrow>\<^sub>N\<^sub>S 0"
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   207
  by (simp add: LIMSEQ_NSLIMSEQ_iff [symmetric] tendsto_norm_zero_iff)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   208
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   209
lemma NSLIMSEQ_rabs_zero: "(\<lambda>n. \<bar>f n\<bar>) \<longlonglongrightarrow>\<^sub>N\<^sub>S 0 \<longleftrightarrow> f \<longlonglongrightarrow>\<^sub>N\<^sub>S (0::real)"
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   210
  by (simp add: LIMSEQ_NSLIMSEQ_iff [symmetric] tendsto_rabs_zero_iff)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   211
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   212
text \<open>Generalization to other limits.\<close>
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   213
lemma NSLIMSEQ_imp_rabs: "f \<longlonglongrightarrow>\<^sub>N\<^sub>S l \<Longrightarrow> (\<lambda>n. \<bar>f n\<bar>) \<longlonglongrightarrow>\<^sub>N\<^sub>S \<bar>l\<bar>"
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   214
  for l :: real
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   215
  by (simp add: NSLIMSEQ_def) (auto intro: approx_hrabs simp add: starfun_abs)
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   216
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   217
lemma NSLIMSEQ_inverse_zero: "\<forall>y::real. \<exists>N. \<forall>n \<ge> N. y < f n \<Longrightarrow> (\<lambda>n. inverse (f n)) \<longlonglongrightarrow>\<^sub>N\<^sub>S 0"
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   218
  by (simp add: LIMSEQ_NSLIMSEQ_iff [symmetric] LIMSEQ_inverse_zero)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   219
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   220
lemma NSLIMSEQ_inverse_real_of_nat: "(\<lambda>n. inverse (real (Suc n))) \<longlonglongrightarrow>\<^sub>N\<^sub>S 0"
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   221
  by (simp add: LIMSEQ_NSLIMSEQ_iff [symmetric] LIMSEQ_inverse_real_of_nat del: of_nat_Suc)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   222
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   223
lemma NSLIMSEQ_inverse_real_of_nat_add: "(\<lambda>n. r + inverse (real (Suc n))) \<longlonglongrightarrow>\<^sub>N\<^sub>S r"
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   224
  by (simp add: LIMSEQ_NSLIMSEQ_iff [symmetric] LIMSEQ_inverse_real_of_nat_add del: of_nat_Suc)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   225
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   226
lemma NSLIMSEQ_inverse_real_of_nat_add_minus: "(\<lambda>n. r + - inverse (real (Suc n))) \<longlonglongrightarrow>\<^sub>N\<^sub>S r"
54230
b1d955791529 more simplification rules on unary and binary minus
haftmann
parents: 51525
diff changeset
   227
  using LIMSEQ_inverse_real_of_nat_add_minus by (simp add: LIMSEQ_NSLIMSEQ_iff [symmetric])
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   228
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   229
lemma NSLIMSEQ_inverse_real_of_nat_add_minus_mult:
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   230
  "(\<lambda>n. r * (1 + - inverse (real (Suc n)))) \<longlonglongrightarrow>\<^sub>N\<^sub>S r"
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   231
  using LIMSEQ_inverse_real_of_nat_add_minus_mult
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   232
  by (simp add: LIMSEQ_NSLIMSEQ_iff [symmetric])
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   233
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   234
61975
b4b11391c676 isabelle update_cartouches -c -t;
wenzelm
parents: 61970
diff changeset
   235
subsection \<open>Convergence\<close>
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   236
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   237
lemma nslimI: "X \<longlonglongrightarrow>\<^sub>N\<^sub>S L \<Longrightarrow> nslim X = L"
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   238
  by (simp add: nslim_def) (blast intro: NSLIMSEQ_unique)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   239
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   240
lemma lim_nslim_iff: "lim X = nslim X"
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   241
  by (simp add: lim_def nslim_def LIMSEQ_NSLIMSEQ_iff)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   242
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   243
lemma NSconvergentD: "NSconvergent X \<Longrightarrow> \<exists>L. X \<longlonglongrightarrow>\<^sub>N\<^sub>S L"
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   244
  by (simp add: NSconvergent_def)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   245
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   246
lemma NSconvergentI: "X \<longlonglongrightarrow>\<^sub>N\<^sub>S L \<Longrightarrow> NSconvergent X"
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   247
  by (auto simp add: NSconvergent_def)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   248
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   249
lemma convergent_NSconvergent_iff: "convergent X = NSconvergent X"
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   250
  by (simp add: convergent_def NSconvergent_def LIMSEQ_NSLIMSEQ_iff)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   251
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   252
lemma NSconvergent_NSLIMSEQ_iff: "NSconvergent X \<longleftrightarrow> X \<longlonglongrightarrow>\<^sub>N\<^sub>S nslim X"
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   253
  by (auto intro: theI NSLIMSEQ_unique simp add: NSconvergent_def nslim_def)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   254
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   255
61975
b4b11391c676 isabelle update_cartouches -c -t;
wenzelm
parents: 61970
diff changeset
   256
subsection \<open>Bounded Monotonic Sequences\<close>
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   257
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   258
lemma NSBseqD: "NSBseq X \<Longrightarrow> N \<in> HNatInfinite \<Longrightarrow> ( *f* X) N \<in> HFinite"
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   259
  by (simp add: NSBseq_def)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   260
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   261
lemma Standard_subset_HFinite: "Standard \<subseteq> HFinite"
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   262
  by (auto simp: Standard_def)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   263
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   264
lemma NSBseqD2: "NSBseq X \<Longrightarrow> ( *f* X) N \<in> HFinite"
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   265
  apply (cases "N \<in> HNatInfinite")
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   266
   apply (erule (1) NSBseqD)
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   267
  apply (rule subsetD [OF Standard_subset_HFinite])
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   268
  apply (simp add: HNatInfinite_def Nats_eq_Standard)
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   269
  done
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   270
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   271
lemma NSBseqI: "\<forall>N \<in> HNatInfinite. ( *f* X) N \<in> HFinite \<Longrightarrow> NSBseq X"
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   272
  by (simp add: NSBseq_def)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   273
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   274
text \<open>The standard definition implies the nonstandard definition.\<close>
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   275
lemma Bseq_NSBseq: "Bseq X \<Longrightarrow> NSBseq X"
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   276
  unfolding NSBseq_def
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   277
proof safe
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   278
  assume X: "Bseq X"
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   279
  fix N
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   280
  assume N: "N \<in> HNatInfinite"
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   281
  from BseqD [OF X] obtain K where "\<forall>n. norm (X n) \<le> K"
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   282
    by fast
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   283
  then have "\<forall>N. hnorm (starfun X N) \<le> star_of K"
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   284
    by transfer
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   285
  then have "hnorm (starfun X N) \<le> star_of K"
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   286
    by simp
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   287
  also have "star_of K < star_of (K + 1)"
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   288
    by simp
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   289
  finally have "\<exists>x\<in>Reals. hnorm (starfun X N) < x"
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   290
    by (rule bexI) simp
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   291
  then show "starfun X N \<in> HFinite"
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   292
    by (simp add: HFinite_def)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   293
qed
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   294
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   295
text \<open>The nonstandard definition implies the standard definition.\<close>
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   296
lemma SReal_less_omega: "r \<in> \<real> \<Longrightarrow> r < \<omega>"
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   297
  using HInfinite_omega
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   298
  by (simp add: HInfinite_def) (simp add: order_less_imp_le)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   299
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   300
lemma NSBseq_Bseq: "NSBseq X \<Longrightarrow> Bseq X"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   301
proof (rule ccontr)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   302
  let ?n = "\<lambda>K. LEAST n. K < norm (X n)"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   303
  assume "NSBseq X"
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   304
  then have finite: "( *f* X) (( *f* ?n) \<omega>) \<in> HFinite"
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   305
    by (rule NSBseqD2)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   306
  assume "\<not> Bseq X"
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   307
  then have "\<forall>K>0. \<exists>n. K < norm (X n)"
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   308
    by (simp add: Bseq_def linorder_not_le)
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   309
  then have "\<forall>K>0. K < norm (X (?n K))"
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   310
    by (auto intro: LeastI_ex)
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   311
  then have "\<forall>K>0. K < hnorm (( *f* X) (( *f* ?n) K))"
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   312
    by transfer
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   313
  then have "\<omega> < hnorm (( *f* X) (( *f* ?n) \<omega>))"
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   314
    by simp
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   315
  then have "\<forall>r\<in>\<real>. r < hnorm (( *f* X) (( *f* ?n) \<omega>))"
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   316
    by (simp add: order_less_trans [OF SReal_less_omega])
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   317
  then have "( *f* X) (( *f* ?n) \<omega>) \<in> HInfinite"
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   318
    by (simp add: HInfinite_def)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   319
  with finite show "False"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   320
    by (simp add: HFinite_HInfinite_iff)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   321
qed
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   322
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   323
text \<open>Equivalence of nonstandard and standard definitions for a bounded sequence.\<close>
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   324
lemma Bseq_NSBseq_iff: "Bseq X = NSBseq X"
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   325
  by (blast intro!: NSBseq_Bseq Bseq_NSBseq)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   326
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   327
text \<open>A convergent sequence is bounded:
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   328
  Boundedness as a necessary condition for convergence.
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   329
  The nonstandard version has no existential, as usual.\<close>
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   330
lemma NSconvergent_NSBseq: "NSconvergent X \<Longrightarrow> NSBseq X"
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   331
  by (simp add: NSconvergent_def NSBseq_def NSLIMSEQ_def)
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   332
    (blast intro: HFinite_star_of approx_sym approx_HFinite)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   333
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   334
text \<open>Standard Version: easily now proved using equivalence of NS and
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   335
 standard definitions.\<close>
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   336
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   337
lemma convergent_Bseq: "convergent X \<Longrightarrow> Bseq X"
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   338
  for X :: "nat \<Rightarrow> 'b::real_normed_vector"
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   339
  by (simp add: NSconvergent_NSBseq convergent_NSconvergent_iff Bseq_NSBseq_iff)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   340
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   341
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   342
subsubsection \<open>Upper Bounds and Lubs of Bounded Sequences\<close>
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   343
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   344
lemma NSBseq_isUb: "NSBseq X \<Longrightarrow> \<exists>U::real. isUb UNIV {x. \<exists>n. X n = x} U"
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   345
  by (simp add: Bseq_NSBseq_iff [symmetric] Bseq_isUb)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   346
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   347
lemma NSBseq_isLub: "NSBseq X \<Longrightarrow> \<exists>U::real. isLub UNIV {x. \<exists>n. X n = x} U"
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   348
  by (simp add: Bseq_NSBseq_iff [symmetric] Bseq_isLub)
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   349
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   350
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   351
subsubsection \<open>A Bounded and Monotonic Sequence Converges\<close>
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   352
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   353
text \<open>The best of both worlds: Easier to prove this result as a standard
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   354
   theorem and then use equivalence to "transfer" it into the
61975
b4b11391c676 isabelle update_cartouches -c -t;
wenzelm
parents: 61970
diff changeset
   355
   equivalent nonstandard form if needed!\<close>
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   356
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   357
lemma Bmonoseq_NSLIMSEQ: "\<forall>n \<ge> m. X n = X m \<Longrightarrow> \<exists>L. X \<longlonglongrightarrow>\<^sub>N\<^sub>S L"
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   358
  by (auto dest!: Bmonoseq_LIMSEQ simp add: LIMSEQ_NSLIMSEQ_iff)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   359
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   360
lemma NSBseq_mono_NSconvergent: "NSBseq X \<Longrightarrow> \<forall>m. \<forall>n \<ge> m. X m \<le> X n \<Longrightarrow> NSconvergent X"
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   361
  for X :: "nat \<Rightarrow> real"
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   362
  by (auto intro: Bseq_mono_convergent
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   363
      simp: convergent_NSconvergent_iff [symmetric] Bseq_NSBseq_iff [symmetric])
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   364
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   365
61975
b4b11391c676 isabelle update_cartouches -c -t;
wenzelm
parents: 61970
diff changeset
   366
subsection \<open>Cauchy Sequences\<close>
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   367
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   368
lemma NSCauchyI:
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   369
  "(\<And>M N. M \<in> HNatInfinite \<Longrightarrow> N \<in> HNatInfinite \<Longrightarrow> starfun X M \<approx> starfun X N) \<Longrightarrow> NSCauchy X"
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   370
  by (simp add: NSCauchy_def)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   371
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   372
lemma NSCauchyD:
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   373
  "NSCauchy X \<Longrightarrow> M \<in> HNatInfinite \<Longrightarrow> N \<in> HNatInfinite \<Longrightarrow> starfun X M \<approx> starfun X N"
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   374
  by (simp add: NSCauchy_def)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   375
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   376
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   377
subsubsection \<open>Equivalence Between NS and Standard\<close>
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   378
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   379
lemma Cauchy_NSCauchy:
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   380
  assumes X: "Cauchy X"
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   381
  shows "NSCauchy X"
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   382
proof (rule NSCauchyI)
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   383
  fix M
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   384
  assume M: "M \<in> HNatInfinite"
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   385
  fix N
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   386
  assume N: "N \<in> HNatInfinite"
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   387
  have "starfun X M - starfun X N \<in> Infinitesimal"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   388
  proof (rule InfinitesimalI2)
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   389
    fix r :: real
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   390
    assume r: "0 < r"
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   391
    from CauchyD [OF X r] obtain k where "\<forall>m\<ge>k. \<forall>n\<ge>k. norm (X m - X n) < r" ..
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   392
    then have "\<forall>m\<ge>star_of k. \<forall>n\<ge>star_of k. hnorm (starfun X m - starfun X n) < star_of r"
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   393
      by transfer
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   394
    then show "hnorm (starfun X M - starfun X N) < star_of r"
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   395
      using M N by (simp add: star_of_le_HNatInfinite)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   396
  qed
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   397
  then show "starfun X M \<approx> starfun X N"
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   398
    by (simp only: approx_def)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   399
qed
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   400
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   401
lemma NSCauchy_Cauchy:
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   402
  assumes X: "NSCauchy X"
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   403
  shows "Cauchy X"
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   404
proof (rule CauchyI)
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   405
  fix r :: real
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   406
  assume r: "0 < r"
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   407
  have "\<exists>k. \<forall>m\<ge>k. \<forall>n\<ge>k. hnorm (starfun X m - starfun X n) < star_of r"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   408
  proof (intro exI allI impI)
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   409
    fix M
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   410
    assume "whn \<le> M"
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   411
    with HNatInfinite_whn have M: "M \<in> HNatInfinite"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   412
      by (rule HNatInfinite_upward_closed)
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   413
    fix N
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   414
    assume "whn \<le> N"
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   415
    with HNatInfinite_whn have N: "N \<in> HNatInfinite"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   416
      by (rule HNatInfinite_upward_closed)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   417
    from X M N have "starfun X M \<approx> starfun X N"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   418
      by (rule NSCauchyD)
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   419
    then have "starfun X M - starfun X N \<in> Infinitesimal"
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   420
      by (simp only: approx_def)
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   421
    then show "hnorm (starfun X M - starfun X N) < star_of r"
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   422
      using r by (rule InfinitesimalD2)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   423
  qed
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   424
  then show "\<exists>k. \<forall>m\<ge>k. \<forall>n\<ge>k. norm (X m - X n) < r"
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   425
    by transfer
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   426
qed
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   427
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   428
theorem NSCauchy_Cauchy_iff: "NSCauchy X = Cauchy X"
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   429
  by (blast intro!: NSCauchy_Cauchy Cauchy_NSCauchy)
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   430
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   431
61975
b4b11391c676 isabelle update_cartouches -c -t;
wenzelm
parents: 61970
diff changeset
   432
subsubsection \<open>Cauchy Sequences are Bounded\<close>
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   433
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   434
text \<open>A Cauchy sequence is bounded -- nonstandard version.\<close>
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   435
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   436
lemma NSCauchy_NSBseq: "NSCauchy X \<Longrightarrow> NSBseq X"
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   437
  by (simp add: Cauchy_Bseq Bseq_NSBseq_iff [symmetric] NSCauchy_Cauchy_iff)
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   438
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   439
61975
b4b11391c676 isabelle update_cartouches -c -t;
wenzelm
parents: 61970
diff changeset
   440
subsubsection \<open>Cauchy Sequences are Convergent\<close>
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   441
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   442
text \<open>Equivalence of Cauchy criterion and convergence:
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   443
  We will prove this using our NS formulation which provides a
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   444
  much easier proof than using the standard definition. We do not
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   445
  need to use properties of subsequences such as boundedness,
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   446
  monotonicity etc... Compare with Harrison's corresponding proof
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   447
  in HOL which is much longer and more complicated. Of course, we do
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   448
  not have problems which he encountered with guessing the right
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   449
  instantiations for his 'espsilon-delta' proof(s) in this case
61975
b4b11391c676 isabelle update_cartouches -c -t;
wenzelm
parents: 61970
diff changeset
   450
  since the NS formulations do not involve existential quantifiers.\<close>
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   451
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   452
lemma NSconvergent_NSCauchy: "NSconvergent X \<Longrightarrow> NSCauchy X"
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   453
  by (simp add: NSconvergent_def NSLIMSEQ_def NSCauchy_def) (auto intro: approx_trans2)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   454
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   455
lemma real_NSCauchy_NSconvergent: "NSCauchy X \<Longrightarrow> NSconvergent X"
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   456
  for X :: "nat \<Rightarrow> real"
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   457
  apply (simp add: NSconvergent_def NSLIMSEQ_def)
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   458
  apply (frule NSCauchy_NSBseq)
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   459
  apply (simp add: NSBseq_def NSCauchy_def)
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   460
  apply (drule HNatInfinite_whn [THEN [2] bspec])
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   461
  apply (drule HNatInfinite_whn [THEN [2] bspec])
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   462
  apply (auto dest!: st_part_Ex simp add: SReal_iff)
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   463
  apply (blast intro: approx_trans3)
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   464
  done
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   465
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   466
lemma NSCauchy_NSconvergent: "NSCauchy X \<Longrightarrow> NSconvergent X"
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   467
  for X :: "nat \<Rightarrow> 'a::banach"
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   468
  apply (drule NSCauchy_Cauchy [THEN Cauchy_convergent])
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   469
  apply (erule convergent_NSconvergent_iff [THEN iffD1])
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   470
  done
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   471
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   472
lemma NSCauchy_NSconvergent_iff: "NSCauchy X = NSconvergent X"
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   473
  for X :: "nat \<Rightarrow> 'a::banach"
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   474
  by (fast intro: NSCauchy_NSconvergent NSconvergent_NSCauchy)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   475
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   476
61975
b4b11391c676 isabelle update_cartouches -c -t;
wenzelm
parents: 61970
diff changeset
   477
subsection \<open>Power Sequences\<close>
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   478
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   479
text \<open>The sequence @{term "x^n"} tends to 0 if @{term "0\<le>x"} and @{term
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   480
  "x<1"}.  Proof will use (NS) Cauchy equivalence for convergence and
61975
b4b11391c676 isabelle update_cartouches -c -t;
wenzelm
parents: 61970
diff changeset
   481
  also fact that bounded and monotonic sequence converges.\<close>
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   482
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   483
text \<open>We now use NS criterion to bring proof of theorem through.\<close>
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   484
lemma NSLIMSEQ_realpow_zero: "0 \<le> x \<Longrightarrow> x < 1 \<Longrightarrow> (\<lambda>n. x ^ n) \<longlonglongrightarrow>\<^sub>N\<^sub>S 0"
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   485
  for x :: real
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   486
  apply (simp add: NSLIMSEQ_def)
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   487
  apply (auto dest!: convergent_realpow simp add: convergent_NSconvergent_iff)
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   488
  apply (frule NSconvergentD)
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   489
  apply (auto simp add: NSLIMSEQ_def NSCauchy_NSconvergent_iff [symmetric] NSCauchy_def starfun_pow)
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   490
  apply (frule HNatInfinite_add_one)
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   491
  apply (drule bspec, assumption)
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   492
  apply (drule bspec, assumption)
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   493
  apply (drule_tac x = "N + 1" in bspec, assumption)
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   494
  apply (simp add: hyperpow_add)
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   495
  apply (drule approx_mult_subst_star_of, assumption)
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   496
  apply (drule approx_trans3, assumption)
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   497
  apply (auto simp del: star_of_mult simp add: star_of_mult [symmetric])
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   498
  done
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   499
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   500
lemma NSLIMSEQ_rabs_realpow_zero: "\<bar>c\<bar> < 1 \<Longrightarrow> (\<lambda>n. \<bar>c\<bar> ^ n) \<longlonglongrightarrow>\<^sub>N\<^sub>S 0"
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   501
  for c :: real
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   502
  by (simp add: LIMSEQ_rabs_realpow_zero LIMSEQ_NSLIMSEQ_iff [symmetric])
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   503
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   504
lemma NSLIMSEQ_rabs_realpow_zero2: "\<bar>c\<bar> < 1 \<Longrightarrow> (\<lambda>n. c ^ n) \<longlonglongrightarrow>\<^sub>N\<^sub>S 0"
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   505
  for c :: real
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   506
  by (simp add: LIMSEQ_rabs_realpow_zero2 LIMSEQ_NSLIMSEQ_iff [symmetric])
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   507
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   508
(***---------------------------------------------------------------
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   509
    Theorems proved by Harrison in HOL that we do not need
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   510
    in order to prove equivalence between Cauchy criterion
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   511
    and convergence:
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   512
 -- Show that every sequence contains a monotonic subsequence
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   513
Goal "\<exists>f. subseq f & monoseq (%n. s (f n))"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   514
 -- Show that a subsequence of a bounded sequence is bounded
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   515
Goal "Bseq X ==> Bseq (%n. X (f n))";
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   516
 -- Show we can take subsequential terms arbitrarily far
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   517
    up a sequence
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   518
Goal "subseq f ==> n \<le> f(n)";
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   519
Goal "subseq f ==> \<exists>n. N1 \<le> n & N2 \<le> f(n)";
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   520
 ---------------------------------------------------------------***)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   521
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   522
end