src/HOL/Nonstandard_Analysis/HSEQ.thy
author Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
Mon, 12 Oct 2020 18:59:44 +0200
changeset 72458 b44e894796d5
parent 70228 2d5b122aa0ff
child 80914 d97fdabd9e2b
permissions -rw-r--r--
add reconstruction for the SMT solver veriT * * * Improved veriT reconstruction
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
67443
3abf6a722518 standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents: 67006
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_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
    45
  by (simp add: NSLIMSEQ_def)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    46
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
    47
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
    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_const: "(\<lambda>n. k) \<longlonglongrightarrow>\<^sub>N\<^sub>S k"
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_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
    54
  by (auto intro: approx_add 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_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
    57
  by (simp only: NSLIMSEQ_add NSLIMSEQ_const)
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_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
    60
  for a b :: "'a::real_normed_algebra"
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
    61
  by (auto intro!: approx_mult_HFinite simp add: NSLIMSEQ_def)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    62
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
    63
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
    64
  by (auto 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_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
    67
  by (drule NSLIMSEQ_minus) simp
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_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
    70
  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
    71
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
    72
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
    73
  by (simp add: NSLIMSEQ_diff NSLIMSEQ_const)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    74
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
    75
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
    76
  for a :: "'a::real_normed_div_algebra"
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
    77
  by (simp add: NSLIMSEQ_def star_of_approx_inverse)
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_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
    80
  for a b :: "'a::real_normed_field"
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
    81
  by (simp add: NSLIMSEQ_mult NSLIMSEQ_inverse divide_inverse)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    82
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    83
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
    84
  by transfer simp
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    85
61970
6226261144d7 more symbols;
wenzelm
parents: 61969
diff changeset
    86
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
    87
  by (simp add: NSLIMSEQ_def starfun_hnorm [symmetric] approx_hnorm)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    88
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
    89
text \<open>Uniqueness of limit.\<close>
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
    90
lemma NSLIMSEQ_unique: "X \<longlonglongrightarrow>\<^sub>N\<^sub>S a \<Longrightarrow> X \<longlonglongrightarrow>\<^sub>N\<^sub>S b \<Longrightarrow> a = b"
70228
2d5b122aa0ff De-applying and combining lemmas to make structured proofs
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
    91
  unfolding NSLIMSEQ_def
2d5b122aa0ff De-applying and combining lemmas to make structured proofs
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
    92
  using HNatInfinite_whn approx_trans3 star_of_approx_iff by blast
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    93
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
    94
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
    95
  for a :: "'a::{real_normed_algebra,power}"
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
    96
  by (induct m) (auto intro: NSLIMSEQ_mult NSLIMSEQ_const)
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
    97
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
    98
text \<open>We can now try and derive a few properties of sequences,
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
    99
  starting with the limit comparison property for sequences.\<close>
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   100
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   101
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
   102
  for l m :: real
70228
2d5b122aa0ff De-applying and combining lemmas to make structured proofs
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   103
  unfolding NSLIMSEQ_def
2d5b122aa0ff De-applying and combining lemmas to make structured proofs
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   104
  by (metis HNatInfinite_whn bex_Infinitesimal_iff2 hypnat_of_nat_le_whn hypreal_of_real_le_add_Infininitesimal_cancel2 starfun_le_mono)
2d5b122aa0ff De-applying and combining lemmas to make structured proofs
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   105
 
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   106
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
   107
  for a r :: real
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   108
  by (erule NSLIMSEQ_le [OF NSLIMSEQ_const]) auto
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_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
   111
  for a r :: real
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   112
  by (erule NSLIMSEQ_le [OF _ NSLIMSEQ_const]) auto
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   113
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   114
text \<open>Shift a convergent series by 1:
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   115
  By the equivalence between Cauchiness and convergence and because
61975
b4b11391c676 isabelle update_cartouches -c -t;
wenzelm
parents: 61970
diff changeset
   116
  the successor of an infinite hypernatural is also infinite.\<close>
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   117
70228
2d5b122aa0ff De-applying and combining lemmas to make structured proofs
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   118
lemma NSLIMSEQ_Suc_iff: "((\<lambda>n. f (Suc n)) \<longlonglongrightarrow>\<^sub>N\<^sub>S l) \<longleftrightarrow> (f \<longlonglongrightarrow>\<^sub>N\<^sub>S l)"
2d5b122aa0ff De-applying and combining lemmas to make structured proofs
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   119
proof
2d5b122aa0ff De-applying and combining lemmas to make structured proofs
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   120
  assume *: "f \<longlonglongrightarrow>\<^sub>N\<^sub>S l"
2d5b122aa0ff De-applying and combining lemmas to make structured proofs
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   121
  show "(\<lambda>n. f(Suc n)) \<longlonglongrightarrow>\<^sub>N\<^sub>S l"
2d5b122aa0ff De-applying and combining lemmas to make structured proofs
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   122
  proof (rule NSLIMSEQ_I)
2d5b122aa0ff De-applying and combining lemmas to make structured proofs
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   123
    fix N
2d5b122aa0ff De-applying and combining lemmas to make structured proofs
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   124
    assume "N \<in> HNatInfinite"
2d5b122aa0ff De-applying and combining lemmas to make structured proofs
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   125
    then have "(*f* f) (N + 1) \<approx> star_of l"
2d5b122aa0ff De-applying and combining lemmas to make structured proofs
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   126
      by (simp add: HNatInfinite_add NSLIMSEQ_D *)
2d5b122aa0ff De-applying and combining lemmas to make structured proofs
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   127
    then show "(*f* (\<lambda>n. f (Suc n))) N \<approx> star_of l"
2d5b122aa0ff De-applying and combining lemmas to make structured proofs
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   128
      by (simp add: starfun_shift_one)
2d5b122aa0ff De-applying and combining lemmas to make structured proofs
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   129
  qed
2d5b122aa0ff De-applying and combining lemmas to make structured proofs
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   130
next
2d5b122aa0ff De-applying and combining lemmas to make structured proofs
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   131
  assume *: "(\<lambda>n. f(Suc n)) \<longlonglongrightarrow>\<^sub>N\<^sub>S l"
2d5b122aa0ff De-applying and combining lemmas to make structured proofs
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   132
  show "f \<longlonglongrightarrow>\<^sub>N\<^sub>S l"
2d5b122aa0ff De-applying and combining lemmas to make structured proofs
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   133
  proof (rule NSLIMSEQ_I)
2d5b122aa0ff De-applying and combining lemmas to make structured proofs
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   134
    fix N
2d5b122aa0ff De-applying and combining lemmas to make structured proofs
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   135
    assume "N \<in> HNatInfinite"
2d5b122aa0ff De-applying and combining lemmas to make structured proofs
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   136
    then have "(*f* (\<lambda>n. f (Suc n))) (N - 1) \<approx> star_of l"
2d5b122aa0ff De-applying and combining lemmas to make structured proofs
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   137
      using * by (simp add: HNatInfinite_diff NSLIMSEQ_D)
2d5b122aa0ff De-applying and combining lemmas to make structured proofs
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   138
    then show "(*f* f) N \<approx> star_of l"
2d5b122aa0ff De-applying and combining lemmas to make structured proofs
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   139
      by (simp add: \<open>N \<in> HNatInfinite\<close> one_le_HNatInfinite starfun_shift_one)
2d5b122aa0ff De-applying and combining lemmas to make structured proofs
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   140
  qed
2d5b122aa0ff De-applying and combining lemmas to make structured proofs
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   141
qed
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   142
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   143
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 68614
diff changeset
   144
subsubsection \<open>Equivalence of \<^term>\<open>LIMSEQ\<close> and \<^term>\<open>NSLIMSEQ\<close>\<close>
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   145
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   146
lemma LIMSEQ_NSLIMSEQ:
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   147
  assumes X: "X \<longlonglongrightarrow> L"
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   148
  shows "X \<longlonglongrightarrow>\<^sub>N\<^sub>S L"
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   149
proof (rule NSLIMSEQ_I)
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   150
  fix N
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   151
  assume N: "N \<in> HNatInfinite"
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   152
  have "starfun X N - star_of L \<in> Infinitesimal"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   153
  proof (rule InfinitesimalI2)
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   154
    fix r :: real
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   155
    assume r: "0 < r"
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   156
    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
   157
    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
   158
      by transfer
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   159
    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
   160
      using N by (simp add: star_of_le_HNatInfinite)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   161
  qed
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   162
  then show "starfun X N \<approx> star_of L"
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   163
    by (simp only: approx_def)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   164
qed
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   165
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   166
lemma NSLIMSEQ_LIMSEQ:
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   167
  assumes X: "X \<longlonglongrightarrow>\<^sub>N\<^sub>S L"
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   168
  shows "X \<longlonglongrightarrow> L"
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   169
proof (rule LIMSEQ_I)
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   170
  fix r :: real
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   171
  assume r: "0 < r"
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   172
  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
   173
  proof (intro exI allI impI)
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   174
    fix n
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   175
    assume "whn \<le> n"
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   176
    with HNatInfinite_whn have "n \<in> HNatInfinite"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   177
      by (rule HNatInfinite_upward_closed)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   178
    with X have "starfun X n \<approx> star_of L"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   179
      by (rule NSLIMSEQ_D)
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   180
    then have "starfun X n - star_of L \<in> Infinitesimal"
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   181
      by (simp only: approx_def)
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   182
    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
   183
      using r by (rule InfinitesimalD2)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   184
  qed
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   185
  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
   186
    by transfer
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   187
qed
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   188
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   189
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
   190
  by (blast intro: LIMSEQ_NSLIMSEQ NSLIMSEQ_LIMSEQ)
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   191
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   192
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 68614
diff changeset
   193
subsubsection \<open>Derived theorems about \<^term>\<open>NSLIMSEQ\<close>\<close>
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   194
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   195
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
   196
  seems more complicated than the standard one above!\<close>
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   197
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
   198
  by (simp add: LIMSEQ_NSLIMSEQ_iff [symmetric] tendsto_norm_zero_iff)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   199
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   200
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
   201
  by (simp add: LIMSEQ_NSLIMSEQ_iff [symmetric] tendsto_rabs_zero_iff)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   202
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   203
text \<open>Generalization to other limits.\<close>
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   204
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
   205
  for l :: real
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   206
  by (simp add: NSLIMSEQ_def) (auto intro: approx_hrabs simp add: starfun_abs)
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   207
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   208
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
   209
  by (simp add: LIMSEQ_NSLIMSEQ_iff [symmetric] LIMSEQ_inverse_zero)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   210
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   211
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
   212
  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
   213
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   214
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
   215
  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
   216
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   217
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
   218
  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
   219
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   220
lemma NSLIMSEQ_inverse_real_of_nat_add_minus_mult:
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   221
  "(\<lambda>n. r * (1 + - inverse (real (Suc n)))) \<longlonglongrightarrow>\<^sub>N\<^sub>S r"
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   222
  using LIMSEQ_inverse_real_of_nat_add_minus_mult
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   223
  by (simp add: LIMSEQ_NSLIMSEQ_iff [symmetric])
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   224
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   225
61975
b4b11391c676 isabelle update_cartouches -c -t;
wenzelm
parents: 61970
diff changeset
   226
subsection \<open>Convergence\<close>
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   227
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   228
lemma nslimI: "X \<longlonglongrightarrow>\<^sub>N\<^sub>S L \<Longrightarrow> nslim X = L"
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   229
  by (simp add: nslim_def) (blast intro: NSLIMSEQ_unique)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   230
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   231
lemma lim_nslim_iff: "lim X = nslim X"
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   232
  by (simp add: lim_def nslim_def LIMSEQ_NSLIMSEQ_iff)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   233
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   234
lemma NSconvergentD: "NSconvergent X \<Longrightarrow> \<exists>L. X \<longlonglongrightarrow>\<^sub>N\<^sub>S L"
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   235
  by (simp add: NSconvergent_def)
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 NSconvergentI: "X \<longlonglongrightarrow>\<^sub>N\<^sub>S L \<Longrightarrow> NSconvergent X"
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   238
  by (auto simp add: NSconvergent_def)
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 convergent_NSconvergent_iff: "convergent X = NSconvergent X"
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   241
  by (simp add: convergent_def NSconvergent_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 NSconvergent_NSLIMSEQ_iff: "NSconvergent X \<longleftrightarrow> X \<longlonglongrightarrow>\<^sub>N\<^sub>S nslim X"
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   244
  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
   245
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   246
61975
b4b11391c676 isabelle update_cartouches -c -t;
wenzelm
parents: 61970
diff changeset
   247
subsection \<open>Bounded Monotonic Sequences\<close>
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   248
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   249
lemma NSBseqD: "NSBseq X \<Longrightarrow> N \<in> HNatInfinite \<Longrightarrow> ( *f* X) N \<in> HFinite"
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   250
  by (simp add: NSBseq_def)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   251
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   252
lemma Standard_subset_HFinite: "Standard \<subseteq> HFinite"
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   253
  by (auto simp: Standard_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
lemma NSBseqD2: "NSBseq X \<Longrightarrow> ( *f* X) N \<in> HFinite"
70228
2d5b122aa0ff De-applying and combining lemmas to make structured proofs
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   256
  using HNatInfinite_def NSBseq_def Nats_eq_Standard Standard_starfun Standard_subset_HFinite by blast
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 NSBseqI: "\<forall>N \<in> HNatInfinite. ( *f* X) N \<in> HFinite \<Longrightarrow> NSBseq X"
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
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   261
text \<open>The standard definition implies the nonstandard definition.\<close>
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   262
lemma Bseq_NSBseq: "Bseq X \<Longrightarrow> NSBseq X"
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   263
  unfolding NSBseq_def
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   264
proof safe
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   265
  assume X: "Bseq X"
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   266
  fix N
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   267
  assume N: "N \<in> HNatInfinite"
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   268
  from BseqD [OF X] obtain K where "\<forall>n. norm (X n) \<le> K"
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   269
    by fast
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   270
  then have "\<forall>N. hnorm (starfun X N) \<le> star_of K"
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   271
    by transfer
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   272
  then have "hnorm (starfun X N) \<le> star_of K"
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   273
    by simp
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   274
  also have "star_of K < star_of (K + 1)"
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   275
    by simp
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   276
  finally have "\<exists>x\<in>Reals. hnorm (starfun X N) < x"
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   277
    by (rule bexI) simp
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   278
  then show "starfun X N \<in> HFinite"
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   279
    by (simp add: HFinite_def)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   280
qed
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   281
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   282
text \<open>The nonstandard definition implies the standard definition.\<close>
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   283
lemma SReal_less_omega: "r \<in> \<real> \<Longrightarrow> r < \<omega>"
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   284
  using HInfinite_omega
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   285
  by (simp add: HInfinite_def) (simp add: order_less_imp_le)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   286
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   287
lemma NSBseq_Bseq: "NSBseq X \<Longrightarrow> Bseq X"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   288
proof (rule ccontr)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   289
  let ?n = "\<lambda>K. LEAST n. K < norm (X n)"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   290
  assume "NSBseq X"
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   291
  then have finite: "( *f* X) (( *f* ?n) \<omega>) \<in> HFinite"
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   292
    by (rule NSBseqD2)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   293
  assume "\<not> Bseq X"
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   294
  then have "\<forall>K>0. \<exists>n. K < norm (X n)"
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   295
    by (simp add: Bseq_def linorder_not_le)
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   296
  then have "\<forall>K>0. K < norm (X (?n K))"
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   297
    by (auto intro: LeastI_ex)
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   298
  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
   299
    by transfer
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   300
  then have "\<omega> < hnorm (( *f* X) (( *f* ?n) \<omega>))"
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   301
    by simp
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   302
  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
   303
    by (simp add: order_less_trans [OF SReal_less_omega])
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   304
  then have "( *f* X) (( *f* ?n) \<omega>) \<in> HInfinite"
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   305
    by (simp add: HInfinite_def)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   306
  with finite show "False"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   307
    by (simp add: HFinite_HInfinite_iff)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   308
qed
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   309
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   310
text \<open>Equivalence of nonstandard and standard definitions for a bounded sequence.\<close>
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   311
lemma Bseq_NSBseq_iff: "Bseq X = NSBseq X"
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   312
  by (blast intro!: NSBseq_Bseq Bseq_NSBseq)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   313
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   314
text \<open>A convergent sequence is bounded:
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   315
  Boundedness as a necessary condition for convergence.
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   316
  The nonstandard version has no existential, as usual.\<close>
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   317
lemma NSconvergent_NSBseq: "NSconvergent X \<Longrightarrow> NSBseq X"
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   318
  by (simp add: NSconvergent_def NSBseq_def NSLIMSEQ_def)
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   319
    (blast intro: HFinite_star_of approx_sym approx_HFinite)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   320
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   321
text \<open>Standard Version: easily now proved using equivalence of NS and
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   322
 standard definitions.\<close>
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   323
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   324
lemma convergent_Bseq: "convergent X \<Longrightarrow> Bseq X"
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   325
  for X :: "nat \<Rightarrow> 'b::real_normed_vector"
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   326
  by (simp add: NSconvergent_NSBseq convergent_NSconvergent_iff Bseq_NSBseq_iff)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   327
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   328
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   329
subsubsection \<open>Upper Bounds and Lubs of Bounded Sequences\<close>
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   330
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   331
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
   332
  by (simp add: Bseq_NSBseq_iff [symmetric] Bseq_isUb)
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
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
   335
  by (simp add: Bseq_NSBseq_iff [symmetric] Bseq_isLub)
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   336
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   337
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   338
subsubsection \<open>A Bounded and Monotonic Sequence Converges\<close>
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   339
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   340
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
   341
   theorem and then use equivalence to "transfer" it into the
61975
b4b11391c676 isabelle update_cartouches -c -t;
wenzelm
parents: 61970
diff changeset
   342
   equivalent nonstandard form if needed!\<close>
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   343
68614
3cb44b0abc5c more de-applying
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   344
lemma Bmonoseq_NSLIMSEQ: "\<forall>\<^sub>F k in sequentially. X k = X m \<Longrightarrow> X \<longlonglongrightarrow>\<^sub>N\<^sub>S X m"
3cb44b0abc5c more de-applying
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   345
  unfolding LIMSEQ_NSLIMSEQ_iff[symmetric]
3cb44b0abc5c more de-applying
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   346
  by (simp add: eventually_mono eventually_nhds_x_imp_x filterlim_iff)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   347
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   348
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
   349
  for X :: "nat \<Rightarrow> real"
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   350
  by (auto intro: Bseq_mono_convergent
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   351
      simp: convergent_NSconvergent_iff [symmetric] Bseq_NSBseq_iff [symmetric])
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   352
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   353
61975
b4b11391c676 isabelle update_cartouches -c -t;
wenzelm
parents: 61970
diff changeset
   354
subsection \<open>Cauchy Sequences\<close>
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   355
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   356
lemma NSCauchyI:
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   357
  "(\<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
   358
  by (simp add: NSCauchy_def)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   359
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   360
lemma NSCauchyD:
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   361
  "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
   362
  by (simp add: NSCauchy_def)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   363
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   364
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   365
subsubsection \<open>Equivalence Between NS and Standard\<close>
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   366
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   367
lemma Cauchy_NSCauchy:
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   368
  assumes X: "Cauchy X"
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   369
  shows "NSCauchy X"
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   370
proof (rule NSCauchyI)
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   371
  fix M
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   372
  assume M: "M \<in> HNatInfinite"
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   373
  fix N
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   374
  assume N: "N \<in> HNatInfinite"
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   375
  have "starfun X M - starfun X N \<in> Infinitesimal"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   376
  proof (rule InfinitesimalI2)
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   377
    fix r :: real
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   378
    assume r: "0 < r"
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   379
    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
   380
    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
   381
      by transfer
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   382
    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
   383
      using M N by (simp add: star_of_le_HNatInfinite)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   384
  qed
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   385
  then show "starfun X M \<approx> starfun X N"
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   386
    by (simp only: approx_def)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   387
qed
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   388
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   389
lemma NSCauchy_Cauchy:
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   390
  assumes X: "NSCauchy X"
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   391
  shows "Cauchy X"
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   392
proof (rule CauchyI)
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   393
  fix r :: real
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   394
  assume r: "0 < r"
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   395
  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
   396
  proof (intro exI allI impI)
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   397
    fix M
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   398
    assume "whn \<le> M"
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   399
    with HNatInfinite_whn have M: "M \<in> HNatInfinite"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   400
      by (rule HNatInfinite_upward_closed)
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   401
    fix N
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   402
    assume "whn \<le> N"
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   403
    with HNatInfinite_whn have N: "N \<in> HNatInfinite"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   404
      by (rule HNatInfinite_upward_closed)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   405
    from X M N have "starfun X M \<approx> starfun X N"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   406
      by (rule NSCauchyD)
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   407
    then have "starfun X M - starfun X N \<in> Infinitesimal"
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   408
      by (simp only: approx_def)
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   409
    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
   410
      using r by (rule InfinitesimalD2)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   411
  qed
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   412
  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
   413
    by transfer
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   414
qed
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   415
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   416
theorem NSCauchy_Cauchy_iff: "NSCauchy X = Cauchy X"
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   417
  by (blast intro!: NSCauchy_Cauchy Cauchy_NSCauchy)
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   418
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   419
61975
b4b11391c676 isabelle update_cartouches -c -t;
wenzelm
parents: 61970
diff changeset
   420
subsubsection \<open>Cauchy Sequences are Bounded\<close>
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   421
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   422
text \<open>A Cauchy sequence is bounded -- nonstandard version.\<close>
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   423
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   424
lemma NSCauchy_NSBseq: "NSCauchy X \<Longrightarrow> NSBseq X"
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   425
  by (simp add: Cauchy_Bseq Bseq_NSBseq_iff [symmetric] NSCauchy_Cauchy_iff)
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   426
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   427
61975
b4b11391c676 isabelle update_cartouches -c -t;
wenzelm
parents: 61970
diff changeset
   428
subsubsection \<open>Cauchy Sequences are Convergent\<close>
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   429
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   430
text \<open>Equivalence of Cauchy criterion and convergence:
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   431
  We will prove this using our NS formulation which provides a
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   432
  much easier proof than using the standard definition. We do not
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   433
  need to use properties of subsequences such as boundedness,
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   434
  monotonicity etc... Compare with Harrison's corresponding proof
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   435
  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
   436
  not have problems which he encountered with guessing the right
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   437
  instantiations for his 'espsilon-delta' proof(s) in this case
61975
b4b11391c676 isabelle update_cartouches -c -t;
wenzelm
parents: 61970
diff changeset
   438
  since the NS formulations do not involve existential quantifiers.\<close>
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   439
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   440
lemma NSconvergent_NSCauchy: "NSconvergent X \<Longrightarrow> NSCauchy X"
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   441
  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
   442
68614
3cb44b0abc5c more de-applying
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   443
lemma real_NSCauchy_NSconvergent: 
3cb44b0abc5c more de-applying
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   444
  fixes X :: "nat \<Rightarrow> real"
3cb44b0abc5c more de-applying
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   445
  assumes "NSCauchy X" shows "NSconvergent X"
3cb44b0abc5c more de-applying
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   446
  unfolding NSconvergent_def NSLIMSEQ_def
3cb44b0abc5c more de-applying
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   447
proof -
3cb44b0abc5c more de-applying
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   448
  have "( *f* X) whn \<in> HFinite"
3cb44b0abc5c more de-applying
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   449
    by (simp add: NSBseqD2 NSCauchy_NSBseq assms)
3cb44b0abc5c more de-applying
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   450
  moreover have "\<forall>N\<in>HNatInfinite. ( *f* X) whn \<approx> ( *f* X) N"
3cb44b0abc5c more de-applying
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   451
    using HNatInfinite_whn NSCauchy_def assms by blast
3cb44b0abc5c more de-applying
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   452
  ultimately show "\<exists>L. \<forall>N\<in>HNatInfinite. ( *f* X) N \<approx> hypreal_of_real L"
3cb44b0abc5c more de-applying
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   453
    by (force dest!: st_part_Ex simp add: SReal_iff intro: approx_trans3)
3cb44b0abc5c more de-applying
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   454
qed
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   455
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   456
lemma NSCauchy_NSconvergent: "NSCauchy X \<Longrightarrow> NSconvergent X"
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   457
  for X :: "nat \<Rightarrow> 'a::banach"
68614
3cb44b0abc5c more de-applying
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   458
  using Cauchy_convergent NSCauchy_Cauchy convergent_NSconvergent_iff by auto
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   459
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   460
lemma NSCauchy_NSconvergent_iff: "NSCauchy X = NSconvergent X"
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   461
  for X :: "nat \<Rightarrow> 'a::banach"
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   462
  by (fast intro: NSCauchy_NSconvergent NSconvergent_NSCauchy)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   463
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   464
61975
b4b11391c676 isabelle update_cartouches -c -t;
wenzelm
parents: 61970
diff changeset
   465
subsection \<open>Power Sequences\<close>
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   466
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 68614
diff changeset
   467
text \<open>The sequence \<^term>\<open>x^n\<close> tends to 0 if \<^term>\<open>0\<le>x\<close> and \<^term>\<open>x<1\<close>.  Proof will use (NS) Cauchy equivalence for convergence and
61975
b4b11391c676 isabelle update_cartouches -c -t;
wenzelm
parents: 61970
diff changeset
   468
  also fact that bounded and monotonic sequence converges.\<close>
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   469
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   470
text \<open>We now use NS criterion to bring proof of theorem through.\<close>
68614
3cb44b0abc5c more de-applying
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   471
lemma NSLIMSEQ_realpow_zero:
3cb44b0abc5c more de-applying
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   472
  fixes x :: real
3cb44b0abc5c more de-applying
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   473
  assumes "0 \<le> x" "x < 1" shows "(\<lambda>n. x ^ n) \<longlonglongrightarrow>\<^sub>N\<^sub>S 0"
3cb44b0abc5c more de-applying
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   474
proof -
3cb44b0abc5c more de-applying
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   475
  have "( *f* (^) x) N \<approx> 0"
3cb44b0abc5c more de-applying
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   476
    if N: "N \<in> HNatInfinite" and x: "NSconvergent ((^) x)" for N
3cb44b0abc5c more de-applying
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   477
  proof -
3cb44b0abc5c more de-applying
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   478
    have "hypreal_of_real x pow N \<approx> hypreal_of_real x pow (N + 1)"
3cb44b0abc5c more de-applying
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   479
      by (metis HNatInfinite_add N NSCauchy_NSconvergent_iff NSCauchy_def starfun_pow x)
3cb44b0abc5c more de-applying
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   480
    moreover obtain L where L: "hypreal_of_real x pow N \<approx> hypreal_of_real L"
3cb44b0abc5c more de-applying
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   481
      using NSconvergentD [OF x] N by (auto simp add: NSLIMSEQ_def starfun_pow)
3cb44b0abc5c more de-applying
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   482
    ultimately have "hypreal_of_real x pow N \<approx> hypreal_of_real L * hypreal_of_real x"
3cb44b0abc5c more de-applying
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   483
      by (simp add: approx_mult_subst_star_of hyperpow_add)
3cb44b0abc5c more de-applying
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   484
    then have "hypreal_of_real L \<approx> hypreal_of_real L * hypreal_of_real x"
3cb44b0abc5c more de-applying
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   485
      using L approx_trans3 by blast
3cb44b0abc5c more de-applying
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   486
    then show ?thesis
3cb44b0abc5c more de-applying
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   487
      by (metis L \<open>x < 1\<close> hyperpow_def less_irrefl mult.right_neutral mult_left_cancel star_of_approx_iff star_of_mult star_of_simps(9) starfun2_star_of)
3cb44b0abc5c more de-applying
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   488
  qed
3cb44b0abc5c more de-applying
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   489
  with assms show ?thesis
3cb44b0abc5c more de-applying
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   490
    by (force dest!: convergent_realpow simp add: NSLIMSEQ_def convergent_NSconvergent_iff)
3cb44b0abc5c more de-applying
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   491
qed
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   492
68614
3cb44b0abc5c more de-applying
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   493
lemma NSLIMSEQ_abs_realpow_zero: "\<bar>c\<bar> < 1 \<Longrightarrow> (\<lambda>n. \<bar>c\<bar> ^ n) \<longlonglongrightarrow>\<^sub>N\<^sub>S 0"
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 63579
diff changeset
   494
  for c :: real
68614
3cb44b0abc5c more de-applying
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   495
  by (simp add: LIMSEQ_abs_realpow_zero LIMSEQ_NSLIMSEQ_iff [symmetric])
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   496
68614
3cb44b0abc5c more de-applying
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   497
lemma NSLIMSEQ_abs_realpow_zero2: "\<bar>c\<bar> < 1 \<Longrightarrow> (\<lambda>n. c ^ n) \<longlonglongrightarrow>\<^sub>N\<^sub>S 0"
3cb44b0abc5c more de-applying
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   498
  for c :: real
3cb44b0abc5c more de-applying
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   499
  by (simp add: LIMSEQ_abs_realpow_zero2 LIMSEQ_NSLIMSEQ_iff [symmetric])
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   500
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   501
end