src/HOL/Nonstandard_Analysis/NSCA.thy
author immler
Sun, 03 Nov 2019 21:46:46 -0500
changeset 71034 e0755162093f
parent 70217 1f04832cbfcf
permissions -rw-r--r--
replace approximation oracle by less ad-hoc @{computation}s
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
62479
716336f19aa9 clarified session;
wenzelm
parents: 61982
diff changeset
     1
(*  Title:      HOL/Nonstandard_Analysis/NSCA.thy
716336f19aa9 clarified session;
wenzelm
parents: 61982
diff changeset
     2
    Author:     Jacques D. Fleuriot
716336f19aa9 clarified session;
wenzelm
parents: 61982
diff changeset
     3
    Copyright:  2001, 2002 University of Edinburgh
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
     4
*)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
     5
61975
b4b11391c676 isabelle update_cartouches -c -t;
wenzelm
parents: 61070
diff changeset
     6
section\<open>Non-Standard Complex Analysis\<close>
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
     7
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
     8
theory NSCA
28952
15a4b2cf8c34 made repository layout more coherent with logical distribution structure; stripped some $Id$s
haftmann
parents: 28562
diff changeset
     9
imports NSComplex HTranscendental
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    10
begin
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    11
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    12
abbreviation
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    13
   (* standard complex numbers reagarded as an embedded subset of NS complex *)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    14
   SComplex  :: "hcomplex set" where
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    15
   "SComplex \<equiv> Standard"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    16
67443
3abf6a722518 standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents: 67091
diff changeset
    17
definition \<comment> \<open>standard part map\<close>
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    18
  stc :: "hcomplex => hcomplex" where 
67091
1393c2340eec more symbols;
wenzelm
parents: 63901
diff changeset
    19
  "stc x = (SOME r. x \<in> HFinite \<and> r\<in>SComplex \<and> r \<approx> x)"
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    20
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    21
61975
b4b11391c676 isabelle update_cartouches -c -t;
wenzelm
parents: 61070
diff changeset
    22
subsection\<open>Closure Laws for SComplex, the Standard Complex Numbers\<close>
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    23
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    24
lemma SComplex_minus_iff [simp]: "(-x \<in> SComplex) = (x \<in> SComplex)"
70216
40f19372a723 A bit of de-applying
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
    25
  using Standard_minus by fastforce
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    26
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    27
lemma SComplex_add_cancel:
70216
40f19372a723 A bit of de-applying
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
    28
  "\<lbrakk>x + y \<in> SComplex; y \<in> SComplex\<rbrakk> \<Longrightarrow> x \<in> SComplex"
40f19372a723 A bit of de-applying
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
    29
  using Standard_diff by fastforce
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    30
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    31
lemma SReal_hcmod_hcomplex_of_complex [simp]:
70216
40f19372a723 A bit of de-applying
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
    32
  "hcmod (hcomplex_of_complex r) \<in> \<real>"
40f19372a723 A bit of de-applying
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
    33
  by (simp add: Reals_eq_Standard)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    34
70216
40f19372a723 A bit of de-applying
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
    35
lemma SReal_hcmod_numeral: "hcmod (numeral w ::hcomplex) \<in> \<real>"
40f19372a723 A bit of de-applying
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
    36
  by simp
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    37
70216
40f19372a723 A bit of de-applying
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
    38
lemma SReal_hcmod_SComplex: "x \<in> SComplex \<Longrightarrow> hcmod x \<in> \<real>"
40f19372a723 A bit of de-applying
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
    39
  by (simp add: Reals_eq_Standard)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    40
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 37887
diff changeset
    41
lemma SComplex_divide_numeral:
70216
40f19372a723 A bit of de-applying
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
    42
  "r \<in> SComplex \<Longrightarrow> r/(numeral w::hcomplex) \<in> SComplex"
40f19372a723 A bit of de-applying
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
    43
  by simp
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    44
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    45
lemma SComplex_UNIV_complex:
70216
40f19372a723 A bit of de-applying
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
    46
  "{x. hcomplex_of_complex x \<in> SComplex} = (UNIV::complex set)"
40f19372a723 A bit of de-applying
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
    47
  by simp
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    48
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    49
lemma SComplex_iff: "(x \<in> SComplex) = (\<exists>y. x = hcomplex_of_complex y)"
70216
40f19372a723 A bit of de-applying
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
    50
  by (simp add: Standard_def image_def)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    51
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    52
lemma hcomplex_of_complex_image:
70216
40f19372a723 A bit of de-applying
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
    53
  "range hcomplex_of_complex = SComplex"
40f19372a723 A bit of de-applying
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
    54
  by (simp add: Standard_def)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    55
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    56
lemma inv_hcomplex_of_complex_image: "inv hcomplex_of_complex `SComplex = UNIV"
59658
0cc388370041 sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
paulson <lp15@cam.ac.uk>
parents: 58878
diff changeset
    57
by (auto simp add: Standard_def image_def) (metis inj_star_of inv_f_f)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    58
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    59
lemma SComplex_hcomplex_of_complex_image: 
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67443
diff changeset
    60
      "\<lbrakk>\<exists>x. x \<in> P; P \<le> SComplex\<rbrakk> \<Longrightarrow> \<exists>Q. P = hcomplex_of_complex ` Q"
70216
40f19372a723 A bit of de-applying
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
    61
  by (metis Standard_def subset_imageE)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    62
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    63
lemma SComplex_SReal_dense:
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67443
diff changeset
    64
     "\<lbrakk>x \<in> SComplex; y \<in> SComplex; hcmod x < hcmod y  
ce654b0e6d69 more symbols;
wenzelm
parents: 67443
diff changeset
    65
      \<rbrakk> \<Longrightarrow> \<exists>r \<in> Reals. hcmod x< r \<and> r < hcmod y"
70216
40f19372a723 A bit of de-applying
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
    66
  by (simp add: SReal_dense SReal_hcmod_SComplex)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    67
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    68
61975
b4b11391c676 isabelle update_cartouches -c -t;
wenzelm
parents: 61070
diff changeset
    69
subsection\<open>The Finite Elements form a Subring\<close>
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    70
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    71
lemma HFinite_hcmod_hcomplex_of_complex [simp]:
70216
40f19372a723 A bit of de-applying
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
    72
  "hcmod (hcomplex_of_complex r) \<in> HFinite"
40f19372a723 A bit of de-applying
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
    73
  by (auto intro!: SReal_subset_HFinite [THEN subsetD])
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    74
70217
1f04832cbfcf more tidying up
paulson <lp15@cam.ac.uk>
parents: 70216
diff changeset
    75
lemma HFinite_hcmod_iff [simp]: "hcmod x \<in> HFinite \<longleftrightarrow> x \<in> HFinite"
70216
40f19372a723 A bit of de-applying
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
    76
  by (simp add: HFinite_def)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    77
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    78
lemma HFinite_bounded_hcmod:
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67443
diff changeset
    79
  "\<lbrakk>x \<in> HFinite; y \<le> hcmod x; 0 \<le> y\<rbrakk> \<Longrightarrow> y \<in> HFinite"
70217
1f04832cbfcf more tidying up
paulson <lp15@cam.ac.uk>
parents: 70216
diff changeset
    80
  using HFinite_bounded HFinite_hcmod_iff by blast
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    81
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    82
61975
b4b11391c676 isabelle update_cartouches -c -t;
wenzelm
parents: 61070
diff changeset
    83
subsection\<open>The Complex Infinitesimals form a Subring\<close>
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    84
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    85
lemma Infinitesimal_hcmod_iff: 
70216
40f19372a723 A bit of de-applying
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
    86
  "(z \<in> Infinitesimal) = (hcmod z \<in> Infinitesimal)"
40f19372a723 A bit of de-applying
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
    87
  by (simp add: Infinitesimal_def)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    88
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    89
lemma HInfinite_hcmod_iff: "(z \<in> HInfinite) = (hcmod z \<in> HInfinite)"
70216
40f19372a723 A bit of de-applying
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
    90
  by (simp add: HInfinite_def)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    91
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    92
lemma HFinite_diff_Infinitesimal_hcmod:
70216
40f19372a723 A bit of de-applying
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
    93
  "x \<in> HFinite - Infinitesimal \<Longrightarrow> hcmod x \<in> HFinite - Infinitesimal"
70217
1f04832cbfcf more tidying up
paulson <lp15@cam.ac.uk>
parents: 70216
diff changeset
    94
  by (simp add: Infinitesimal_hcmod_iff)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    95
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    96
lemma hcmod_less_Infinitesimal:
70216
40f19372a723 A bit of de-applying
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
    97
  "\<lbrakk>e \<in> Infinitesimal; hcmod x < hcmod e\<rbrakk> \<Longrightarrow> x \<in> Infinitesimal"
40f19372a723 A bit of de-applying
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
    98
  by (auto elim: hrabs_less_Infinitesimal simp add: Infinitesimal_hcmod_iff)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    99
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   100
lemma hcmod_le_Infinitesimal:
70216
40f19372a723 A bit of de-applying
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   101
  "\<lbrakk>e \<in> Infinitesimal; hcmod x \<le> hcmod e\<rbrakk> \<Longrightarrow> x \<in> Infinitesimal"
40f19372a723 A bit of de-applying
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   102
  by (auto elim: hrabs_le_Infinitesimal simp add: Infinitesimal_hcmod_iff)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   103
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   104
61975
b4b11391c676 isabelle update_cartouches -c -t;
wenzelm
parents: 61070
diff changeset
   105
subsection\<open>The ``Infinitely Close'' Relation\<close>
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   106
70216
40f19372a723 A bit of de-applying
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   107
lemma approx_SComplex_mult_cancel_zero:
40f19372a723 A bit of de-applying
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   108
  "\<lbrakk>a \<in> SComplex; a \<noteq> 0; a*x \<approx> 0\<rbrakk> \<Longrightarrow> x \<approx> 0"
40f19372a723 A bit of de-applying
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   109
  by (metis Infinitesimal_mult_disj SComplex_iff mem_infmal_iff star_of_Infinitesimal_iff_0 star_zero_def)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   110
70216
40f19372a723 A bit of de-applying
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   111
lemma approx_mult_SComplex1: "\<lbrakk>a \<in> SComplex; x \<approx> 0\<rbrakk> \<Longrightarrow> x*a \<approx> 0"
40f19372a723 A bit of de-applying
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   112
  using SComplex_iff approx_mult_subst_star_of by fastforce
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   113
70216
40f19372a723 A bit of de-applying
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   114
lemma approx_mult_SComplex2: "\<lbrakk>a \<in> SComplex; x \<approx> 0\<rbrakk> \<Longrightarrow> a*x \<approx> 0"
40f19372a723 A bit of de-applying
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   115
  by (metis approx_mult_SComplex1 mult.commute)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   116
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   117
lemma approx_mult_SComplex_zero_cancel_iff [simp]:
70216
40f19372a723 A bit of de-applying
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   118
  "\<lbrakk>a \<in> SComplex; a \<noteq> 0\<rbrakk> \<Longrightarrow> (a*x \<approx> 0) = (x \<approx> 0)"
40f19372a723 A bit of de-applying
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   119
  using approx_SComplex_mult_cancel_zero approx_mult_SComplex2 by blast
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   120
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   121
lemma approx_SComplex_mult_cancel:
70216
40f19372a723 A bit of de-applying
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   122
     "\<lbrakk>a \<in> SComplex; a \<noteq> 0; a*w \<approx> a*z\<rbrakk> \<Longrightarrow> w \<approx> z"
40f19372a723 A bit of de-applying
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   123
  by (metis approx_SComplex_mult_cancel_zero approx_minus_iff right_diff_distrib)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   124
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   125
lemma approx_SComplex_mult_cancel_iff1 [simp]:
70216
40f19372a723 A bit of de-applying
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   126
     "\<lbrakk>a \<in> SComplex; a \<noteq> 0\<rbrakk> \<Longrightarrow> (a*w \<approx> a*z) = (w \<approx> z)"
40f19372a723 A bit of de-applying
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   127
  by (metis HFinite_star_of SComplex_iff approx_SComplex_mult_cancel approx_mult2)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   128
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   129
(* TODO: generalize following theorems: hcmod -> hnorm *)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   130
61982
3af5a06577c7 more symbols;
wenzelm
parents: 61981
diff changeset
   131
lemma approx_hcmod_approx_zero: "(x \<approx> y) = (hcmod (y - x) \<approx> 0)"
70216
40f19372a723 A bit of de-applying
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   132
  by (simp add: Infinitesimal_hcmod_iff approx_def hnorm_minus_commute)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   133
61982
3af5a06577c7 more symbols;
wenzelm
parents: 61981
diff changeset
   134
lemma approx_approx_zero_iff: "(x \<approx> 0) = (hcmod x \<approx> 0)"
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   135
by (simp add: approx_hcmod_approx_zero)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   136
61982
3af5a06577c7 more symbols;
wenzelm
parents: 61981
diff changeset
   137
lemma approx_minus_zero_cancel_iff [simp]: "(-x \<approx> 0) = (x \<approx> 0)"
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   138
by (simp add: approx_def)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   139
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   140
lemma Infinitesimal_hcmod_add_diff:
70216
40f19372a723 A bit of de-applying
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   141
     "u \<approx> 0 \<Longrightarrow> hcmod(x + u) - hcmod x \<in> Infinitesimal"
40f19372a723 A bit of de-applying
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   142
  by (metis add.commute add.left_neutral approx_add_right_iff approx_def approx_hnorm)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   143
70216
40f19372a723 A bit of de-applying
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   144
lemma approx_hcmod_add_hcmod: "u \<approx> 0 \<Longrightarrow> hcmod(x + u) \<approx> hcmod x"
40f19372a723 A bit of de-applying
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   145
  using Infinitesimal_hcmod_add_diff approx_def by blast
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   146
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   147
61975
b4b11391c676 isabelle update_cartouches -c -t;
wenzelm
parents: 61070
diff changeset
   148
subsection\<open>Zero is the Only Infinitesimal Complex Number\<close>
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   149
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   150
lemma Infinitesimal_less_SComplex:
70216
40f19372a723 A bit of de-applying
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   151
  "\<lbrakk>x \<in> SComplex; y \<in> Infinitesimal; 0 < hcmod x\<rbrakk> \<Longrightarrow> hcmod y < hcmod x"
40f19372a723 A bit of de-applying
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   152
  by (auto intro: Infinitesimal_less_SReal SReal_hcmod_SComplex simp add: Infinitesimal_hcmod_iff)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   153
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   154
lemma SComplex_Int_Infinitesimal_zero: "SComplex Int Infinitesimal = {0}"
70216
40f19372a723 A bit of de-applying
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   155
  by (auto simp add: Standard_def Infinitesimal_hcmod_iff)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   156
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   157
lemma SComplex_Infinitesimal_zero:
70216
40f19372a723 A bit of de-applying
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   158
  "\<lbrakk>x \<in> SComplex; x \<in> Infinitesimal\<rbrakk> \<Longrightarrow> x = 0"
40f19372a723 A bit of de-applying
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   159
  using SComplex_iff by auto
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   160
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   161
lemma SComplex_HFinite_diff_Infinitesimal:
70216
40f19372a723 A bit of de-applying
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   162
  "\<lbrakk>x \<in> SComplex; x \<noteq> 0\<rbrakk> \<Longrightarrow> x \<in> HFinite - Infinitesimal"
40f19372a723 A bit of de-applying
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   163
  using SComplex_iff by auto
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   164
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 37887
diff changeset
   165
lemma numeral_not_Infinitesimal [simp]:
70216
40f19372a723 A bit of de-applying
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   166
  "numeral w \<noteq> (0::hcomplex) \<Longrightarrow> (numeral w::hcomplex) \<notin> Infinitesimal"
40f19372a723 A bit of de-applying
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   167
  by (fast dest: Standard_numeral [THEN SComplex_Infinitesimal_zero])
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   168
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   169
lemma approx_SComplex_not_zero:
70216
40f19372a723 A bit of de-applying
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   170
  "\<lbrakk>y \<in> SComplex; x \<approx> y; y\<noteq> 0\<rbrakk> \<Longrightarrow> x \<noteq> 0"
40f19372a723 A bit of de-applying
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   171
  by (auto dest: SComplex_Infinitesimal_zero approx_sym [THEN mem_infmal_iff [THEN iffD2]])
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   172
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   173
lemma SComplex_approx_iff:
70216
40f19372a723 A bit of de-applying
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   174
  "\<lbrakk>x \<in> SComplex; y \<in> SComplex\<rbrakk> \<Longrightarrow> (x \<approx> y) = (x = y)"
40f19372a723 A bit of de-applying
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   175
  by (auto simp add: Standard_def)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   176
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   177
lemma approx_unique_complex:
70216
40f19372a723 A bit of de-applying
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   178
  "\<lbrakk>r \<in> SComplex; s \<in> SComplex; r \<approx> x; s \<approx> x\<rbrakk> \<Longrightarrow> r = s"
40f19372a723 A bit of de-applying
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   179
  by (blast intro: SComplex_approx_iff [THEN iffD1] approx_trans2)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   180
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67613
diff changeset
   181
subsection \<open>Properties of \<^term>\<open>hRe\<close>, \<^term>\<open>hIm\<close> and \<^term>\<open>HComplex\<close>\<close>
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   182
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   183
lemma abs_hRe_le_hcmod: "\<And>x. \<bar>hRe x\<bar> \<le> hcmod x"
70216
40f19372a723 A bit of de-applying
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   184
  by transfer (rule abs_Re_le_cmod)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   185
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   186
lemma abs_hIm_le_hcmod: "\<And>x. \<bar>hIm x\<bar> \<le> hcmod x"
70216
40f19372a723 A bit of de-applying
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   187
  by transfer (rule abs_Im_le_cmod)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   188
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   189
lemma Infinitesimal_hRe: "x \<in> Infinitesimal \<Longrightarrow> hRe x \<in> Infinitesimal"
70216
40f19372a723 A bit of de-applying
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   190
  using Infinitesimal_hcmod_iff abs_hRe_le_hcmod hrabs_le_Infinitesimal by blast
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   191
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   192
lemma Infinitesimal_hIm: "x \<in> Infinitesimal \<Longrightarrow> hIm x \<in> Infinitesimal"
70216
40f19372a723 A bit of de-applying
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   193
  using Infinitesimal_hcmod_iff abs_hIm_le_hcmod hrabs_le_Infinitesimal by blast
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   194
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   195
lemma Infinitesimal_HComplex:
70217
1f04832cbfcf more tidying up
paulson <lp15@cam.ac.uk>
parents: 70216
diff changeset
   196
  assumes x: "x \<in> Infinitesimal" and y: "y \<in> Infinitesimal"
1f04832cbfcf more tidying up
paulson <lp15@cam.ac.uk>
parents: 70216
diff changeset
   197
  shows "HComplex x y \<in> Infinitesimal"
1f04832cbfcf more tidying up
paulson <lp15@cam.ac.uk>
parents: 70216
diff changeset
   198
proof -
1f04832cbfcf more tidying up
paulson <lp15@cam.ac.uk>
parents: 70216
diff changeset
   199
  have "hcmod (HComplex 0 y) \<in> Infinitesimal"
1f04832cbfcf more tidying up
paulson <lp15@cam.ac.uk>
parents: 70216
diff changeset
   200
    by (simp add: hcmod_i y)
1f04832cbfcf more tidying up
paulson <lp15@cam.ac.uk>
parents: 70216
diff changeset
   201
  moreover have "hcmod (hcomplex_of_hypreal x) \<in> Infinitesimal" 
1f04832cbfcf more tidying up
paulson <lp15@cam.ac.uk>
parents: 70216
diff changeset
   202
    using Infinitesimal_hcmod_iff Infinitesimal_of_hypreal_iff x by blast
1f04832cbfcf more tidying up
paulson <lp15@cam.ac.uk>
parents: 70216
diff changeset
   203
  ultimately have "hcmod (HComplex x y) \<in> Infinitesimal"
1f04832cbfcf more tidying up
paulson <lp15@cam.ac.uk>
parents: 70216
diff changeset
   204
    by (metis Infinitesimal_add Infinitesimal_hcmod_iff add.right_neutral hcomplex_of_hypreal_add_HComplex)
1f04832cbfcf more tidying up
paulson <lp15@cam.ac.uk>
parents: 70216
diff changeset
   205
  then show ?thesis
1f04832cbfcf more tidying up
paulson <lp15@cam.ac.uk>
parents: 70216
diff changeset
   206
    by (simp add: Infinitesimal_hnorm_iff)
1f04832cbfcf more tidying up
paulson <lp15@cam.ac.uk>
parents: 70216
diff changeset
   207
qed
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   208
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   209
lemma hcomplex_Infinitesimal_iff:
70217
1f04832cbfcf more tidying up
paulson <lp15@cam.ac.uk>
parents: 70216
diff changeset
   210
  "(x \<in> Infinitesimal) \<longleftrightarrow> (hRe x \<in> Infinitesimal \<and> hIm x \<in> Infinitesimal)"
70216
40f19372a723 A bit of de-applying
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   211
  using Infinitesimal_HComplex Infinitesimal_hIm Infinitesimal_hRe by fastforce
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   212
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   213
lemma hRe_diff [simp]: "\<And>x y. hRe (x - y) = hRe x - hRe y"
70216
40f19372a723 A bit of de-applying
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   214
  by transfer simp
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   215
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   216
lemma hIm_diff [simp]: "\<And>x y. hIm (x - y) = hIm x - hIm y"
70216
40f19372a723 A bit of de-applying
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   217
  by transfer simp
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   218
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   219
lemma approx_hRe: "x \<approx> y \<Longrightarrow> hRe x \<approx> hRe y"
70216
40f19372a723 A bit of de-applying
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   220
  unfolding approx_def by (drule Infinitesimal_hRe) simp
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   221
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   222
lemma approx_hIm: "x \<approx> y \<Longrightarrow> hIm x \<approx> hIm y"
70216
40f19372a723 A bit of de-applying
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   223
  unfolding approx_def by (drule Infinitesimal_hIm) simp
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
lemma approx_HComplex:
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   226
  "\<lbrakk>a \<approx> b; c \<approx> d\<rbrakk> \<Longrightarrow> HComplex a c \<approx> HComplex b d"
70216
40f19372a723 A bit of de-applying
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   227
  unfolding approx_def by (simp add: Infinitesimal_HComplex)
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 hcomplex_approx_iff:
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   230
  "(x \<approx> y) = (hRe x \<approx> hRe y \<and> hIm x \<approx> hIm y)"
70216
40f19372a723 A bit of de-applying
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   231
  unfolding approx_def by (simp add: hcomplex_Infinitesimal_iff)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   232
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   233
lemma HFinite_hRe: "x \<in> HFinite \<Longrightarrow> hRe x \<in> HFinite"
70216
40f19372a723 A bit of de-applying
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   234
  using HFinite_bounded_hcmod abs_ge_zero abs_hRe_le_hcmod by blast
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   235
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   236
lemma HFinite_hIm: "x \<in> HFinite \<Longrightarrow> hIm x \<in> HFinite"
70216
40f19372a723 A bit of de-applying
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   237
  using HFinite_bounded_hcmod abs_ge_zero abs_hIm_le_hcmod by blast
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   238
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   239
lemma HFinite_HComplex:
70217
1f04832cbfcf more tidying up
paulson <lp15@cam.ac.uk>
parents: 70216
diff changeset
   240
  assumes "x \<in> HFinite" "y \<in> HFinite"
1f04832cbfcf more tidying up
paulson <lp15@cam.ac.uk>
parents: 70216
diff changeset
   241
  shows "HComplex x y \<in> HFinite"
1f04832cbfcf more tidying up
paulson <lp15@cam.ac.uk>
parents: 70216
diff changeset
   242
proof -
1f04832cbfcf more tidying up
paulson <lp15@cam.ac.uk>
parents: 70216
diff changeset
   243
  have "HComplex x 0 \<in> HFinite" "HComplex 0 y \<in> HFinite"
1f04832cbfcf more tidying up
paulson <lp15@cam.ac.uk>
parents: 70216
diff changeset
   244
    using HFinite_hcmod_iff assms hcmod_i by fastforce+
1f04832cbfcf more tidying up
paulson <lp15@cam.ac.uk>
parents: 70216
diff changeset
   245
  then have "HComplex x 0 + HComplex 0 y \<in> HFinite"
1f04832cbfcf more tidying up
paulson <lp15@cam.ac.uk>
parents: 70216
diff changeset
   246
    using HFinite_add by blast
1f04832cbfcf more tidying up
paulson <lp15@cam.ac.uk>
parents: 70216
diff changeset
   247
  then show ?thesis
1f04832cbfcf more tidying up
paulson <lp15@cam.ac.uk>
parents: 70216
diff changeset
   248
    by simp
1f04832cbfcf more tidying up
paulson <lp15@cam.ac.uk>
parents: 70216
diff changeset
   249
qed
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   250
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   251
lemma hcomplex_HFinite_iff:
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   252
  "(x \<in> HFinite) = (hRe x \<in> HFinite \<and> hIm x \<in> HFinite)"
70216
40f19372a723 A bit of de-applying
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   253
  using HFinite_HComplex HFinite_hIm HFinite_hRe by fastforce
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 hcomplex_HInfinite_iff:
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   256
  "(x \<in> HInfinite) = (hRe x \<in> HInfinite \<or> hIm x \<in> HInfinite)"
70216
40f19372a723 A bit of de-applying
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   257
  by (simp add: HInfinite_HFinite_iff hcomplex_HFinite_iff)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   258
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   259
lemma hcomplex_of_hypreal_approx_iff [simp]:
70216
40f19372a723 A bit of de-applying
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   260
  "(hcomplex_of_hypreal x \<approx> hcomplex_of_hypreal z) = (x \<approx> z)"
40f19372a723 A bit of de-applying
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   261
  by (simp add: hcomplex_approx_iff)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   262
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   263
(* Here we go - easy proof now!! *)
70217
1f04832cbfcf more tidying up
paulson <lp15@cam.ac.uk>
parents: 70216
diff changeset
   264
lemma stc_part_Ex:
1f04832cbfcf more tidying up
paulson <lp15@cam.ac.uk>
parents: 70216
diff changeset
   265
  assumes "x \<in> HFinite" 
1f04832cbfcf more tidying up
paulson <lp15@cam.ac.uk>
parents: 70216
diff changeset
   266
  shows "\<exists>t \<in> SComplex. x \<approx> t"
1f04832cbfcf more tidying up
paulson <lp15@cam.ac.uk>
parents: 70216
diff changeset
   267
proof -
1f04832cbfcf more tidying up
paulson <lp15@cam.ac.uk>
parents: 70216
diff changeset
   268
  let ?t = "HComplex (st (hRe x)) (st (hIm x))"
1f04832cbfcf more tidying up
paulson <lp15@cam.ac.uk>
parents: 70216
diff changeset
   269
  have "?t \<in> SComplex"
1f04832cbfcf more tidying up
paulson <lp15@cam.ac.uk>
parents: 70216
diff changeset
   270
    using HFinite_hIm HFinite_hRe Reals_eq_Standard assms st_SReal by auto
1f04832cbfcf more tidying up
paulson <lp15@cam.ac.uk>
parents: 70216
diff changeset
   271
  moreover have "x \<approx> ?t"
1f04832cbfcf more tidying up
paulson <lp15@cam.ac.uk>
parents: 70216
diff changeset
   272
    by (simp add: HFinite_hIm HFinite_hRe assms hcomplex_approx_iff st_HFinite st_eq_approx)
1f04832cbfcf more tidying up
paulson <lp15@cam.ac.uk>
parents: 70216
diff changeset
   273
  ultimately show ?thesis ..
1f04832cbfcf more tidying up
paulson <lp15@cam.ac.uk>
parents: 70216
diff changeset
   274
qed
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   275
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67443
diff changeset
   276
lemma stc_part_Ex1: "x \<in> HFinite \<Longrightarrow> \<exists>!t. t \<in> SComplex \<and> x \<approx> t"
70216
40f19372a723 A bit of de-applying
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   277
  using approx_sym approx_unique_complex stc_part_Ex by blast
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   278
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   279
61975
b4b11391c676 isabelle update_cartouches -c -t;
wenzelm
parents: 61070
diff changeset
   280
subsection\<open>Theorems About Monads\<close>
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   281
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67443
diff changeset
   282
lemma monad_zero_hcmod_iff: "(x \<in> monad 0) = (hcmod x \<in> monad 0)"
70216
40f19372a723 A bit of de-applying
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   283
  by (simp add: Infinitesimal_monad_zero_iff [symmetric] Infinitesimal_hcmod_iff)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   284
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   285
61975
b4b11391c676 isabelle update_cartouches -c -t;
wenzelm
parents: 61070
diff changeset
   286
subsection\<open>Theorems About Standard Part\<close>
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   287
70216
40f19372a723 A bit of de-applying
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   288
lemma stc_approx_self: "x \<in> HFinite \<Longrightarrow> stc x \<approx> x"
40f19372a723 A bit of de-applying
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   289
  unfolding stc_def
40f19372a723 A bit of de-applying
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   290
  by (metis (no_types, lifting) approx_reorient someI_ex stc_part_Ex1)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   291
70216
40f19372a723 A bit of de-applying
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   292
lemma stc_SComplex: "x \<in> HFinite \<Longrightarrow> stc x \<in> SComplex"
40f19372a723 A bit of de-applying
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   293
  unfolding stc_def
40f19372a723 A bit of de-applying
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   294
  by (metis (no_types, lifting) SComplex_iff approx_sym someI_ex stc_part_Ex)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   295
70216
40f19372a723 A bit of de-applying
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   296
lemma stc_HFinite: "x \<in> HFinite \<Longrightarrow> stc x \<in> HFinite"
40f19372a723 A bit of de-applying
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   297
  by (erule stc_SComplex [THEN Standard_subset_HFinite [THEN subsetD]])
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   298
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   299
lemma stc_unique: "\<lbrakk>y \<in> SComplex; y \<approx> x\<rbrakk> \<Longrightarrow> stc x = y"
70216
40f19372a723 A bit of de-applying
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   300
  by (metis SComplex_approx_iff SComplex_iff approx_monad_iff approx_star_of_HFinite stc_SComplex stc_approx_self)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   301
70216
40f19372a723 A bit of de-applying
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   302
lemma stc_SComplex_eq [simp]: "x \<in> SComplex \<Longrightarrow> stc x = x"
40f19372a723 A bit of de-applying
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   303
  by (simp add: stc_unique)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   304
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   305
lemma stc_eq_approx:
70216
40f19372a723 A bit of de-applying
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   306
  "\<lbrakk>x \<in> HFinite; y \<in> HFinite; stc x = stc y\<rbrakk> \<Longrightarrow> x \<approx> y"
40f19372a723 A bit of de-applying
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   307
  by (auto dest!: stc_approx_self elim!: approx_trans3)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   308
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   309
lemma approx_stc_eq:
70216
40f19372a723 A bit of de-applying
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   310
     "\<lbrakk>x \<in> HFinite; y \<in> HFinite; x \<approx> y\<rbrakk> \<Longrightarrow> stc x = stc y"
40f19372a723 A bit of de-applying
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   311
  by (metis approx_sym approx_trans3 stc_part_Ex1 stc_unique)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   312
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   313
lemma stc_eq_approx_iff:
70216
40f19372a723 A bit of de-applying
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   314
  "\<lbrakk>x \<in> HFinite; y \<in> HFinite\<rbrakk> \<Longrightarrow> (x \<approx> y) = (stc x = stc y)"
40f19372a723 A bit of de-applying
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   315
  by (blast intro: approx_stc_eq stc_eq_approx)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   316
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   317
lemma stc_Infinitesimal_add_SComplex:
70216
40f19372a723 A bit of de-applying
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   318
  "\<lbrakk>x \<in> SComplex; e \<in> Infinitesimal\<rbrakk> \<Longrightarrow> stc(x + e) = x"
40f19372a723 A bit of de-applying
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   319
  using Infinitesimal_add_approx_self stc_unique by blast
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   320
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   321
lemma stc_Infinitesimal_add_SComplex2:
70216
40f19372a723 A bit of de-applying
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   322
  "\<lbrakk>x \<in> SComplex; e \<in> Infinitesimal\<rbrakk> \<Longrightarrow> stc(e + x) = x"
40f19372a723 A bit of de-applying
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   323
  using Infinitesimal_add_approx_self2 stc_unique by blast
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   324
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   325
lemma HFinite_stc_Infinitesimal_add:
70216
40f19372a723 A bit of de-applying
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   326
  "x \<in> HFinite \<Longrightarrow> \<exists>e \<in> Infinitesimal. x = stc(x) + e"
40f19372a723 A bit of de-applying
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   327
  by (blast dest!: stc_approx_self [THEN approx_sym] bex_Infinitesimal_iff2 [THEN iffD2])
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   328
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   329
lemma stc_add:
70216
40f19372a723 A bit of de-applying
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   330
  "\<lbrakk>x \<in> HFinite; y \<in> HFinite\<rbrakk> \<Longrightarrow> stc (x + y) = stc(x) + stc(y)"
40f19372a723 A bit of de-applying
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   331
  by (simp add: stc_unique stc_SComplex stc_approx_self approx_add)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   332
70216
40f19372a723 A bit of de-applying
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   333
lemma stc_zero: "stc 0 = 0"
40f19372a723 A bit of de-applying
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   334
  by simp
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   335
70216
40f19372a723 A bit of de-applying
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   336
lemma stc_one: "stc 1 = 1"
40f19372a723 A bit of de-applying
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   337
  by simp
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   338
70216
40f19372a723 A bit of de-applying
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   339
lemma stc_minus: "y \<in> HFinite \<Longrightarrow> stc(-y) = -stc(y)"
40f19372a723 A bit of de-applying
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   340
  by (simp add: stc_unique stc_SComplex stc_approx_self approx_minus)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   341
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   342
lemma stc_diff: 
70216
40f19372a723 A bit of de-applying
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   343
  "\<lbrakk>x \<in> HFinite; y \<in> HFinite\<rbrakk> \<Longrightarrow> stc (x-y) = stc(x) - stc(y)"
40f19372a723 A bit of de-applying
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   344
  by (simp add: stc_unique stc_SComplex stc_approx_self approx_diff)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   345
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   346
lemma stc_mult:
70216
40f19372a723 A bit of de-applying
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   347
  "\<lbrakk>x \<in> HFinite; y \<in> HFinite\<rbrakk>  
40f19372a723 A bit of de-applying
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   348
               \<Longrightarrow> stc (x * y) = stc(x) * stc(y)"
40f19372a723 A bit of de-applying
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   349
  by (simp add: stc_unique stc_SComplex stc_approx_self approx_mult_HFinite)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   350
70216
40f19372a723 A bit of de-applying
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   351
lemma stc_Infinitesimal: "x \<in> Infinitesimal \<Longrightarrow> stc x = 0"
40f19372a723 A bit of de-applying
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   352
  by (simp add: stc_unique mem_infmal_iff)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   353
70216
40f19372a723 A bit of de-applying
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   354
lemma stc_not_Infinitesimal: "stc(x) \<noteq> 0 \<Longrightarrow> x \<notin> Infinitesimal"
40f19372a723 A bit of de-applying
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   355
  by (fast intro: stc_Infinitesimal)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   356
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   357
lemma stc_inverse:
70216
40f19372a723 A bit of de-applying
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   358
  "\<lbrakk>x \<in> HFinite; stc x \<noteq> 0\<rbrakk>  \<Longrightarrow> stc(inverse x) = inverse (stc x)"
70217
1f04832cbfcf more tidying up
paulson <lp15@cam.ac.uk>
parents: 70216
diff changeset
   359
  by (simp add: stc_unique stc_SComplex stc_approx_self approx_inverse stc_not_Infinitesimal)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   360
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   361
lemma stc_divide [simp]:
70216
40f19372a723 A bit of de-applying
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   362
  "\<lbrakk>x \<in> HFinite; y \<in> HFinite; stc y \<noteq> 0\<rbrakk>  
40f19372a723 A bit of de-applying
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   363
      \<Longrightarrow> stc(x/y) = (stc x) / (stc y)"
40f19372a723 A bit of de-applying
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   364
  by (simp add: divide_inverse stc_mult stc_not_Infinitesimal HFinite_inverse stc_inverse)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   365
70216
40f19372a723 A bit of de-applying
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   366
lemma stc_idempotent [simp]: "x \<in> HFinite \<Longrightarrow> stc(stc(x)) = stc(x)"
40f19372a723 A bit of de-applying
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   367
  by (blast intro: stc_HFinite stc_approx_self approx_stc_eq)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   368
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   369
lemma HFinite_HFinite_hcomplex_of_hypreal:
70216
40f19372a723 A bit of de-applying
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   370
  "z \<in> HFinite \<Longrightarrow> hcomplex_of_hypreal z \<in> HFinite"
40f19372a723 A bit of de-applying
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   371
  by (simp add: hcomplex_HFinite_iff)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   372
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   373
lemma SComplex_SReal_hcomplex_of_hypreal:
70216
40f19372a723 A bit of de-applying
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   374
     "x \<in> \<real> \<Longrightarrow>  hcomplex_of_hypreal x \<in> SComplex"
40f19372a723 A bit of de-applying
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   375
  by (simp add: Reals_eq_Standard)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   376
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   377
lemma stc_hcomplex_of_hypreal: 
70216
40f19372a723 A bit of de-applying
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   378
 "z \<in> HFinite \<Longrightarrow> stc(hcomplex_of_hypreal z) = hcomplex_of_hypreal (st z)"
70217
1f04832cbfcf more tidying up
paulson <lp15@cam.ac.uk>
parents: 70216
diff changeset
   379
  by (simp add: SComplex_SReal_hcomplex_of_hypreal st_SReal st_approx_self stc_unique)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   380
70217
1f04832cbfcf more tidying up
paulson <lp15@cam.ac.uk>
parents: 70216
diff changeset
   381
lemma hmod_stc_eq:
1f04832cbfcf more tidying up
paulson <lp15@cam.ac.uk>
parents: 70216
diff changeset
   382
  assumes "x \<in> HFinite" 
1f04832cbfcf more tidying up
paulson <lp15@cam.ac.uk>
parents: 70216
diff changeset
   383
  shows "hcmod(stc x) = st(hcmod x)"
1f04832cbfcf more tidying up
paulson <lp15@cam.ac.uk>
parents: 70216
diff changeset
   384
    by (metis SReal_hcmod_SComplex approx_HFinite approx_hnorm assms st_unique stc_SComplex_eq stc_eq_approx_iff stc_part_Ex)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   385
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   386
lemma Infinitesimal_hcnj_iff [simp]:
70217
1f04832cbfcf more tidying up
paulson <lp15@cam.ac.uk>
parents: 70216
diff changeset
   387
  "(hcnj z \<in> Infinitesimal) \<longleftrightarrow> (z \<in> Infinitesimal)"
1f04832cbfcf more tidying up
paulson <lp15@cam.ac.uk>
parents: 70216
diff changeset
   388
  by (simp add: Infinitesimal_hcmod_iff)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   389
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   390
end