src/HOL/Complex/NSCA.thy
author huffman
Thu, 27 Mar 2008 00:27:16 +0100
changeset 26420 57a626f64875
parent 26119 cb9bdde1b444
child 27148 5b78e50adc49
permissions -rw-r--r--
make preorder locale into a superclass of class po
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
13957
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
     1
(*  Title       : NSCA.thy
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
     2
    Author      : Jacques D. Fleuriot
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
     3
    Copyright   : 2001,2002 University of Edinburgh
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
     4
*)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
     5
14408
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
     6
header{*Non-Standard Complex Analysis*}
13957
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
     7
15131
c69542757a4d New theory header syntax.
nipkow
parents: 15013
diff changeset
     8
theory NSCA
22655
83878e551c8c minimize imports
huffman
parents: 21839
diff changeset
     9
imports NSComplex "../Hyperreal/HTranscendental"
15131
c69542757a4d New theory header syntax.
nipkow
parents: 15013
diff changeset
    10
begin
13957
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    11
21830
e38f0226e956 SComplex abbreviates Standard
huffman
parents: 21810
diff changeset
    12
abbreviation
13957
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    13
   (* standard complex numbers reagarded as an embedded subset of NS complex *)
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 20724
diff changeset
    14
   SComplex  :: "hcomplex set" where
21830
e38f0226e956 SComplex abbreviates Standard
huffman
parents: 21810
diff changeset
    15
   "SComplex \<equiv> Standard"
13957
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    16
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 20724
diff changeset
    17
definition
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 20724
diff changeset
    18
   stc :: "hcomplex => hcomplex" where
14408
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
    19
    --{* standard part map*}
20559
2116b7a371c7 removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents: 20557
diff changeset
    20
   "stc x = (SOME r. x \<in> HFinite & r:SComplex & r @= x)"
14408
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
    21
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
    22
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
    23
subsection{*Closure Laws for SComplex, the Standard Complex Numbers*}
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
    24
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
    25
lemma SComplex_minus_iff [simp]: "(-x \<in> SComplex) = (x \<in> SComplex)"
22883
005be8dafce0 hcomplex_of_hypreal abbreviates of_hypreal; removed redundant lemmas
huffman
parents: 22655
diff changeset
    26
by (auto, drule Standard_minus, auto)
14408
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
    27
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
    28
lemma SComplex_add_cancel:
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
    29
     "[| x + y \<in> SComplex; y \<in> SComplex |] ==> x \<in> SComplex"
22883
005be8dafce0 hcomplex_of_hypreal abbreviates of_hypreal; removed redundant lemmas
huffman
parents: 22655
diff changeset
    30
by (drule (1) Standard_diff, simp)
14408
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
    31
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
    32
lemma SReal_hcmod_hcomplex_of_complex [simp]:
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
    33
     "hcmod (hcomplex_of_complex r) \<in> Reals"
21830
e38f0226e956 SComplex abbreviates Standard
huffman
parents: 21810
diff changeset
    34
by (simp add: Reals_eq_Standard)
14408
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
    35
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
    36
lemma SReal_hcmod_number_of [simp]: "hcmod (number_of w ::hcomplex) \<in> Reals"
22883
005be8dafce0 hcomplex_of_hypreal abbreviates of_hypreal; removed redundant lemmas
huffman
parents: 22655
diff changeset
    37
by (simp add: Reals_eq_Standard)
14408
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
    38
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
    39
lemma SReal_hcmod_SComplex: "x \<in> SComplex ==> hcmod x \<in> Reals"
21830
e38f0226e956 SComplex abbreviates Standard
huffman
parents: 21810
diff changeset
    40
by (simp add: Reals_eq_Standard)
14408
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
    41
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
    42
lemma SComplex_divide_number_of:
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
    43
     "r \<in> SComplex ==> r/(number_of w::hcomplex) \<in> SComplex"
21830
e38f0226e956 SComplex abbreviates Standard
huffman
parents: 21810
diff changeset
    44
by simp
14408
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
    45
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
    46
lemma SComplex_UNIV_complex:
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
    47
     "{x. hcomplex_of_complex x \<in> SComplex} = (UNIV::complex set)"
21830
e38f0226e956 SComplex abbreviates Standard
huffman
parents: 21810
diff changeset
    48
by simp
14408
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
    49
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
    50
lemma SComplex_iff: "(x \<in> SComplex) = (\<exists>y. x = hcomplex_of_complex y)"
21830
e38f0226e956 SComplex abbreviates Standard
huffman
parents: 21810
diff changeset
    51
by (simp add: Standard_def image_def)
14408
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
    52
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
    53
lemma hcomplex_of_complex_image:
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
    54
     "hcomplex_of_complex `(UNIV::complex set) = SComplex"
21830
e38f0226e956 SComplex abbreviates Standard
huffman
parents: 21810
diff changeset
    55
by (simp add: Standard_def)
14408
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
    56
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
    57
lemma inv_hcomplex_of_complex_image: "inv hcomplex_of_complex `SComplex = UNIV"
21830
e38f0226e956 SComplex abbreviates Standard
huffman
parents: 21810
diff changeset
    58
apply (auto simp add: Standard_def image_def)
14408
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
    59
apply (rule inj_hcomplex_of_complex [THEN inv_f_f, THEN subst], blast)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
    60
done
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
    61
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
    62
lemma SComplex_hcomplex_of_complex_image: 
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
    63
      "[| \<exists>x. x: P; P \<le> SComplex |] ==> \<exists>Q. P = hcomplex_of_complex ` Q"
21830
e38f0226e956 SComplex abbreviates Standard
huffman
parents: 21810
diff changeset
    64
apply (simp add: Standard_def, blast)
14408
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
    65
done
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
    66
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
    67
lemma SComplex_SReal_dense:
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
    68
     "[| x \<in> SComplex; y \<in> SComplex; hcmod x < hcmod y  
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
    69
      |] ==> \<exists>r \<in> Reals. hcmod x< r & r < hcmod y"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
    70
apply (auto intro: SReal_dense simp add: SReal_hcmod_SComplex)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
    71
done
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
    72
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
    73
lemma SComplex_hcmod_SReal: 
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
    74
      "z \<in> SComplex ==> hcmod z \<in> Reals"
21830
e38f0226e956 SComplex abbreviates Standard
huffman
parents: 21810
diff changeset
    75
by (simp add: Reals_eq_Standard)
14408
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
    76
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
    77
(*
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
    78
Goalw [SComplex_def,SReal_def] "hcmod z \<in> Reals ==> z \<in> SComplex"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
    79
by (res_inst_tac [("z","z")] eq_Abs_hcomplex 1);
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
    80
by (auto_tac (claset(),simpset() addsimps [hcmod,hypreal_of_real_def,
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
    81
    hcomplex_of_complex_def,cmod_def]));
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
    82
*)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
    83
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
    84
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
    85
subsection{*The Finite Elements form a Subring*}
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
    86
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
    87
lemma HFinite_hcmod_hcomplex_of_complex [simp]:
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
    88
     "hcmod (hcomplex_of_complex r) \<in> HFinite"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
    89
by (auto intro!: SReal_subset_HFinite [THEN subsetD])
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
    90
20559
2116b7a371c7 removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents: 20557
diff changeset
    91
lemma HFinite_hcmod_iff: "(x \<in> HFinite) = (hcmod x \<in> HFinite)"
2116b7a371c7 removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents: 20557
diff changeset
    92
by (simp add: HFinite_def)
14408
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
    93
20559
2116b7a371c7 removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents: 20557
diff changeset
    94
lemma HFinite_bounded_hcmod:
2116b7a371c7 removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents: 20557
diff changeset
    95
  "[|x \<in> HFinite; y \<le> hcmod x; 0 \<le> y |] ==> y: HFinite"
2116b7a371c7 removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents: 20557
diff changeset
    96
by (auto intro: HFinite_bounded simp add: HFinite_hcmod_iff)
14408
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
    97
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
    98
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
    99
subsection{*The Complex Infinitesimals form a Subring*}
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   100
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   101
lemma hcomplex_sum_of_halves: "x/(2::hcomplex) + x/(2::hcomplex) = x"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   102
by auto
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   103
20559
2116b7a371c7 removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents: 20557
diff changeset
   104
lemma Infinitesimal_hcmod_iff: 
2116b7a371c7 removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents: 20557
diff changeset
   105
   "(z \<in> Infinitesimal) = (hcmod z \<in> Infinitesimal)"
2116b7a371c7 removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents: 20557
diff changeset
   106
by (simp add: Infinitesimal_def)
14408
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   107
20559
2116b7a371c7 removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents: 20557
diff changeset
   108
lemma HInfinite_hcmod_iff: "(z \<in> HInfinite) = (hcmod z \<in> HInfinite)"
2116b7a371c7 removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents: 20557
diff changeset
   109
by (simp add: HInfinite_def)
14408
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   110
20559
2116b7a371c7 removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents: 20557
diff changeset
   111
lemma HFinite_diff_Infinitesimal_hcmod:
2116b7a371c7 removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents: 20557
diff changeset
   112
     "x \<in> HFinite - Infinitesimal ==> hcmod x \<in> HFinite - Infinitesimal"
2116b7a371c7 removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents: 20557
diff changeset
   113
by (simp add: HFinite_hcmod_iff Infinitesimal_hcmod_iff)
14408
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   114
20559
2116b7a371c7 removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents: 20557
diff changeset
   115
lemma hcmod_less_Infinitesimal:
2116b7a371c7 removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents: 20557
diff changeset
   116
     "[| e \<in> Infinitesimal; hcmod x < hcmod e |] ==> x \<in> Infinitesimal"
2116b7a371c7 removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents: 20557
diff changeset
   117
by (auto elim: hrabs_less_Infinitesimal simp add: Infinitesimal_hcmod_iff)
14408
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   118
20559
2116b7a371c7 removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents: 20557
diff changeset
   119
lemma hcmod_le_Infinitesimal:
2116b7a371c7 removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents: 20557
diff changeset
   120
     "[| e \<in> Infinitesimal; hcmod x \<le> hcmod e |] ==> x \<in> Infinitesimal"
2116b7a371c7 removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents: 20557
diff changeset
   121
by (auto elim: hrabs_le_Infinitesimal simp add: Infinitesimal_hcmod_iff)
14408
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   122
20559
2116b7a371c7 removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents: 20557
diff changeset
   123
lemma Infinitesimal_interval_hcmod:
2116b7a371c7 removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents: 20557
diff changeset
   124
     "[| e \<in> Infinitesimal;  
2116b7a371c7 removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents: 20557
diff changeset
   125
          e' \<in> Infinitesimal;  
2116b7a371c7 removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents: 20557
diff changeset
   126
          hcmod e' < hcmod x ; hcmod x < hcmod e  
2116b7a371c7 removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents: 20557
diff changeset
   127
       |] ==> x \<in> Infinitesimal"
2116b7a371c7 removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents: 20557
diff changeset
   128
by (auto intro: Infinitesimal_interval simp add: Infinitesimal_hcmod_iff)
14408
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   129
20559
2116b7a371c7 removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents: 20557
diff changeset
   130
lemma Infinitesimal_interval2_hcmod:
2116b7a371c7 removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents: 20557
diff changeset
   131
     "[| e \<in> Infinitesimal;  
2116b7a371c7 removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents: 20557
diff changeset
   132
         e' \<in> Infinitesimal;  
14408
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   133
         hcmod e' \<le> hcmod x ; hcmod x \<le> hcmod e  
20559
2116b7a371c7 removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents: 20557
diff changeset
   134
      |] ==> x \<in> Infinitesimal"
2116b7a371c7 removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents: 20557
diff changeset
   135
by (auto intro: Infinitesimal_interval2 simp add: Infinitesimal_hcmod_iff)
14408
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   136
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   137
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   138
subsection{*The ``Infinitely Close'' Relation*}
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   139
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   140
(*
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   141
Goalw [capprox_def,approx_def] "(z @c= w) = (hcmod z @= hcmod w)"
20559
2116b7a371c7 removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents: 20557
diff changeset
   142
by (auto_tac (claset(),simpset() addsimps [Infinitesimal_hcmod_iff]));
14408
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   143
*)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   144
20559
2116b7a371c7 removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents: 20557
diff changeset
   145
lemma approx_SComplex_mult_cancel_zero:
2116b7a371c7 removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents: 20557
diff changeset
   146
     "[| a \<in> SComplex; a \<noteq> 0; a*x @= 0 |] ==> x @= 0"
22883
005be8dafce0 hcomplex_of_hypreal abbreviates of_hypreal; removed redundant lemmas
huffman
parents: 22655
diff changeset
   147
apply (drule Standard_inverse [THEN Standard_subset_HFinite [THEN subsetD]])
20559
2116b7a371c7 removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents: 20557
diff changeset
   148
apply (auto dest: approx_mult2 simp add: mult_assoc [symmetric])
14408
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   149
done
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   150
20559
2116b7a371c7 removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents: 20557
diff changeset
   151
lemma approx_mult_SComplex1: "[| a \<in> SComplex; x @= 0 |] ==> x*a @= 0"
22883
005be8dafce0 hcomplex_of_hypreal abbreviates of_hypreal; removed redundant lemmas
huffman
parents: 22655
diff changeset
   152
by (auto dest: Standard_subset_HFinite [THEN subsetD] approx_mult1)
14408
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   153
20559
2116b7a371c7 removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents: 20557
diff changeset
   154
lemma approx_mult_SComplex2: "[| a \<in> SComplex; x @= 0 |] ==> a*x @= 0"
22883
005be8dafce0 hcomplex_of_hypreal abbreviates of_hypreal; removed redundant lemmas
huffman
parents: 22655
diff changeset
   155
by (auto dest: Standard_subset_HFinite [THEN subsetD] approx_mult2)
14408
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   156
20559
2116b7a371c7 removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents: 20557
diff changeset
   157
lemma approx_mult_SComplex_zero_cancel_iff [simp]:
2116b7a371c7 removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents: 20557
diff changeset
   158
     "[|a \<in> SComplex; a \<noteq> 0 |] ==> (a*x @= 0) = (x @= 0)"
2116b7a371c7 removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents: 20557
diff changeset
   159
by (blast intro: approx_SComplex_mult_cancel_zero approx_mult_SComplex2)
14408
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   160
20559
2116b7a371c7 removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents: 20557
diff changeset
   161
lemma approx_SComplex_mult_cancel:
2116b7a371c7 removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents: 20557
diff changeset
   162
     "[| a \<in> SComplex; a \<noteq> 0; a* w @= a*z |] ==> w @= z"
22883
005be8dafce0 hcomplex_of_hypreal abbreviates of_hypreal; removed redundant lemmas
huffman
parents: 22655
diff changeset
   163
apply (drule Standard_inverse [THEN Standard_subset_HFinite [THEN subsetD]])
20559
2116b7a371c7 removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents: 20557
diff changeset
   164
apply (auto dest: approx_mult2 simp add: mult_assoc [symmetric])
14408
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   165
done
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   166
20559
2116b7a371c7 removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents: 20557
diff changeset
   167
lemma approx_SComplex_mult_cancel_iff1 [simp]:
2116b7a371c7 removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents: 20557
diff changeset
   168
     "[| a \<in> SComplex; a \<noteq> 0|] ==> (a* w @= a*z) = (w @= z)"
22883
005be8dafce0 hcomplex_of_hypreal abbreviates of_hypreal; removed redundant lemmas
huffman
parents: 22655
diff changeset
   169
by (auto intro!: approx_mult2 Standard_subset_HFinite [THEN subsetD]
20559
2116b7a371c7 removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents: 20557
diff changeset
   170
            intro: approx_SComplex_mult_cancel)
14408
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   171
21810
b2d23672b003 generalized some lemmas; removed redundant lemmas; cleaned up some proofs
huffman
parents: 21404
diff changeset
   172
(* TODO: generalize following theorems: hcmod -> hnorm *)
b2d23672b003 generalized some lemmas; removed redundant lemmas; cleaned up some proofs
huffman
parents: 21404
diff changeset
   173
20559
2116b7a371c7 removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents: 20557
diff changeset
   174
lemma approx_hcmod_approx_zero: "(x @= y) = (hcmod (y - x) @= 0)"
22883
005be8dafce0 hcomplex_of_hypreal abbreviates of_hypreal; removed redundant lemmas
huffman
parents: 22655
diff changeset
   175
apply (subst hnorm_minus_commute)
20559
2116b7a371c7 removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents: 20557
diff changeset
   176
apply (simp add: approx_def Infinitesimal_hcmod_iff diff_minus)
14408
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   177
done
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   178
20559
2116b7a371c7 removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents: 20557
diff changeset
   179
lemma approx_approx_zero_iff: "(x @= 0) = (hcmod x @= 0)"
2116b7a371c7 removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents: 20557
diff changeset
   180
by (simp add: approx_hcmod_approx_zero)
14408
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   181
20559
2116b7a371c7 removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents: 20557
diff changeset
   182
lemma approx_minus_zero_cancel_iff [simp]: "(-x @= 0) = (x @= 0)"
2116b7a371c7 removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents: 20557
diff changeset
   183
by (simp add: approx_def)
14408
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   184
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   185
lemma Infinitesimal_hcmod_add_diff:
20559
2116b7a371c7 removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents: 20557
diff changeset
   186
     "u @= 0 ==> hcmod(x + u) - hcmod x \<in> Infinitesimal"
2116b7a371c7 removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents: 20557
diff changeset
   187
apply (drule approx_approx_zero_iff [THEN iffD1])
14408
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   188
apply (rule_tac e = "hcmod u" and e' = "- hcmod u" in Infinitesimal_interval2)
20559
2116b7a371c7 removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents: 20557
diff changeset
   189
apply (auto simp add: mem_infmal_iff [symmetric] diff_def)
14408
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   190
apply (rule_tac c1 = "hcmod x" in add_le_cancel_left [THEN iffD1])
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   191
apply (auto simp add: diff_minus [symmetric])
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   192
done
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   193
20559
2116b7a371c7 removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents: 20557
diff changeset
   194
lemma approx_hcmod_add_hcmod: "u @= 0 ==> hcmod(x + u) @= hcmod x"
14408
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   195
apply (rule approx_minus_iff [THEN iffD2])
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   196
apply (auto intro: Infinitesimal_hcmod_add_diff simp add: mem_infmal_iff [symmetric] diff_minus [symmetric])
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   197
done
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   198
13957
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   199
14408
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   200
subsection{*Zero is the Only Infinitesimal Complex Number*}
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   201
20559
2116b7a371c7 removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents: 20557
diff changeset
   202
lemma Infinitesimal_less_SComplex:
2116b7a371c7 removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents: 20557
diff changeset
   203
   "[| x \<in> SComplex; y \<in> Infinitesimal; 0 < hcmod x |] ==> hcmod y < hcmod x"
2116b7a371c7 removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents: 20557
diff changeset
   204
by (auto intro: Infinitesimal_less_SReal SComplex_hcmod_SReal simp add: Infinitesimal_hcmod_iff)
14408
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   205
20559
2116b7a371c7 removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents: 20557
diff changeset
   206
lemma SComplex_Int_Infinitesimal_zero: "SComplex Int Infinitesimal = {0}"
21830
e38f0226e956 SComplex abbreviates Standard
huffman
parents: 21810
diff changeset
   207
by (auto simp add: Standard_def Infinitesimal_hcmod_iff)
14408
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   208
20559
2116b7a371c7 removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents: 20557
diff changeset
   209
lemma SComplex_Infinitesimal_zero:
2116b7a371c7 removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents: 20557
diff changeset
   210
     "[| x \<in> SComplex; x \<in> Infinitesimal|] ==> x = 0"
2116b7a371c7 removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents: 20557
diff changeset
   211
by (cut_tac SComplex_Int_Infinitesimal_zero, blast)
14408
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   212
20559
2116b7a371c7 removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents: 20557
diff changeset
   213
lemma SComplex_HFinite_diff_Infinitesimal:
2116b7a371c7 removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents: 20557
diff changeset
   214
     "[| x \<in> SComplex; x \<noteq> 0 |] ==> x \<in> HFinite - Infinitesimal"
22883
005be8dafce0 hcomplex_of_hypreal abbreviates of_hypreal; removed redundant lemmas
huffman
parents: 22655
diff changeset
   215
by (auto dest: SComplex_Infinitesimal_zero Standard_subset_HFinite [THEN subsetD])
14408
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   216
20559
2116b7a371c7 removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents: 20557
diff changeset
   217
lemma hcomplex_of_complex_HFinite_diff_Infinitesimal:
2116b7a371c7 removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents: 20557
diff changeset
   218
     "hcomplex_of_complex x \<noteq> 0 
2116b7a371c7 removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents: 20557
diff changeset
   219
      ==> hcomplex_of_complex x \<in> HFinite - Infinitesimal"
2116b7a371c7 removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents: 20557
diff changeset
   220
by (rule SComplex_HFinite_diff_Infinitesimal, auto)
14408
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   221
20559
2116b7a371c7 removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents: 20557
diff changeset
   222
lemma number_of_not_Infinitesimal [simp]:
2116b7a371c7 removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents: 20557
diff changeset
   223
     "number_of w \<noteq> (0::hcomplex) ==> (number_of w::hcomplex) \<notin> Infinitesimal"
22883
005be8dafce0 hcomplex_of_hypreal abbreviates of_hypreal; removed redundant lemmas
huffman
parents: 22655
diff changeset
   224
by (fast dest: Standard_number_of [THEN SComplex_Infinitesimal_zero])
14408
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   225
20559
2116b7a371c7 removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents: 20557
diff changeset
   226
lemma approx_SComplex_not_zero:
2116b7a371c7 removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents: 20557
diff changeset
   227
     "[| y \<in> SComplex; x @= y; y\<noteq> 0 |] ==> x \<noteq> 0"
2116b7a371c7 removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents: 20557
diff changeset
   228
by (auto dest: SComplex_Infinitesimal_zero approx_sym [THEN mem_infmal_iff [THEN iffD2]])
14408
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   229
20559
2116b7a371c7 removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents: 20557
diff changeset
   230
lemma SComplex_approx_iff:
2116b7a371c7 removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents: 20557
diff changeset
   231
     "[|x \<in> SComplex; y \<in> SComplex|] ==> (x @= y) = (x = y)"
21830
e38f0226e956 SComplex abbreviates Standard
huffman
parents: 21810
diff changeset
   232
by (auto simp add: Standard_def)
14408
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   233
20559
2116b7a371c7 removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents: 20557
diff changeset
   234
lemma number_of_Infinitesimal_iff [simp]:
2116b7a371c7 removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents: 20557
diff changeset
   235
     "((number_of w :: hcomplex) \<in> Infinitesimal) =
2116b7a371c7 removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents: 20557
diff changeset
   236
      (number_of w = (0::hcomplex))"
14408
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   237
apply (rule iffI)
22883
005be8dafce0 hcomplex_of_hypreal abbreviates of_hypreal; removed redundant lemmas
huffman
parents: 22655
diff changeset
   238
apply (fast dest: Standard_number_of [THEN SComplex_Infinitesimal_zero])
14408
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   239
apply (simp (no_asm_simp))
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   240
done
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   241
20559
2116b7a371c7 removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents: 20557
diff changeset
   242
lemma approx_unique_complex:
2116b7a371c7 removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents: 20557
diff changeset
   243
     "[| r \<in> SComplex; s \<in> SComplex; r @= x; s @= x|] ==> r = s"
2116b7a371c7 removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents: 20557
diff changeset
   244
by (blast intro: SComplex_approx_iff [THEN iffD1] approx_trans2)
14408
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   245
21839
54018ed3b99d added lemmas about hRe, hIm, HComplex; removed all uses of star_n
huffman
parents: 21830
diff changeset
   246
subsection {* Properties of @{term hRe}, @{term hIm} and @{term HComplex} *}
54018ed3b99d added lemmas about hRe, hIm, HComplex; removed all uses of star_n
huffman
parents: 21830
diff changeset
   247
54018ed3b99d added lemmas about hRe, hIm, HComplex; removed all uses of star_n
huffman
parents: 21830
diff changeset
   248
54018ed3b99d added lemmas about hRe, hIm, HComplex; removed all uses of star_n
huffman
parents: 21830
diff changeset
   249
lemma abs_hRe_le_hcmod: "\<And>x. \<bar>hRe x\<bar> \<le> hcmod x"
54018ed3b99d added lemmas about hRe, hIm, HComplex; removed all uses of star_n
huffman
parents: 21830
diff changeset
   250
by transfer (rule abs_Re_le_cmod)
54018ed3b99d added lemmas about hRe, hIm, HComplex; removed all uses of star_n
huffman
parents: 21830
diff changeset
   251
54018ed3b99d added lemmas about hRe, hIm, HComplex; removed all uses of star_n
huffman
parents: 21830
diff changeset
   252
lemma abs_hIm_le_hcmod: "\<And>x. \<bar>hIm x\<bar> \<le> hcmod x"
54018ed3b99d added lemmas about hRe, hIm, HComplex; removed all uses of star_n
huffman
parents: 21830
diff changeset
   253
by transfer (rule abs_Im_le_cmod)
54018ed3b99d added lemmas about hRe, hIm, HComplex; removed all uses of star_n
huffman
parents: 21830
diff changeset
   254
54018ed3b99d added lemmas about hRe, hIm, HComplex; removed all uses of star_n
huffman
parents: 21830
diff changeset
   255
lemma Infinitesimal_hRe: "x \<in> Infinitesimal \<Longrightarrow> hRe x \<in> Infinitesimal"
54018ed3b99d added lemmas about hRe, hIm, HComplex; removed all uses of star_n
huffman
parents: 21830
diff changeset
   256
apply (rule InfinitesimalI2, simp)
54018ed3b99d added lemmas about hRe, hIm, HComplex; removed all uses of star_n
huffman
parents: 21830
diff changeset
   257
apply (rule order_le_less_trans [OF abs_hRe_le_hcmod])
54018ed3b99d added lemmas about hRe, hIm, HComplex; removed all uses of star_n
huffman
parents: 21830
diff changeset
   258
apply (erule (1) InfinitesimalD2)
54018ed3b99d added lemmas about hRe, hIm, HComplex; removed all uses of star_n
huffman
parents: 21830
diff changeset
   259
done
54018ed3b99d added lemmas about hRe, hIm, HComplex; removed all uses of star_n
huffman
parents: 21830
diff changeset
   260
54018ed3b99d added lemmas about hRe, hIm, HComplex; removed all uses of star_n
huffman
parents: 21830
diff changeset
   261
lemma Infinitesimal_hIm: "x \<in> Infinitesimal \<Longrightarrow> hIm x \<in> Infinitesimal"
54018ed3b99d added lemmas about hRe, hIm, HComplex; removed all uses of star_n
huffman
parents: 21830
diff changeset
   262
apply (rule InfinitesimalI2, simp)
54018ed3b99d added lemmas about hRe, hIm, HComplex; removed all uses of star_n
huffman
parents: 21830
diff changeset
   263
apply (rule order_le_less_trans [OF abs_hIm_le_hcmod])
54018ed3b99d added lemmas about hRe, hIm, HComplex; removed all uses of star_n
huffman
parents: 21830
diff changeset
   264
apply (erule (1) InfinitesimalD2)
54018ed3b99d added lemmas about hRe, hIm, HComplex; removed all uses of star_n
huffman
parents: 21830
diff changeset
   265
done
54018ed3b99d added lemmas about hRe, hIm, HComplex; removed all uses of star_n
huffman
parents: 21830
diff changeset
   266
22956
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22883
diff changeset
   267
lemma real_sqrt_lessI: "\<lbrakk>0 < u; x < u\<twosuperior>\<rbrakk> \<Longrightarrow> sqrt x < u"
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22883
diff changeset
   268
(* TODO: this belongs somewhere else *)
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22883
diff changeset
   269
by (frule real_sqrt_less_mono) simp
21839
54018ed3b99d added lemmas about hRe, hIm, HComplex; removed all uses of star_n
huffman
parents: 21830
diff changeset
   270
54018ed3b99d added lemmas about hRe, hIm, HComplex; removed all uses of star_n
huffman
parents: 21830
diff changeset
   271
lemma hypreal_sqrt_lessI:
22956
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22883
diff changeset
   272
  "\<And>x u. \<lbrakk>0 < u; x < u\<twosuperior>\<rbrakk> \<Longrightarrow> ( *f* sqrt) x < u"
21839
54018ed3b99d added lemmas about hRe, hIm, HComplex; removed all uses of star_n
huffman
parents: 21830
diff changeset
   273
by transfer (rule real_sqrt_lessI)
54018ed3b99d added lemmas about hRe, hIm, HComplex; removed all uses of star_n
huffman
parents: 21830
diff changeset
   274
 
54018ed3b99d added lemmas about hRe, hIm, HComplex; removed all uses of star_n
huffman
parents: 21830
diff changeset
   275
lemma hypreal_sqrt_ge_zero: "\<And>x. 0 \<le> x \<Longrightarrow> 0 \<le> ( *f* sqrt) x"
54018ed3b99d added lemmas about hRe, hIm, HComplex; removed all uses of star_n
huffman
parents: 21830
diff changeset
   276
by transfer (rule real_sqrt_ge_zero)
54018ed3b99d added lemmas about hRe, hIm, HComplex; removed all uses of star_n
huffman
parents: 21830
diff changeset
   277
54018ed3b99d added lemmas about hRe, hIm, HComplex; removed all uses of star_n
huffman
parents: 21830
diff changeset
   278
lemma Infinitesimal_sqrt:
54018ed3b99d added lemmas about hRe, hIm, HComplex; removed all uses of star_n
huffman
parents: 21830
diff changeset
   279
  "\<lbrakk>x \<in> Infinitesimal; 0 \<le> x\<rbrakk> \<Longrightarrow> ( *f* sqrt) x \<in> Infinitesimal"
54018ed3b99d added lemmas about hRe, hIm, HComplex; removed all uses of star_n
huffman
parents: 21830
diff changeset
   280
apply (rule InfinitesimalI2)
54018ed3b99d added lemmas about hRe, hIm, HComplex; removed all uses of star_n
huffman
parents: 21830
diff changeset
   281
apply (drule_tac r="r\<twosuperior>" in InfinitesimalD2, simp)
54018ed3b99d added lemmas about hRe, hIm, HComplex; removed all uses of star_n
huffman
parents: 21830
diff changeset
   282
apply (simp add: hypreal_sqrt_ge_zero)
54018ed3b99d added lemmas about hRe, hIm, HComplex; removed all uses of star_n
huffman
parents: 21830
diff changeset
   283
apply (rule hypreal_sqrt_lessI, simp_all)
14408
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   284
done
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   285
21839
54018ed3b99d added lemmas about hRe, hIm, HComplex; removed all uses of star_n
huffman
parents: 21830
diff changeset
   286
lemma Infinitesimal_HComplex:
54018ed3b99d added lemmas about hRe, hIm, HComplex; removed all uses of star_n
huffman
parents: 21830
diff changeset
   287
  "\<lbrakk>x \<in> Infinitesimal; y \<in> Infinitesimal\<rbrakk> \<Longrightarrow> HComplex x y \<in> Infinitesimal"
54018ed3b99d added lemmas about hRe, hIm, HComplex; removed all uses of star_n
huffman
parents: 21830
diff changeset
   288
apply (rule Infinitesimal_hcmod_iff [THEN iffD2])
54018ed3b99d added lemmas about hRe, hIm, HComplex; removed all uses of star_n
huffman
parents: 21830
diff changeset
   289
apply (simp add: hcmod_i)
54018ed3b99d added lemmas about hRe, hIm, HComplex; removed all uses of star_n
huffman
parents: 21830
diff changeset
   290
apply (rule Infinitesimal_sqrt)
54018ed3b99d added lemmas about hRe, hIm, HComplex; removed all uses of star_n
huffman
parents: 21830
diff changeset
   291
apply (rule Infinitesimal_add)
54018ed3b99d added lemmas about hRe, hIm, HComplex; removed all uses of star_n
huffman
parents: 21830
diff changeset
   292
apply (erule Infinitesimal_hrealpow, simp)
54018ed3b99d added lemmas about hRe, hIm, HComplex; removed all uses of star_n
huffman
parents: 21830
diff changeset
   293
apply (erule Infinitesimal_hrealpow, simp)
54018ed3b99d added lemmas about hRe, hIm, HComplex; removed all uses of star_n
huffman
parents: 21830
diff changeset
   294
apply (rule add_nonneg_nonneg)
54018ed3b99d added lemmas about hRe, hIm, HComplex; removed all uses of star_n
huffman
parents: 21830
diff changeset
   295
apply (rule zero_le_power2)
54018ed3b99d added lemmas about hRe, hIm, HComplex; removed all uses of star_n
huffman
parents: 21830
diff changeset
   296
apply (rule zero_le_power2)
54018ed3b99d added lemmas about hRe, hIm, HComplex; removed all uses of star_n
huffman
parents: 21830
diff changeset
   297
done
54018ed3b99d added lemmas about hRe, hIm, HComplex; removed all uses of star_n
huffman
parents: 21830
diff changeset
   298
54018ed3b99d added lemmas about hRe, hIm, HComplex; removed all uses of star_n
huffman
parents: 21830
diff changeset
   299
lemma hcomplex_Infinitesimal_iff:
54018ed3b99d added lemmas about hRe, hIm, HComplex; removed all uses of star_n
huffman
parents: 21830
diff changeset
   300
  "(x \<in> Infinitesimal) = (hRe x \<in> Infinitesimal \<and> hIm x \<in> Infinitesimal)"
54018ed3b99d added lemmas about hRe, hIm, HComplex; removed all uses of star_n
huffman
parents: 21830
diff changeset
   301
apply (safe intro!: Infinitesimal_hRe Infinitesimal_hIm)
54018ed3b99d added lemmas about hRe, hIm, HComplex; removed all uses of star_n
huffman
parents: 21830
diff changeset
   302
apply (drule (1) Infinitesimal_HComplex, simp)
14408
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   303
done
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   304
21839
54018ed3b99d added lemmas about hRe, hIm, HComplex; removed all uses of star_n
huffman
parents: 21830
diff changeset
   305
lemma hRe_diff [simp]: "\<And>x y. hRe (x - y) = hRe x - hRe y"
54018ed3b99d added lemmas about hRe, hIm, HComplex; removed all uses of star_n
huffman
parents: 21830
diff changeset
   306
by transfer (rule complex_Re_diff)
54018ed3b99d added lemmas about hRe, hIm, HComplex; removed all uses of star_n
huffman
parents: 21830
diff changeset
   307
54018ed3b99d added lemmas about hRe, hIm, HComplex; removed all uses of star_n
huffman
parents: 21830
diff changeset
   308
lemma hIm_diff [simp]: "\<And>x y. hIm (x - y) = hIm x - hIm y"
54018ed3b99d added lemmas about hRe, hIm, HComplex; removed all uses of star_n
huffman
parents: 21830
diff changeset
   309
by transfer (rule complex_Im_diff)
54018ed3b99d added lemmas about hRe, hIm, HComplex; removed all uses of star_n
huffman
parents: 21830
diff changeset
   310
54018ed3b99d added lemmas about hRe, hIm, HComplex; removed all uses of star_n
huffman
parents: 21830
diff changeset
   311
lemma approx_hRe: "x \<approx> y \<Longrightarrow> hRe x \<approx> hRe y"
54018ed3b99d added lemmas about hRe, hIm, HComplex; removed all uses of star_n
huffman
parents: 21830
diff changeset
   312
unfolding approx_def by (drule Infinitesimal_hRe) simp
54018ed3b99d added lemmas about hRe, hIm, HComplex; removed all uses of star_n
huffman
parents: 21830
diff changeset
   313
54018ed3b99d added lemmas about hRe, hIm, HComplex; removed all uses of star_n
huffman
parents: 21830
diff changeset
   314
lemma approx_hIm: "x \<approx> y \<Longrightarrow> hIm x \<approx> hIm y"
54018ed3b99d added lemmas about hRe, hIm, HComplex; removed all uses of star_n
huffman
parents: 21830
diff changeset
   315
unfolding approx_def by (drule Infinitesimal_hIm) simp
54018ed3b99d added lemmas about hRe, hIm, HComplex; removed all uses of star_n
huffman
parents: 21830
diff changeset
   316
54018ed3b99d added lemmas about hRe, hIm, HComplex; removed all uses of star_n
huffman
parents: 21830
diff changeset
   317
lemma approx_HComplex:
54018ed3b99d added lemmas about hRe, hIm, HComplex; removed all uses of star_n
huffman
parents: 21830
diff changeset
   318
  "\<lbrakk>a \<approx> b; c \<approx> d\<rbrakk> \<Longrightarrow> HComplex a c \<approx> HComplex b d"
54018ed3b99d added lemmas about hRe, hIm, HComplex; removed all uses of star_n
huffman
parents: 21830
diff changeset
   319
unfolding approx_def by (simp add: Infinitesimal_HComplex)
54018ed3b99d added lemmas about hRe, hIm, HComplex; removed all uses of star_n
huffman
parents: 21830
diff changeset
   320
54018ed3b99d added lemmas about hRe, hIm, HComplex; removed all uses of star_n
huffman
parents: 21830
diff changeset
   321
lemma hcomplex_approx_iff:
54018ed3b99d added lemmas about hRe, hIm, HComplex; removed all uses of star_n
huffman
parents: 21830
diff changeset
   322
  "(x \<approx> y) = (hRe x \<approx> hRe y \<and> hIm x \<approx> hIm y)"
54018ed3b99d added lemmas about hRe, hIm, HComplex; removed all uses of star_n
huffman
parents: 21830
diff changeset
   323
unfolding approx_def by (simp add: hcomplex_Infinitesimal_iff)
54018ed3b99d added lemmas about hRe, hIm, HComplex; removed all uses of star_n
huffman
parents: 21830
diff changeset
   324
54018ed3b99d added lemmas about hRe, hIm, HComplex; removed all uses of star_n
huffman
parents: 21830
diff changeset
   325
lemma HFinite_hRe: "x \<in> HFinite \<Longrightarrow> hRe x \<in> HFinite"
54018ed3b99d added lemmas about hRe, hIm, HComplex; removed all uses of star_n
huffman
parents: 21830
diff changeset
   326
apply (auto simp add: HFinite_def SReal_def)
54018ed3b99d added lemmas about hRe, hIm, HComplex; removed all uses of star_n
huffman
parents: 21830
diff changeset
   327
apply (rule_tac x="star_of r" in exI, simp)
54018ed3b99d added lemmas about hRe, hIm, HComplex; removed all uses of star_n
huffman
parents: 21830
diff changeset
   328
apply (erule order_le_less_trans [OF abs_hRe_le_hcmod])
14408
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   329
done
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   330
21839
54018ed3b99d added lemmas about hRe, hIm, HComplex; removed all uses of star_n
huffman
parents: 21830
diff changeset
   331
lemma HFinite_hIm: "x \<in> HFinite \<Longrightarrow> hIm x \<in> HFinite"
54018ed3b99d added lemmas about hRe, hIm, HComplex; removed all uses of star_n
huffman
parents: 21830
diff changeset
   332
apply (auto simp add: HFinite_def SReal_def)
54018ed3b99d added lemmas about hRe, hIm, HComplex; removed all uses of star_n
huffman
parents: 21830
diff changeset
   333
apply (rule_tac x="star_of r" in exI, simp)
54018ed3b99d added lemmas about hRe, hIm, HComplex; removed all uses of star_n
huffman
parents: 21830
diff changeset
   334
apply (erule order_le_less_trans [OF abs_hIm_le_hcmod])
14408
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   335
done
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   336
21839
54018ed3b99d added lemmas about hRe, hIm, HComplex; removed all uses of star_n
huffman
parents: 21830
diff changeset
   337
lemma HFinite_HComplex:
54018ed3b99d added lemmas about hRe, hIm, HComplex; removed all uses of star_n
huffman
parents: 21830
diff changeset
   338
  "\<lbrakk>x \<in> HFinite; y \<in> HFinite\<rbrakk> \<Longrightarrow> HComplex x y \<in> HFinite"
54018ed3b99d added lemmas about hRe, hIm, HComplex; removed all uses of star_n
huffman
parents: 21830
diff changeset
   339
apply (subgoal_tac "HComplex x 0 + HComplex 0 y \<in> HFinite", simp)
54018ed3b99d added lemmas about hRe, hIm, HComplex; removed all uses of star_n
huffman
parents: 21830
diff changeset
   340
apply (rule HFinite_add)
54018ed3b99d added lemmas about hRe, hIm, HComplex; removed all uses of star_n
huffman
parents: 21830
diff changeset
   341
apply (simp add: HFinite_hcmod_iff hcmod_i)
54018ed3b99d added lemmas about hRe, hIm, HComplex; removed all uses of star_n
huffman
parents: 21830
diff changeset
   342
apply (simp add: HFinite_hcmod_iff hcmod_i)
54018ed3b99d added lemmas about hRe, hIm, HComplex; removed all uses of star_n
huffman
parents: 21830
diff changeset
   343
done
54018ed3b99d added lemmas about hRe, hIm, HComplex; removed all uses of star_n
huffman
parents: 21830
diff changeset
   344
54018ed3b99d added lemmas about hRe, hIm, HComplex; removed all uses of star_n
huffman
parents: 21830
diff changeset
   345
lemma hcomplex_HFinite_iff:
54018ed3b99d added lemmas about hRe, hIm, HComplex; removed all uses of star_n
huffman
parents: 21830
diff changeset
   346
  "(x \<in> HFinite) = (hRe x \<in> HFinite \<and> hIm x \<in> HFinite)"
54018ed3b99d added lemmas about hRe, hIm, HComplex; removed all uses of star_n
huffman
parents: 21830
diff changeset
   347
apply (safe intro!: HFinite_hRe HFinite_hIm)
54018ed3b99d added lemmas about hRe, hIm, HComplex; removed all uses of star_n
huffman
parents: 21830
diff changeset
   348
apply (drule (1) HFinite_HComplex, simp)
54018ed3b99d added lemmas about hRe, hIm, HComplex; removed all uses of star_n
huffman
parents: 21830
diff changeset
   349
done
54018ed3b99d added lemmas about hRe, hIm, HComplex; removed all uses of star_n
huffman
parents: 21830
diff changeset
   350
54018ed3b99d added lemmas about hRe, hIm, HComplex; removed all uses of star_n
huffman
parents: 21830
diff changeset
   351
lemma hcomplex_HInfinite_iff:
54018ed3b99d added lemmas about hRe, hIm, HComplex; removed all uses of star_n
huffman
parents: 21830
diff changeset
   352
  "(x \<in> HInfinite) = (hRe x \<in> HInfinite \<or> hIm x \<in> HInfinite)"
54018ed3b99d added lemmas about hRe, hIm, HComplex; removed all uses of star_n
huffman
parents: 21830
diff changeset
   353
by (simp add: HInfinite_HFinite_iff hcomplex_HFinite_iff)
54018ed3b99d added lemmas about hRe, hIm, HComplex; removed all uses of star_n
huffman
parents: 21830
diff changeset
   354
20559
2116b7a371c7 removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents: 20557
diff changeset
   355
lemma hcomplex_of_hypreal_approx_iff [simp]:
2116b7a371c7 removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents: 20557
diff changeset
   356
     "(hcomplex_of_hypreal x @= hcomplex_of_hypreal z) = (x @= z)"
21839
54018ed3b99d added lemmas about hRe, hIm, HComplex; removed all uses of star_n
huffman
parents: 21830
diff changeset
   357
by (simp add: hcomplex_approx_iff)
14408
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   358
21839
54018ed3b99d added lemmas about hRe, hIm, HComplex; removed all uses of star_n
huffman
parents: 21830
diff changeset
   359
lemma Standard_HComplex:
54018ed3b99d added lemmas about hRe, hIm, HComplex; removed all uses of star_n
huffman
parents: 21830
diff changeset
   360
  "\<lbrakk>x \<in> Standard; y \<in> Standard\<rbrakk> \<Longrightarrow> HComplex x y \<in> Standard"
54018ed3b99d added lemmas about hRe, hIm, HComplex; removed all uses of star_n
huffman
parents: 21830
diff changeset
   361
by (simp add: HComplex_def)
14408
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   362
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   363
(* Here we go - easy proof now!! *)
20559
2116b7a371c7 removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents: 20557
diff changeset
   364
lemma stc_part_Ex: "x:HFinite ==> \<exists>t \<in> SComplex. x @= t"
21839
54018ed3b99d added lemmas about hRe, hIm, HComplex; removed all uses of star_n
huffman
parents: 21830
diff changeset
   365
apply (simp add: hcomplex_HFinite_iff hcomplex_approx_iff)
54018ed3b99d added lemmas about hRe, hIm, HComplex; removed all uses of star_n
huffman
parents: 21830
diff changeset
   366
apply (rule_tac x="HComplex (st (hRe x)) (st (hIm x))" in bexI)
54018ed3b99d added lemmas about hRe, hIm, HComplex; removed all uses of star_n
huffman
parents: 21830
diff changeset
   367
apply (simp add: st_approx_self [THEN approx_sym])
54018ed3b99d added lemmas about hRe, hIm, HComplex; removed all uses of star_n
huffman
parents: 21830
diff changeset
   368
apply (simp add: Standard_HComplex st_SReal [unfolded Reals_eq_Standard])
14408
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   369
done
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   370
20559
2116b7a371c7 removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents: 20557
diff changeset
   371
lemma stc_part_Ex1: "x:HFinite ==> EX! t. t \<in> SComplex &  x @= t"
14408
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   372
apply (drule stc_part_Ex, safe)
20559
2116b7a371c7 removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents: 20557
diff changeset
   373
apply (drule_tac [2] approx_sym, drule_tac [2] approx_sym, drule_tac [2] approx_sym)
2116b7a371c7 removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents: 20557
diff changeset
   374
apply (auto intro!: approx_unique_complex)
14408
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   375
done
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   376
20559
2116b7a371c7 removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents: 20557
diff changeset
   377
lemmas hcomplex_of_complex_approx_inverse =
2116b7a371c7 removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents: 20557
diff changeset
   378
  hcomplex_of_complex_HFinite_diff_Infinitesimal [THEN [2] approx_inverse]
14408
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   379
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   380
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   381
subsection{*Theorems About Monads*}
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   382
20559
2116b7a371c7 removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents: 20557
diff changeset
   383
lemma monad_zero_hcmod_iff: "(x \<in> monad 0) = (hcmod x:monad 0)"
2116b7a371c7 removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents: 20557
diff changeset
   384
by (simp add: Infinitesimal_monad_zero_iff [symmetric] Infinitesimal_hcmod_iff)
14408
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   385
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   386
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   387
subsection{*Theorems About Standard Part*}
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   388
20559
2116b7a371c7 removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents: 20557
diff changeset
   389
lemma stc_approx_self: "x \<in> HFinite ==> stc x @= x"
14408
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   390
apply (simp add: stc_def)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   391
apply (frule stc_part_Ex, safe)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   392
apply (rule someI2)
20559
2116b7a371c7 removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents: 20557
diff changeset
   393
apply (auto intro: approx_sym)
14408
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   394
done
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   395
20559
2116b7a371c7 removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents: 20557
diff changeset
   396
lemma stc_SComplex: "x \<in> HFinite ==> stc x \<in> SComplex"
14408
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   397
apply (simp add: stc_def)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   398
apply (frule stc_part_Ex, safe)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   399
apply (rule someI2)
20559
2116b7a371c7 removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents: 20557
diff changeset
   400
apply (auto intro: approx_sym)
14408
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   401
done
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   402
20559
2116b7a371c7 removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents: 20557
diff changeset
   403
lemma stc_HFinite: "x \<in> HFinite ==> stc x \<in> HFinite"
22883
005be8dafce0 hcomplex_of_hypreal abbreviates of_hypreal; removed redundant lemmas
huffman
parents: 22655
diff changeset
   404
by (erule stc_SComplex [THEN Standard_subset_HFinite [THEN subsetD]])
14408
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   405
20724
a1a8ba09e0ea add lemma stc_unique; shorten stc proofs
huffman
parents: 20686
diff changeset
   406
lemma stc_unique: "\<lbrakk>y \<in> SComplex; y \<approx> x\<rbrakk> \<Longrightarrow> stc x = y"
22883
005be8dafce0 hcomplex_of_hypreal abbreviates of_hypreal; removed redundant lemmas
huffman
parents: 22655
diff changeset
   407
apply (frule Standard_subset_HFinite [THEN subsetD])
20724
a1a8ba09e0ea add lemma stc_unique; shorten stc proofs
huffman
parents: 20686
diff changeset
   408
apply (drule (1) approx_HFinite)
a1a8ba09e0ea add lemma stc_unique; shorten stc proofs
huffman
parents: 20686
diff changeset
   409
apply (unfold stc_def)
14408
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   410
apply (rule some_equality)
20724
a1a8ba09e0ea add lemma stc_unique; shorten stc proofs
huffman
parents: 20686
diff changeset
   411
apply (auto intro: approx_unique_complex)
a1a8ba09e0ea add lemma stc_unique; shorten stc proofs
huffman
parents: 20686
diff changeset
   412
done
a1a8ba09e0ea add lemma stc_unique; shorten stc proofs
huffman
parents: 20686
diff changeset
   413
a1a8ba09e0ea add lemma stc_unique; shorten stc proofs
huffman
parents: 20686
diff changeset
   414
lemma stc_SComplex_eq [simp]: "x \<in> SComplex ==> stc x = x"
a1a8ba09e0ea add lemma stc_unique; shorten stc proofs
huffman
parents: 20686
diff changeset
   415
apply (erule stc_unique)
a1a8ba09e0ea add lemma stc_unique; shorten stc proofs
huffman
parents: 20686
diff changeset
   416
apply (rule approx_refl)
14408
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   417
done
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   418
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   419
lemma stc_hcomplex_of_complex:
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   420
     "stc (hcomplex_of_complex x) = hcomplex_of_complex x"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   421
by auto
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   422
20559
2116b7a371c7 removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents: 20557
diff changeset
   423
lemma stc_eq_approx:
2116b7a371c7 removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents: 20557
diff changeset
   424
     "[| x \<in> HFinite; y \<in> HFinite; stc x = stc y |] ==> x @= y"
2116b7a371c7 removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents: 20557
diff changeset
   425
by (auto dest!: stc_approx_self elim!: approx_trans3)
14408
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   426
20559
2116b7a371c7 removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents: 20557
diff changeset
   427
lemma approx_stc_eq:
2116b7a371c7 removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents: 20557
diff changeset
   428
     "[| x \<in> HFinite; y \<in> HFinite; x @= y |] ==> stc x = stc y"
2116b7a371c7 removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents: 20557
diff changeset
   429
by (blast intro: approx_trans approx_trans2 SComplex_approx_iff [THEN iffD1]
2116b7a371c7 removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents: 20557
diff changeset
   430
          dest: stc_approx_self stc_SComplex)
13957
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   431
20559
2116b7a371c7 removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents: 20557
diff changeset
   432
lemma stc_eq_approx_iff:
2116b7a371c7 removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents: 20557
diff changeset
   433
     "[| x \<in> HFinite; y \<in> HFinite|] ==> (x @= y) = (stc x = stc y)"
2116b7a371c7 removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents: 20557
diff changeset
   434
by (blast intro: approx_stc_eq stc_eq_approx)
14408
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   435
20559
2116b7a371c7 removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents: 20557
diff changeset
   436
lemma stc_Infinitesimal_add_SComplex:
2116b7a371c7 removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents: 20557
diff changeset
   437
     "[| x \<in> SComplex; e \<in> Infinitesimal |] ==> stc(x + e) = x"
20724
a1a8ba09e0ea add lemma stc_unique; shorten stc proofs
huffman
parents: 20686
diff changeset
   438
apply (erule stc_unique)
a1a8ba09e0ea add lemma stc_unique; shorten stc proofs
huffman
parents: 20686
diff changeset
   439
apply (erule Infinitesimal_add_approx_self)
14408
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   440
done
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   441
20559
2116b7a371c7 removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents: 20557
diff changeset
   442
lemma stc_Infinitesimal_add_SComplex2:
2116b7a371c7 removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents: 20557
diff changeset
   443
     "[| x \<in> SComplex; e \<in> Infinitesimal |] ==> stc(e + x) = x"
20724
a1a8ba09e0ea add lemma stc_unique; shorten stc proofs
huffman
parents: 20686
diff changeset
   444
apply (erule stc_unique)
a1a8ba09e0ea add lemma stc_unique; shorten stc proofs
huffman
parents: 20686
diff changeset
   445
apply (erule Infinitesimal_add_approx_self2)
14408
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   446
done
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   447
20559
2116b7a371c7 removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents: 20557
diff changeset
   448
lemma HFinite_stc_Infinitesimal_add:
2116b7a371c7 removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents: 20557
diff changeset
   449
     "x \<in> HFinite ==> \<exists>e \<in> Infinitesimal. x = stc(x) + e"
2116b7a371c7 removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents: 20557
diff changeset
   450
by (blast dest!: stc_approx_self [THEN approx_sym] bex_Infinitesimal_iff2 [THEN iffD2])
14408
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   451
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   452
lemma stc_add:
20559
2116b7a371c7 removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents: 20557
diff changeset
   453
     "[| x \<in> HFinite; y \<in> HFinite |] ==> stc (x + y) = stc(x) + stc(y)"
22883
005be8dafce0 hcomplex_of_hypreal abbreviates of_hypreal; removed redundant lemmas
huffman
parents: 22655
diff changeset
   454
by (simp add: stc_unique stc_SComplex stc_approx_self approx_add)
14408
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   455
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   456
lemma stc_number_of [simp]: "stc (number_of w) = number_of w"
22883
005be8dafce0 hcomplex_of_hypreal abbreviates of_hypreal; removed redundant lemmas
huffman
parents: 22655
diff changeset
   457
by (rule Standard_number_of [THEN stc_SComplex_eq])
14408
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   458
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   459
lemma stc_zero [simp]: "stc 0 = 0"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   460
by simp
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   461
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   462
lemma stc_one [simp]: "stc 1 = 1"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   463
by simp
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   464
20559
2116b7a371c7 removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents: 20557
diff changeset
   465
lemma stc_minus: "y \<in> HFinite ==> stc(-y) = -stc(y)"
20724
a1a8ba09e0ea add lemma stc_unique; shorten stc proofs
huffman
parents: 20686
diff changeset
   466
by (simp add: stc_unique stc_SComplex stc_approx_self approx_minus)
14408
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   467
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   468
lemma stc_diff: 
20559
2116b7a371c7 removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents: 20557
diff changeset
   469
     "[| x \<in> HFinite; y \<in> HFinite |] ==> stc (x-y) = stc(x) - stc(y)"
22883
005be8dafce0 hcomplex_of_hypreal abbreviates of_hypreal; removed redundant lemmas
huffman
parents: 22655
diff changeset
   470
by (simp add: stc_unique stc_SComplex stc_approx_self approx_diff)
14408
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   471
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   472
lemma stc_mult:
20559
2116b7a371c7 removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents: 20557
diff changeset
   473
     "[| x \<in> HFinite; y \<in> HFinite |]  
14408
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   474
               ==> stc (x * y) = stc(x) * stc(y)"
22883
005be8dafce0 hcomplex_of_hypreal abbreviates of_hypreal; removed redundant lemmas
huffman
parents: 22655
diff changeset
   475
by (simp add: stc_unique stc_SComplex stc_approx_self approx_mult_HFinite)
14408
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   476
20559
2116b7a371c7 removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents: 20557
diff changeset
   477
lemma stc_Infinitesimal: "x \<in> Infinitesimal ==> stc x = 0"
20724
a1a8ba09e0ea add lemma stc_unique; shorten stc proofs
huffman
parents: 20686
diff changeset
   478
by (simp add: stc_unique mem_infmal_iff)
14408
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   479
20559
2116b7a371c7 removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents: 20557
diff changeset
   480
lemma stc_not_Infinitesimal: "stc(x) \<noteq> 0 ==> x \<notin> Infinitesimal"
2116b7a371c7 removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents: 20557
diff changeset
   481
by (fast intro: stc_Infinitesimal)
14408
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   482
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   483
lemma stc_inverse:
20559
2116b7a371c7 removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents: 20557
diff changeset
   484
     "[| x \<in> HFinite; stc x \<noteq> 0 |]  
14408
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   485
      ==> stc(inverse x) = inverse (stc x)"
20724
a1a8ba09e0ea add lemma stc_unique; shorten stc proofs
huffman
parents: 20686
diff changeset
   486
apply (drule stc_not_Infinitesimal)
22883
005be8dafce0 hcomplex_of_hypreal abbreviates of_hypreal; removed redundant lemmas
huffman
parents: 22655
diff changeset
   487
apply (simp add: stc_unique stc_SComplex stc_approx_self approx_inverse)
14408
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   488
done
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   489
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   490
lemma stc_divide [simp]:
20559
2116b7a371c7 removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents: 20557
diff changeset
   491
     "[| x \<in> HFinite; y \<in> HFinite; stc y \<noteq> 0 |]  
14408
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   492
      ==> stc(x/y) = (stc x) / (stc y)"
20559
2116b7a371c7 removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents: 20557
diff changeset
   493
by (simp add: divide_inverse stc_mult stc_not_Infinitesimal HFinite_inverse stc_inverse)
14408
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   494
20559
2116b7a371c7 removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents: 20557
diff changeset
   495
lemma stc_idempotent [simp]: "x \<in> HFinite ==> stc(stc(x)) = stc(x)"
2116b7a371c7 removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents: 20557
diff changeset
   496
by (blast intro: stc_HFinite stc_approx_self approx_stc_eq)
14408
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   497
20559
2116b7a371c7 removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents: 20557
diff changeset
   498
lemma HFinite_HFinite_hcomplex_of_hypreal:
2116b7a371c7 removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents: 20557
diff changeset
   499
     "z \<in> HFinite ==> hcomplex_of_hypreal z \<in> HFinite"
21839
54018ed3b99d added lemmas about hRe, hIm, HComplex; removed all uses of star_n
huffman
parents: 21830
diff changeset
   500
by (simp add: hcomplex_HFinite_iff)
14408
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   501
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   502
lemma SComplex_SReal_hcomplex_of_hypreal:
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   503
     "x \<in> Reals ==>  hcomplex_of_hypreal x \<in> SComplex"
22883
005be8dafce0 hcomplex_of_hypreal abbreviates of_hypreal; removed redundant lemmas
huffman
parents: 22655
diff changeset
   504
apply (rule Standard_of_hypreal)
005be8dafce0 hcomplex_of_hypreal abbreviates of_hypreal; removed redundant lemmas
huffman
parents: 22655
diff changeset
   505
apply (simp add: Reals_eq_Standard)
005be8dafce0 hcomplex_of_hypreal abbreviates of_hypreal; removed redundant lemmas
huffman
parents: 22655
diff changeset
   506
done
14408
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   507
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   508
lemma stc_hcomplex_of_hypreal: 
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   509
 "z \<in> HFinite ==> stc(hcomplex_of_hypreal z) = hcomplex_of_hypreal (st z)"
20724
a1a8ba09e0ea add lemma stc_unique; shorten stc proofs
huffman
parents: 20686
diff changeset
   510
apply (rule stc_unique)
a1a8ba09e0ea add lemma stc_unique; shorten stc proofs
huffman
parents: 20686
diff changeset
   511
apply (rule SComplex_SReal_hcomplex_of_hypreal)
a1a8ba09e0ea add lemma stc_unique; shorten stc proofs
huffman
parents: 20686
diff changeset
   512
apply (erule st_SReal)
a1a8ba09e0ea add lemma stc_unique; shorten stc proofs
huffman
parents: 20686
diff changeset
   513
apply (simp add: hcomplex_of_hypreal_approx_iff st_approx_self)
14408
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   514
done
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   515
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   516
(*
20559
2116b7a371c7 removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents: 20557
diff changeset
   517
Goal "x \<in> HFinite ==> hcmod(stc x) = st(hcmod x)"
2116b7a371c7 removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents: 20557
diff changeset
   518
by (dtac stc_approx_self 1)
2116b7a371c7 removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents: 20557
diff changeset
   519
by (auto_tac (claset(),simpset() addsimps [bex_Infinitesimal_iff2 RS sym]));
14408
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   520
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   521
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   522
approx_hcmod_add_hcmod
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   523
*)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   524
20559
2116b7a371c7 removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents: 20557
diff changeset
   525
lemma Infinitesimal_hcnj_iff [simp]:
2116b7a371c7 removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents: 20557
diff changeset
   526
     "(hcnj z \<in> Infinitesimal) = (z \<in> Infinitesimal)"
2116b7a371c7 removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents: 20557
diff changeset
   527
by (simp add: Infinitesimal_hcmod_iff)
14408
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   528
20559
2116b7a371c7 removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents: 20557
diff changeset
   529
lemma Infinitesimal_hcomplex_of_hypreal_epsilon [simp]:
2116b7a371c7 removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents: 20557
diff changeset
   530
     "hcomplex_of_hypreal epsilon \<in> Infinitesimal"
2116b7a371c7 removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents: 20557
diff changeset
   531
by (simp add: Infinitesimal_hcmod_iff)
14408
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   532
13957
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   533
end