src/HOL/Complex/NSCA.thy
author huffman
Sat, 16 Sep 2006 23:46:38 +0200
changeset 20557 81dd3679f92c
parent 20552 2c31dd358c21
child 20559 2116b7a371c7
permissions -rw-r--r--
complex_of_real abbreviates of_real::real=>complex; cmod abbreviates norm::complex=>real; removed several redundant lemmas
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
15140
322485b816ac import -> imports
nipkow
parents: 15131
diff changeset
     9
imports NSComplex
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
19765
dfe940911617 misc cleanup;
wenzelm
parents: 17429
diff changeset
    12
definition
13957
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    13
14653
0848ab6fe5fc constdefs: proper order;
wenzelm
parents: 14469
diff changeset
    14
   CInfinitesimal  :: "hcomplex set"
19765
dfe940911617 misc cleanup;
wenzelm
parents: 17429
diff changeset
    15
   "CInfinitesimal = {x. \<forall>r \<in> Reals. 0 < r --> hcmod x < r}"
14653
0848ab6fe5fc constdefs: proper order;
wenzelm
parents: 14469
diff changeset
    16
14408
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
    17
    capprox    :: "[hcomplex,hcomplex] => bool"  (infixl "@c=" 50)  
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
    18
      --{*the ``infinitely close'' relation*}
19765
dfe940911617 misc cleanup;
wenzelm
parents: 17429
diff changeset
    19
      "x @c= y = ((x - y) \<in> CInfinitesimal)"     
13957
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    20
  
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    21
   (* standard complex numbers reagarded as an embedded subset of NS complex *)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    22
   SComplex  :: "hcomplex set"
19765
dfe940911617 misc cleanup;
wenzelm
parents: 17429
diff changeset
    23
   "SComplex = {x. \<exists>r. x = hcomplex_of_complex r}"
13957
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    24
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    25
   CFinite :: "hcomplex set"
19765
dfe940911617 misc cleanup;
wenzelm
parents: 17429
diff changeset
    26
   "CFinite = {x. \<exists>r \<in> Reals. hcmod x < r}"
13957
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    27
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    28
   CInfinite :: "hcomplex set"
19765
dfe940911617 misc cleanup;
wenzelm
parents: 17429
diff changeset
    29
   "CInfinite = {x. \<forall>r \<in> Reals. r < hcmod x}"
13957
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    30
14408
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
    31
   stc :: "hcomplex => hcomplex"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
    32
    --{* standard part map*}
19765
dfe940911617 misc cleanup;
wenzelm
parents: 17429
diff changeset
    33
   "stc x = (SOME r. x \<in> CFinite & r:SComplex & r @c= x)"
13957
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    34
14408
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
    35
   cmonad    :: "hcomplex => hcomplex set"
19765
dfe940911617 misc cleanup;
wenzelm
parents: 17429
diff changeset
    36
   "cmonad x = {y. x @c= y}"
13957
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    37
14408
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
    38
   cgalaxy   :: "hcomplex => hcomplex set"
19765
dfe940911617 misc cleanup;
wenzelm
parents: 17429
diff changeset
    39
   "cgalaxy x = {y. (x - y) \<in> CFinite}"
14408
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
    40
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
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
    43
subsection{*Closure Laws for SComplex, the Standard Complex Numbers*}
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
    44
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
    45
lemma SComplex_add: "[| x \<in> SComplex; y \<in> SComplex |] ==> x + y \<in> SComplex"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
    46
apply (simp add: SComplex_def, safe)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
    47
apply (rule_tac x = "r + ra" in exI, simp)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
    48
done
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_mult: "[| x \<in> SComplex; y \<in> SComplex |] ==> x * y \<in> SComplex"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
    51
apply (simp add: SComplex_def, safe)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
    52
apply (rule_tac x = "r * ra" in exI, simp)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
    53
done
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
    54
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
    55
lemma SComplex_inverse: "x \<in> SComplex ==> inverse x \<in> SComplex"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
    56
apply (simp add: SComplex_def)
17373
27509e72f29e removed duplicated lemmas; convert more proofs to transfer principle
huffman
parents: 17318
diff changeset
    57
apply (blast intro: star_of_inverse [symmetric])
14408
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
    58
done
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
    59
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
    60
lemma SComplex_divide: "[| x \<in> SComplex;  y \<in> SComplex |] ==> x/y \<in> SComplex"
14430
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14408
diff changeset
    61
by (simp add: SComplex_mult SComplex_inverse divide_inverse)
14408
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
    62
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
    63
lemma SComplex_minus: "x \<in> SComplex ==> -x \<in> SComplex"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
    64
apply (simp add: SComplex_def)
17373
27509e72f29e removed duplicated lemmas; convert more proofs to transfer principle
huffman
parents: 17318
diff changeset
    65
apply (blast intro: star_of_minus [symmetric])
14408
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
    66
done
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
    67
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
    68
lemma SComplex_minus_iff [simp]: "(-x \<in> SComplex) = (x \<in> SComplex)"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
    69
apply auto
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
    70
apply (erule_tac [2] SComplex_minus)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
    71
apply (drule SComplex_minus, auto)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
    72
done
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
    73
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
    74
lemma SComplex_diff: "[| x \<in> SComplex; y \<in> SComplex |] ==> x - y \<in> SComplex"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
    75
by (simp add: diff_minus SComplex_add) 
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
lemma SComplex_add_cancel:
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
    78
     "[| x + y \<in> SComplex; y \<in> SComplex |] ==> x \<in> SComplex"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
    79
by (drule SComplex_diff, assumption, simp)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
    80
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
    81
lemma SReal_hcmod_hcomplex_of_complex [simp]:
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
    82
     "hcmod (hcomplex_of_complex r) \<in> Reals"
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17300
diff changeset
    83
by (auto simp add: hcmod SReal_def star_of_def)
14408
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
lemma SReal_hcmod_number_of [simp]: "hcmod (number_of w ::hcomplex) \<in> Reals"
17373
27509e72f29e removed duplicated lemmas; convert more proofs to transfer principle
huffman
parents: 17318
diff changeset
    86
apply (subst star_of_number_of [symmetric])
14408
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
    87
apply (rule SReal_hcmod_hcomplex_of_complex)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
    88
done
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
    89
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
    90
lemma SReal_hcmod_SComplex: "x \<in> SComplex ==> hcmod x \<in> Reals"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
    91
by (auto simp add: SComplex_def)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
    92
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
    93
lemma SComplex_hcomplex_of_complex [simp]: "hcomplex_of_complex x \<in> SComplex"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
    94
by (simp add: SComplex_def)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
    95
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
    96
lemma SComplex_number_of [simp]: "(number_of w ::hcomplex) \<in> SComplex"
17373
27509e72f29e removed duplicated lemmas; convert more proofs to transfer principle
huffman
parents: 17318
diff changeset
    97
apply (subst star_of_number_of [symmetric])
14408
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
    98
apply (rule SComplex_hcomplex_of_complex)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
    99
done
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 SComplex_divide_number_of:
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   102
     "r \<in> SComplex ==> r/(number_of w::hcomplex) \<in> SComplex"
14430
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14408
diff changeset
   103
apply (simp only: divide_inverse)
14408
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   104
apply (blast intro!: SComplex_number_of SComplex_mult SComplex_inverse)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   105
done
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   106
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   107
lemma SComplex_UNIV_complex:
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   108
     "{x. hcomplex_of_complex x \<in> SComplex} = (UNIV::complex set)"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   109
by (simp add: SComplex_def)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   110
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   111
lemma SComplex_iff: "(x \<in> SComplex) = (\<exists>y. x = hcomplex_of_complex y)"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   112
by (simp add: SComplex_def)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   113
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   114
lemma hcomplex_of_complex_image:
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   115
     "hcomplex_of_complex `(UNIV::complex set) = SComplex"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   116
by (auto simp add: SComplex_def)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   117
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   118
lemma inv_hcomplex_of_complex_image: "inv hcomplex_of_complex `SComplex = UNIV"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   119
apply (auto simp add: SComplex_def)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   120
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
   121
done
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   122
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   123
lemma SComplex_hcomplex_of_complex_image: 
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   124
      "[| \<exists>x. x: P; P \<le> SComplex |] ==> \<exists>Q. P = hcomplex_of_complex ` Q"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   125
apply (simp add: SComplex_def, blast)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   126
done
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   127
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   128
lemma SComplex_SReal_dense:
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   129
     "[| x \<in> SComplex; y \<in> SComplex; hcmod x < hcmod y  
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   130
      |] ==> \<exists>r \<in> Reals. hcmod x< r & r < hcmod y"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   131
apply (auto intro: SReal_dense simp add: SReal_hcmod_SComplex)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   132
done
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   133
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   134
lemma SComplex_hcmod_SReal: 
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   135
      "z \<in> SComplex ==> hcmod z \<in> Reals"
17429
e8d6ed3aacfe merged Transfer.thy and StarType.thy into StarDef.thy; renamed Ifun2_of to starfun2; cleaned up
huffman
parents: 17373
diff changeset
   136
by (auto simp add: SComplex_def SReal_def hcmod_def)
14408
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
lemma SComplex_zero [simp]: "0 \<in> SComplex"
17373
27509e72f29e removed duplicated lemmas; convert more proofs to transfer principle
huffman
parents: 17318
diff changeset
   139
by (simp add: SComplex_def)
14408
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
lemma SComplex_one [simp]: "1 \<in> SComplex"
17373
27509e72f29e removed duplicated lemmas; convert more proofs to transfer principle
huffman
parents: 17318
diff changeset
   142
by (simp add: SComplex_def)
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
(*
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   145
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
   146
by (res_inst_tac [("z","z")] eq_Abs_hcomplex 1);
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   147
by (auto_tac (claset(),simpset() addsimps [hcmod,hypreal_of_real_def,
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   148
    hcomplex_of_complex_def,cmod_def]));
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   149
*)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   150
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   151
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   152
subsection{*The Finite Elements form a Subring*}
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   153
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   154
lemma CFinite_add: "[|x \<in> CFinite; y \<in> CFinite|] ==> (x+y) \<in> CFinite"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   155
apply (simp add: CFinite_def)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   156
apply (blast intro!: SReal_add hcmod_add_less)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   157
done
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   158
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   159
lemma CFinite_mult: "[|x \<in> CFinite; y \<in> CFinite|] ==> x*y \<in> CFinite"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   160
apply (simp add: CFinite_def)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   161
apply (blast intro!: SReal_mult hcmod_mult_less)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   162
done
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   163
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   164
lemma CFinite_minus_iff [simp]: "(-x \<in> CFinite) = (x \<in> CFinite)"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   165
by (simp add: CFinite_def)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   166
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   167
lemma SComplex_subset_CFinite [simp]: "SComplex \<le> CFinite"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   168
apply (auto simp add: SComplex_def CFinite_def)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   169
apply (rule_tac x = "1 + hcmod (hcomplex_of_complex r) " in bexI)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   170
apply (auto intro: SReal_add)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   171
done
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   172
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   173
lemma HFinite_hcmod_hcomplex_of_complex [simp]:
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   174
     "hcmod (hcomplex_of_complex r) \<in> HFinite"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   175
by (auto intro!: SReal_subset_HFinite [THEN subsetD])
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   176
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   177
lemma CFinite_hcomplex_of_complex [simp]: "hcomplex_of_complex x \<in> CFinite"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   178
by (auto intro!: SComplex_subset_CFinite [THEN subsetD])
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   179
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   180
lemma CFiniteD: "x \<in> CFinite ==> \<exists>t \<in> Reals. hcmod x < t"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   181
by (simp add: CFinite_def)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   182
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   183
lemma CFinite_hcmod_iff: "(x \<in> CFinite) = (hcmod x \<in> HFinite)"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   184
by (simp add: CFinite_def HFinite_def)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   185
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   186
lemma CFinite_number_of [simp]: "number_of w \<in> CFinite"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   187
by (rule SComplex_number_of [THEN SComplex_subset_CFinite [THEN subsetD]])
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   188
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   189
lemma CFinite_bounded: "[|x \<in> CFinite; y \<le> hcmod x; 0 \<le> y |] ==> y: HFinite"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   190
by (auto intro: HFinite_bounded simp add: CFinite_hcmod_iff)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   191
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   192
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   193
subsection{*The Complex Infinitesimals form a Subring*}
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   194
	 
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   195
lemma CInfinitesimal_zero [iff]: "0 \<in> CInfinitesimal"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   196
by (simp add: CInfinitesimal_def)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   197
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   198
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
   199
by auto
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   200
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   201
lemma CInfinitesimal_hcmod_iff: 
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   202
   "(z \<in> CInfinitesimal) = (hcmod z \<in> Infinitesimal)"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   203
by (simp add: CInfinitesimal_def Infinitesimal_def)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   204
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   205
lemma one_not_CInfinitesimal [simp]: "1 \<notin> CInfinitesimal"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   206
by (simp add: CInfinitesimal_hcmod_iff)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   207
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   208
lemma CInfinitesimal_add:
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   209
     "[| x \<in> CInfinitesimal; y \<in> CInfinitesimal |] ==> (x+y) \<in> CInfinitesimal"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   210
apply (auto simp add: CInfinitesimal_hcmod_iff)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   211
apply (rule hrabs_le_Infinitesimal)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   212
apply (rule_tac y = "hcmod y" in Infinitesimal_add, auto)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   213
done
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   214
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   215
lemma CInfinitesimal_minus_iff [simp]:
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   216
     "(-x:CInfinitesimal) = (x:CInfinitesimal)"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   217
by (simp add: CInfinitesimal_def)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   218
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   219
lemma CInfinitesimal_diff:
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   220
     "[| x \<in> CInfinitesimal;  y \<in> CInfinitesimal |] ==> x-y \<in> CInfinitesimal"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   221
by (simp add: diff_minus CInfinitesimal_add)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   222
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   223
lemma CInfinitesimal_mult:
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   224
     "[| x \<in> CInfinitesimal; y \<in> CInfinitesimal |] ==> x * y \<in> CInfinitesimal"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   225
by (auto intro: Infinitesimal_mult simp add: CInfinitesimal_hcmod_iff hcmod_mult)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   226
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   227
lemma CInfinitesimal_CFinite_mult:
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   228
     "[| x \<in> CInfinitesimal; y \<in> CFinite |] ==> (x * y) \<in> CInfinitesimal"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   229
by (auto intro: Infinitesimal_HFinite_mult simp add: CInfinitesimal_hcmod_iff CFinite_hcmod_iff hcmod_mult)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   230
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   231
lemma CInfinitesimal_CFinite_mult2:
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   232
     "[| x \<in> CInfinitesimal; y \<in> CFinite |] ==> (y * x) \<in> CInfinitesimal"
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17300
diff changeset
   233
by (auto dest: CInfinitesimal_CFinite_mult simp add: mult_commute)
14408
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   234
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   235
lemma CInfinite_hcmod_iff: "(z \<in> CInfinite) = (hcmod z \<in> HInfinite)"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   236
by (simp add: CInfinite_def HInfinite_def)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   237
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   238
lemma CInfinite_inverse_CInfinitesimal:
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   239
     "x \<in> CInfinite ==> inverse x \<in> CInfinitesimal"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   240
by (auto intro: HInfinite_inverse_Infinitesimal simp add: CInfinitesimal_hcmod_iff CInfinite_hcmod_iff hcmod_hcomplex_inverse)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   241
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   242
lemma CInfinite_mult: "[|x \<in> CInfinite; y \<in> CInfinite|] ==> (x*y): CInfinite"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   243
by (auto intro: HInfinite_mult simp add: CInfinite_hcmod_iff hcmod_mult)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   244
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   245
lemma CInfinite_minus_iff [simp]: "(-x \<in> CInfinite) = (x \<in> CInfinite)"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   246
by (simp add: CInfinite_def)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   247
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   248
lemma CFinite_sum_squares:
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   249
     "[|a \<in> CFinite; b \<in> CFinite; c \<in> CFinite|]   
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   250
      ==> a*a + b*b + c*c \<in> CFinite"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   251
by (auto intro: CFinite_mult CFinite_add)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   252
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   253
lemma not_CInfinitesimal_not_zero: "x \<notin> CInfinitesimal ==> x \<noteq> 0"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   254
by auto
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   255
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   256
lemma not_CInfinitesimal_not_zero2: "x \<in> CFinite - CInfinitesimal ==> x \<noteq> 0"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   257
by auto
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   258
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   259
lemma CFinite_diff_CInfinitesimal_hcmod:
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   260
     "x \<in> CFinite - CInfinitesimal ==> hcmod x \<in> HFinite - Infinitesimal"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   261
by (simp add: CFinite_hcmod_iff CInfinitesimal_hcmod_iff)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   262
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   263
lemma hcmod_less_CInfinitesimal:
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   264
     "[| e \<in> CInfinitesimal; hcmod x < hcmod e |] ==> x \<in> CInfinitesimal"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   265
by (auto intro: hrabs_less_Infinitesimal simp add: CInfinitesimal_hcmod_iff)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   266
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   267
lemma hcmod_le_CInfinitesimal:
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   268
     "[| e \<in> CInfinitesimal; hcmod x \<le> hcmod e |] ==> x \<in> CInfinitesimal"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   269
by (auto intro: hrabs_le_Infinitesimal simp add: CInfinitesimal_hcmod_iff)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   270
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   271
lemma CInfinitesimal_interval:
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   272
     "[| e \<in> CInfinitesimal;  
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   273
          e' \<in> CInfinitesimal;  
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   274
          hcmod e' < hcmod x ; hcmod x < hcmod e  
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   275
       |] ==> x \<in> CInfinitesimal"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   276
by (auto intro: Infinitesimal_interval simp add: CInfinitesimal_hcmod_iff)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   277
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   278
lemma CInfinitesimal_interval2:
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   279
     "[| e \<in> CInfinitesimal;  
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   280
         e' \<in> CInfinitesimal;  
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   281
         hcmod e' \<le> hcmod x ; hcmod x \<le> hcmod e  
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   282
      |] ==> x \<in> CInfinitesimal"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   283
by (auto intro: Infinitesimal_interval2 simp add: CInfinitesimal_hcmod_iff)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   284
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   285
lemma not_CInfinitesimal_mult:
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   286
     "[| x \<notin> CInfinitesimal;  y \<notin> CInfinitesimal|] ==> (x*y) \<notin> CInfinitesimal"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   287
apply (auto simp add: CInfinitesimal_hcmod_iff hcmod_mult)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   288
apply (drule not_Infinitesimal_mult, auto)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   289
done
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   290
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   291
lemma CInfinitesimal_mult_disj:
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   292
     "x*y \<in> CInfinitesimal ==> x \<in> CInfinitesimal | y \<in> CInfinitesimal"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   293
by (auto dest: Infinitesimal_mult_disj simp add: CInfinitesimal_hcmod_iff hcmod_mult)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   294
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   295
lemma CFinite_CInfinitesimal_diff_mult:
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   296
     "[| x \<in> CFinite - CInfinitesimal; y \<in> CFinite - CInfinitesimal |] 
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   297
      ==> x*y \<in> CFinite - CInfinitesimal"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   298
by (blast dest: CFinite_mult not_CInfinitesimal_mult)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   299
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   300
lemma CInfinitesimal_subset_CFinite: "CInfinitesimal \<le> CFinite"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   301
by (auto intro: Infinitesimal_subset_HFinite [THEN subsetD]
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   302
         simp add: CInfinitesimal_hcmod_iff CFinite_hcmod_iff)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   303
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   304
lemma CInfinitesimal_hcomplex_of_complex_mult:
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   305
     "x \<in> CInfinitesimal ==> x * hcomplex_of_complex r \<in> CInfinitesimal"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   306
by (auto intro!: Infinitesimal_HFinite_mult simp add: CInfinitesimal_hcmod_iff hcmod_mult)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   307
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   308
lemma CInfinitesimal_hcomplex_of_complex_mult2:
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   309
     "x \<in> CInfinitesimal ==> hcomplex_of_complex r * x \<in> CInfinitesimal"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   310
by (auto intro!: Infinitesimal_HFinite_mult2 simp add: CInfinitesimal_hcmod_iff hcmod_mult)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   311
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   312
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   313
subsection{*The ``Infinitely Close'' Relation*}
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   314
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   315
(*
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   316
Goalw [capprox_def,approx_def] "(z @c= w) = (hcmod z @= hcmod w)"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   317
by (auto_tac (claset(),simpset() addsimps [CInfinitesimal_hcmod_iff]));
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   318
*)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   319
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   320
lemma mem_cinfmal_iff: "x:CInfinitesimal = (x @c= 0)"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   321
by (simp add: CInfinitesimal_hcmod_iff capprox_def)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   322
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   323
lemma capprox_minus_iff: "(x @c= y) = (x + -y @c= 0)"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   324
by (simp add: capprox_def diff_minus)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   325
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   326
lemma capprox_minus_iff2: "(x @c= y) = (-y + x @c= 0)"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   327
by (simp add: capprox_def diff_minus add_commute)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   328
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   329
lemma capprox_refl [simp]: "x @c= x"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   330
by (simp add: capprox_def)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   331
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   332
lemma capprox_sym: "x @c= y ==> y @c= x"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   333
by (simp add: capprox_def CInfinitesimal_def hcmod_diff_commute)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   334
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   335
lemma capprox_trans: "[| x @c= y; y @c= z |] ==> x @c= z"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   336
apply (simp add: capprox_def)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   337
apply (drule CInfinitesimal_add, assumption)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   338
apply (simp add: diff_minus)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   339
done
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   340
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   341
lemma capprox_trans2: "[| r @c= x; s @c= x |] ==> r @c= s"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   342
by (blast intro: capprox_sym capprox_trans)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   343
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   344
lemma capprox_trans3: "[| x @c= r; x @c= s|] ==> r @c= s"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   345
by (blast intro: capprox_sym capprox_trans)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   346
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   347
lemma number_of_capprox_reorient [simp]:
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   348
     "(number_of w @c= x) = (x @c= number_of w)"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   349
by (blast intro: capprox_sym)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   350
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   351
lemma CInfinitesimal_capprox_minus: "(x-y \<in> CInfinitesimal) = (x @c= y)"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   352
by (simp add: diff_minus capprox_minus_iff [symmetric] mem_cinfmal_iff)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   353
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   354
lemma capprox_monad_iff: "(x @c= y) = (cmonad(x)=cmonad(y))"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   355
by (auto simp add: cmonad_def dest: capprox_sym elim!: capprox_trans equalityCE)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   356
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   357
lemma Infinitesimal_capprox:
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   358
     "[| x \<in> CInfinitesimal; y \<in> CInfinitesimal |] ==> x @c= y"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   359
apply (simp add: mem_cinfmal_iff)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   360
apply (blast intro: capprox_trans capprox_sym)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   361
done
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
lemma capprox_add: "[| a @c= b; c @c= d |] ==> a+c @c= b+d"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   364
apply (simp add: capprox_def diff_minus) 
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   365
apply (rule minus_add_distrib [THEN ssubst])
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   366
apply (rule add_assoc [THEN ssubst])
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   367
apply (rule_tac b1 = c in add_left_commute [THEN subst])
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   368
apply (rule add_assoc [THEN subst])
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   369
apply (blast intro: CInfinitesimal_add)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   370
done
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   371
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   372
lemma capprox_minus: "a @c= b ==> -a @c= -b"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   373
apply (rule capprox_minus_iff [THEN iffD2, THEN capprox_sym])
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   374
apply (drule capprox_minus_iff [THEN iffD1])
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   375
apply (simp add: add_commute)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   376
done
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   377
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   378
lemma capprox_minus2: "-a @c= -b ==> a @c= b"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   379
by (auto dest: capprox_minus)
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
lemma capprox_minus_cancel [simp]: "(-a @c= -b) = (a @c= b)"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   382
by (blast intro: capprox_minus capprox_minus2)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   383
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   384
lemma capprox_add_minus: "[| a @c= b; c @c= d |] ==> a + -c @c= b + -d"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   385
by (blast intro!: capprox_add capprox_minus)
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
lemma capprox_mult1: 
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   388
      "[| a @c= b; c \<in> CFinite|] ==> a*c @c= b*c"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   389
apply (simp add: capprox_def diff_minus)
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17300
diff changeset
   390
apply (simp only: CInfinitesimal_CFinite_mult minus_mult_left left_distrib [symmetric])
14408
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   391
done
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   392
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   393
lemma capprox_mult2: "[|a @c= b; c \<in> CFinite|] ==> c*a @c= c*b"
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17300
diff changeset
   394
by (simp add: capprox_mult1 mult_commute)
14408
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   395
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   396
lemma capprox_mult_subst:
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   397
     "[|u @c= v*x; x @c= y; v \<in> CFinite|] ==> u @c= v*y"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   398
by (blast intro: capprox_mult2 capprox_trans)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   399
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   400
lemma capprox_mult_subst2:
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   401
     "[| u @c= x*v; x @c= y; v \<in> CFinite |] ==> u @c= y*v"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   402
by (blast intro: capprox_mult1 capprox_trans)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   403
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   404
lemma capprox_mult_subst_SComplex:
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   405
     "[| u @c= x*hcomplex_of_complex v; x @c= y |] 
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   406
      ==> u @c= y*hcomplex_of_complex v"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   407
by (auto intro: capprox_mult_subst2)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   408
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   409
lemma capprox_eq_imp: "a = b ==> a @c= b"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   410
by (simp add: capprox_def)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   411
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   412
lemma CInfinitesimal_minus_capprox: "x \<in> CInfinitesimal ==> -x @c= x"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   413
by (blast intro: CInfinitesimal_minus_iff [THEN iffD2] mem_cinfmal_iff [THEN iffD1] capprox_trans2)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   414
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   415
lemma bex_CInfinitesimal_iff: "(\<exists>y \<in> CInfinitesimal. x - z = y) = (x @c= z)"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   416
by (unfold capprox_def, blast)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   417
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   418
lemma bex_CInfinitesimal_iff2: "(\<exists>y \<in> CInfinitesimal. x = z + y) = (x @c= z)"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   419
by (simp add: bex_CInfinitesimal_iff [symmetric], force)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   420
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   421
lemma CInfinitesimal_add_capprox:
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   422
     "[| y \<in> CInfinitesimal; x + y = z |] ==> x @c= z"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   423
apply (rule bex_CInfinitesimal_iff [THEN iffD1])
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   424
apply (drule CInfinitesimal_minus_iff [THEN iffD2])
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   425
apply (simp add: eq_commute compare_rls)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   426
done
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   427
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   428
lemma CInfinitesimal_add_capprox_self: "y \<in> CInfinitesimal ==> x @c= x + y"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   429
apply (rule bex_CInfinitesimal_iff [THEN iffD1])
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   430
apply (drule CInfinitesimal_minus_iff [THEN iffD2])
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   431
apply (simp add: eq_commute compare_rls)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   432
done
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   433
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   434
lemma CInfinitesimal_add_capprox_self2: "y \<in> CInfinitesimal ==> x @c= y + x"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   435
by (auto dest: CInfinitesimal_add_capprox_self simp add: add_commute)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   436
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   437
lemma CInfinitesimal_add_minus_capprox_self:
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   438
     "y \<in> CInfinitesimal ==> x @c= x + -y"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   439
by (blast intro!: CInfinitesimal_add_capprox_self CInfinitesimal_minus_iff [THEN iffD2])
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   440
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   441
lemma CInfinitesimal_add_cancel:
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   442
     "[| y \<in> CInfinitesimal; x+y @c= z|] ==> x @c= z"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   443
apply (drule_tac x = x in CInfinitesimal_add_capprox_self [THEN capprox_sym])
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   444
apply (erule capprox_trans3 [THEN capprox_sym], assumption)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   445
done
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   446
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   447
lemma CInfinitesimal_add_right_cancel:
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   448
     "[| y \<in> CInfinitesimal; x @c= z + y|] ==> x @c= z"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   449
apply (drule_tac x = z in CInfinitesimal_add_capprox_self2 [THEN capprox_sym])
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   450
apply (erule capprox_trans3 [THEN capprox_sym])
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   451
apply (simp add: add_commute)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   452
apply (erule capprox_sym)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   453
done
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   454
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   455
lemma capprox_add_left_cancel: "d + b  @c= d + c ==> b @c= c"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   456
apply (drule capprox_minus_iff [THEN iffD1])
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   457
apply (simp add: minus_add_distrib capprox_minus_iff [symmetric] add_ac)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   458
done
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   459
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   460
lemma capprox_add_right_cancel: "b + d @c= c + d ==> b @c= c"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   461
apply (rule capprox_add_left_cancel)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   462
apply (simp add: add_commute)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   463
done
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   464
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   465
lemma capprox_add_mono1: "b @c= c ==> d + b @c= d + c"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   466
apply (rule capprox_minus_iff [THEN iffD2])
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   467
apply (simp add: capprox_minus_iff [symmetric] add_ac)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   468
done
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   469
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   470
lemma capprox_add_mono2: "b @c= c ==> b + a @c= c + a"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   471
apply (simp (no_asm_simp) add: add_commute capprox_add_mono1)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   472
done
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   473
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   474
lemma capprox_add_left_iff [iff]: "(a + b @c= a + c) = (b @c= c)"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   475
by (fast elim: capprox_add_left_cancel capprox_add_mono1)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   476
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   477
lemma capprox_add_right_iff [iff]: "(b + a @c= c + a) = (b @c= c)"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   478
by (simp add: add_commute)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   479
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   480
lemma capprox_CFinite: "[| x \<in> CFinite; x @c= y |] ==> y \<in> CFinite"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   481
apply (drule bex_CInfinitesimal_iff2 [THEN iffD2], safe)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   482
apply (drule CInfinitesimal_subset_CFinite [THEN subsetD, THEN CFinite_minus_iff [THEN iffD2]])
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   483
apply (drule CFinite_add)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   484
apply (assumption, auto)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   485
done
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   486
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   487
lemma capprox_hcomplex_of_complex_CFinite:
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   488
     "x @c= hcomplex_of_complex D ==> x \<in> CFinite"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   489
by (rule capprox_sym [THEN [2] capprox_CFinite], auto)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   490
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   491
lemma capprox_mult_CFinite:
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   492
     "[|a @c= b; c @c= d; b \<in> CFinite; d \<in> CFinite|] ==> a*c @c= b*d"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   493
apply (rule capprox_trans)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   494
apply (rule_tac [2] capprox_mult2)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   495
apply (rule capprox_mult1)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   496
prefer 2 apply (blast intro: capprox_CFinite capprox_sym, auto)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   497
done
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   498
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   499
lemma capprox_mult_hcomplex_of_complex:
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   500
     "[|a @c= hcomplex_of_complex b; c @c= hcomplex_of_complex d |]  
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   501
      ==> a*c @c= hcomplex_of_complex b * hcomplex_of_complex d"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   502
apply (blast intro!: capprox_mult_CFinite capprox_hcomplex_of_complex_CFinite CFinite_hcomplex_of_complex)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   503
done
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   504
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   505
lemma capprox_SComplex_mult_cancel_zero:
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   506
     "[| a \<in> SComplex; a \<noteq> 0; a*x @c= 0 |] ==> x @c= 0"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   507
apply (drule SComplex_inverse [THEN SComplex_subset_CFinite [THEN subsetD]])
15013
34264f5e4691 new treatment of binary numerals
paulson
parents: 14653
diff changeset
   508
apply (auto dest: capprox_mult2 simp add: mult_assoc [symmetric])
14408
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   509
done
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   510
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   511
lemma capprox_mult_SComplex1: "[| a \<in> SComplex; x @c= 0 |] ==> x*a @c= 0"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   512
by (auto dest: SComplex_subset_CFinite [THEN subsetD] capprox_mult1)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   513
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   514
lemma capprox_mult_SComplex2: "[| a \<in> SComplex; x @c= 0 |] ==> a*x @c= 0"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   515
by (auto dest: SComplex_subset_CFinite [THEN subsetD] capprox_mult2)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   516
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   517
lemma capprox_mult_SComplex_zero_cancel_iff [simp]:
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   518
     "[|a \<in> SComplex; a \<noteq> 0 |] ==> (a*x @c= 0) = (x @c= 0)"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   519
by (blast intro: capprox_SComplex_mult_cancel_zero capprox_mult_SComplex2)
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
lemma capprox_SComplex_mult_cancel:
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   522
     "[| a \<in> SComplex; a \<noteq> 0; a* w @c= a*z |] ==> w @c= z"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   523
apply (drule SComplex_inverse [THEN SComplex_subset_CFinite [THEN subsetD]])
15013
34264f5e4691 new treatment of binary numerals
paulson
parents: 14653
diff changeset
   524
apply (auto dest: capprox_mult2 simp add: mult_assoc [symmetric])
14408
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   525
done
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   526
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   527
lemma capprox_SComplex_mult_cancel_iff1 [simp]:
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   528
     "[| a \<in> SComplex; a \<noteq> 0|] ==> (a* w @c= a*z) = (w @c= z)"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   529
by (auto intro!: capprox_mult2 SComplex_subset_CFinite [THEN subsetD]
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   530
            intro: capprox_SComplex_mult_cancel)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   531
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   532
lemma capprox_hcmod_approx_zero: "(x @c= y) = (hcmod (y - x) @= 0)"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   533
apply (rule capprox_minus_iff [THEN ssubst])
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   534
apply (simp add: capprox_def CInfinitesimal_hcmod_iff mem_infmal_iff diff_minus [symmetric] hcmod_diff_commute)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   535
done
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   536
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   537
lemma capprox_approx_zero_iff: "(x @c= 0) = (hcmod x @= 0)"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   538
by (simp add: capprox_hcmod_approx_zero)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   539
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   540
lemma capprox_minus_zero_cancel_iff [simp]: "(-x @c= 0) = (x @c= 0)"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   541
by (simp add: capprox_hcmod_approx_zero)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   542
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   543
lemma Infinitesimal_hcmod_add_diff:
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   544
     "u @c= 0 ==> hcmod(x + u) - hcmod x \<in> Infinitesimal"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   545
apply (rule_tac e = "hcmod u" and e' = "- hcmod u" in Infinitesimal_interval2)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   546
apply (auto dest: capprox_approx_zero_iff [THEN iffD1]
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17300
diff changeset
   547
             simp add: mem_infmal_iff [symmetric] diff_def)
14408
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   548
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
   549
apply (auto simp add: diff_minus [symmetric])
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   550
done
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   551
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   552
lemma approx_hcmod_add_hcmod: "u @c= 0 ==> hcmod(x + u) @= hcmod x"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   553
apply (rule approx_minus_iff [THEN iffD2])
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   554
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
   555
done
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   556
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   557
lemma capprox_hcmod_approx: "x @c= y ==> hcmod x @= hcmod y"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   558
by (auto intro: approx_hcmod_add_hcmod 
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   559
         dest!: bex_CInfinitesimal_iff2 [THEN iffD2]
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   560
         simp add: mem_cinfmal_iff)
13957
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   561
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   562
14408
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   563
subsection{*Zero is the Only Infinitesimal Complex Number*}
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   564
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   565
lemma CInfinitesimal_less_SComplex:
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   566
   "[| x \<in> SComplex; y \<in> CInfinitesimal; 0 < hcmod x |] ==> hcmod y < hcmod x"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   567
by (auto intro!: Infinitesimal_less_SReal SComplex_hcmod_SReal simp add: CInfinitesimal_hcmod_iff)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   568
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   569
lemma SComplex_Int_CInfinitesimal_zero: "SComplex Int CInfinitesimal = {0}"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   570
apply (auto simp add: SComplex_def CInfinitesimal_hcmod_iff)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   571
apply (cut_tac r = r in SReal_hcmod_hcomplex_of_complex)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   572
apply (drule_tac A = Reals in IntI, assumption)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   573
apply (subgoal_tac "hcmod (hcomplex_of_complex r) = 0")
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   574
apply simp
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   575
apply (simp add: SReal_Int_Infinitesimal_zero) 
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   576
done
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   577
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   578
lemma SComplex_CInfinitesimal_zero:
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   579
     "[| x \<in> SComplex; x \<in> CInfinitesimal|] ==> x = 0"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   580
by (cut_tac SComplex_Int_CInfinitesimal_zero, blast)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   581
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   582
lemma SComplex_CFinite_diff_CInfinitesimal:
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   583
     "[| x \<in> SComplex; x \<noteq> 0 |] ==> x \<in> CFinite - CInfinitesimal"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   584
by (auto dest: SComplex_CInfinitesimal_zero SComplex_subset_CFinite [THEN subsetD])
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   585
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   586
lemma hcomplex_of_complex_CFinite_diff_CInfinitesimal:
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   587
     "hcomplex_of_complex x \<noteq> 0 
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   588
      ==> hcomplex_of_complex x \<in> CFinite - CInfinitesimal"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   589
by (rule SComplex_CFinite_diff_CInfinitesimal, auto)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   590
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   591
lemma hcomplex_of_complex_CInfinitesimal_iff_0 [iff]:
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   592
     "(hcomplex_of_complex x \<in> CInfinitesimal) = (x=0)"
17373
27509e72f29e removed duplicated lemmas; convert more proofs to transfer principle
huffman
parents: 17318
diff changeset
   593
apply (auto)
14408
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   594
apply (rule ccontr)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   595
apply (rule hcomplex_of_complex_CFinite_diff_CInfinitesimal [THEN DiffD2], auto)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   596
done
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   597
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   598
lemma number_of_not_CInfinitesimal [simp]:
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   599
     "number_of w \<noteq> (0::hcomplex) ==> number_of w \<notin> CInfinitesimal"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   600
by (fast dest: SComplex_number_of [THEN SComplex_CInfinitesimal_zero])
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   601
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   602
lemma capprox_SComplex_not_zero:
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   603
     "[| y \<in> SComplex; x @c= y; y\<noteq> 0 |] ==> x \<noteq> 0"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   604
by (auto dest: SComplex_CInfinitesimal_zero capprox_sym [THEN mem_cinfmal_iff [THEN iffD2]])
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   605
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   606
lemma CFinite_diff_CInfinitesimal_capprox:
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   607
     "[| x @c= y; y \<in> CFinite - CInfinitesimal |]  
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   608
      ==> x \<in> CFinite - CInfinitesimal"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   609
apply (auto intro: capprox_sym [THEN [2] capprox_CFinite] 
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   610
            simp add: mem_cinfmal_iff)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   611
apply (drule capprox_trans3, assumption)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   612
apply (blast dest: capprox_sym)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   613
done
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   614
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   615
lemma CInfinitesimal_ratio:
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   616
     "[| y \<noteq> 0;  y \<in> CInfinitesimal;  x/y \<in> CFinite |] ==> x \<in> CInfinitesimal"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   617
apply (drule CInfinitesimal_CFinite_mult2, assumption)
15013
34264f5e4691 new treatment of binary numerals
paulson
parents: 14653
diff changeset
   618
apply (simp add: divide_inverse mult_assoc)
14408
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   619
done
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   620
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   621
lemma SComplex_capprox_iff:
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   622
     "[|x \<in> SComplex; y \<in> SComplex|] ==> (x @c= y) = (x = y)"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   623
apply auto
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   624
apply (simp add: capprox_def)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   625
apply (subgoal_tac "x-y = 0", simp) 
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   626
apply (rule SComplex_CInfinitesimal_zero)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   627
apply (simp add: SComplex_diff, assumption)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   628
done
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   629
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   630
lemma number_of_capprox_iff [simp]:
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   631
    "(number_of v @c= number_of w) = (number_of v = (number_of w :: hcomplex))"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   632
by (rule SComplex_capprox_iff, auto)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   633
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   634
lemma number_of_CInfinitesimal_iff [simp]:
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   635
     "(number_of w \<in> CInfinitesimal) = (number_of w = (0::hcomplex))"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   636
apply (rule iffI)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   637
apply (fast dest: SComplex_number_of [THEN SComplex_CInfinitesimal_zero])
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   638
apply (simp (no_asm_simp))
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   639
done
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   640
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   641
lemma hcomplex_of_complex_approx_iff [simp]:
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   642
     "(hcomplex_of_complex k @c= hcomplex_of_complex m) = (k = m)"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   643
apply auto
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   644
apply (rule inj_hcomplex_of_complex [THEN injD])
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   645
apply (rule SComplex_capprox_iff [THEN iffD1], auto)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   646
done
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   647
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   648
lemma hcomplex_of_complex_capprox_number_of_iff [simp]:
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   649
     "(hcomplex_of_complex k @c= number_of w) = (k = number_of w)"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   650
by (subst hcomplex_of_complex_approx_iff [symmetric], auto)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   651
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   652
lemma capprox_unique_complex:
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   653
     "[| r \<in> SComplex; s \<in> SComplex; r @c= x; s @c= x|] ==> r = s"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   654
by (blast intro: SComplex_capprox_iff [THEN iffD1] capprox_trans2)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   655
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   656
lemma hcomplex_capproxD1:
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17300
diff changeset
   657
     "star_n X @c= star_n Y
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17300
diff changeset
   658
      ==> star_n (%n. Re(X n)) @= star_n (%n. Re(Y n))"
20552
2c31dd358c21 generalized types of many constants to work over arbitrary vector spaces;
huffman
parents: 20432
diff changeset
   659
apply (simp add: approx_FreeUltrafilterNat_iff2, safe)
14408
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   660
apply (drule capprox_minus_iff [THEN iffD1])
20552
2c31dd358c21 generalized types of many constants to work over arbitrary vector spaces;
huffman
parents: 20432
diff changeset
   661
apply (simp add: star_n_minus star_n_add mem_cinfmal_iff [symmetric] CInfinitesimal_hcmod_iff hcmod Infinitesimal_FreeUltrafilterNat_iff2)
2c31dd358c21 generalized types of many constants to work over arbitrary vector spaces;
huffman
parents: 20432
diff changeset
   662
apply (drule_tac x = m in spec)
2c31dd358c21 generalized types of many constants to work over arbitrary vector spaces;
huffman
parents: 20432
diff changeset
   663
apply (erule ultra, rule FreeUltrafilterNat_all, clarify)
2c31dd358c21 generalized types of many constants to work over arbitrary vector spaces;
huffman
parents: 20432
diff changeset
   664
apply (rule_tac y="cmod (X n + - Y n)" in order_le_less_trans)
2c31dd358c21 generalized types of many constants to work over arbitrary vector spaces;
huffman
parents: 20432
diff changeset
   665
apply (case_tac "X n")
2c31dd358c21 generalized types of many constants to work over arbitrary vector spaces;
huffman
parents: 20432
diff changeset
   666
apply (case_tac "Y n")
14408
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   667
apply (auto simp add: complex_minus complex_add complex_mod
20552
2c31dd358c21 generalized types of many constants to work over arbitrary vector spaces;
huffman
parents: 20432
diff changeset
   668
            simp del: realpow_Suc)
14408
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   669
done
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   670
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   671
(* same proof *)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   672
lemma hcomplex_capproxD2:
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17300
diff changeset
   673
     "star_n X @c= star_n Y
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17300
diff changeset
   674
      ==> star_n (%n. Im(X n)) @= star_n (%n. Im(Y n))"
20552
2c31dd358c21 generalized types of many constants to work over arbitrary vector spaces;
huffman
parents: 20432
diff changeset
   675
apply (simp add: approx_FreeUltrafilterNat_iff2, safe)
14408
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   676
apply (drule capprox_minus_iff [THEN iffD1])
20552
2c31dd358c21 generalized types of many constants to work over arbitrary vector spaces;
huffman
parents: 20432
diff changeset
   677
apply (simp add: star_n_minus star_n_add mem_cinfmal_iff [symmetric] CInfinitesimal_hcmod_iff hcmod Infinitesimal_FreeUltrafilterNat_iff2)
2c31dd358c21 generalized types of many constants to work over arbitrary vector spaces;
huffman
parents: 20432
diff changeset
   678
apply (drule_tac x = m in spec)
2c31dd358c21 generalized types of many constants to work over arbitrary vector spaces;
huffman
parents: 20432
diff changeset
   679
apply (erule ultra, rule FreeUltrafilterNat_all, clarify)
2c31dd358c21 generalized types of many constants to work over arbitrary vector spaces;
huffman
parents: 20432
diff changeset
   680
apply (rule_tac y="cmod (X n + - Y n)" in order_le_less_trans)
2c31dd358c21 generalized types of many constants to work over arbitrary vector spaces;
huffman
parents: 20432
diff changeset
   681
apply (case_tac "X n")
2c31dd358c21 generalized types of many constants to work over arbitrary vector spaces;
huffman
parents: 20432
diff changeset
   682
apply (case_tac "Y n")
2c31dd358c21 generalized types of many constants to work over arbitrary vector spaces;
huffman
parents: 20432
diff changeset
   683
apply (auto simp add: complex_minus complex_add complex_mod
2c31dd358c21 generalized types of many constants to work over arbitrary vector spaces;
huffman
parents: 20432
diff changeset
   684
            simp del: realpow_Suc)
14408
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   685
done
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   686
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   687
lemma hcomplex_capproxI:
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17300
diff changeset
   688
     "[| star_n (%n. Re(X n)) @= star_n (%n. Re(Y n));  
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17300
diff changeset
   689
         star_n (%n. Im(X n)) @= star_n (%n. Im(Y n))  
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17300
diff changeset
   690
      |] ==> star_n X @c= star_n Y"
14408
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   691
apply (drule approx_minus_iff [THEN iffD1])
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   692
apply (drule approx_minus_iff [THEN iffD1])
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   693
apply (rule capprox_minus_iff [THEN iffD2])
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17300
diff changeset
   694
apply (auto simp add: mem_cinfmal_iff [symmetric] mem_infmal_iff [symmetric] star_n_add star_n_minus CInfinitesimal_hcmod_iff hcmod Infinitesimal_FreeUltrafilterNat_iff)
14408
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   695
apply (drule_tac x = "u/2" in spec)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   696
apply (drule_tac x = "u/2" in spec, auto, ultra)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   697
apply (case_tac "X x")
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   698
apply (case_tac "Y x")
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   699
apply (auto simp add: complex_minus complex_add complex_mod snd_conv fst_conv numeral_2_eq_2)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   700
apply (rename_tac a b c d)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   701
apply (subgoal_tac "sqrt (abs (a + - c) ^ 2 + abs (b + - d) ^ 2) < u")
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   702
apply (rule_tac [2] lemma_sqrt_hcomplex_capprox, auto)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   703
apply (simp add: power2_eq_square)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   704
done
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   705
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   706
lemma capprox_approx_iff:
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17300
diff changeset
   707
     "(star_n X @c= star_n Y) = 
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17300
diff changeset
   708
       (star_n (%n. Re(X n)) @= star_n (%n. Re(Y n)) &  
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17300
diff changeset
   709
        star_n (%n. Im(X n)) @= star_n (%n. Im(Y n)))"
14408
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   710
apply (blast intro: hcomplex_capproxI hcomplex_capproxD1 hcomplex_capproxD2)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   711
done
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   712
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   713
lemma hcomplex_of_hypreal_capprox_iff [simp]:
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   714
     "(hcomplex_of_hypreal x @c= hcomplex_of_hypreal z) = (x @= z)"
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17300
diff changeset
   715
apply (cases x, cases z)
14408
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   716
apply (simp add: hcomplex_of_hypreal capprox_approx_iff)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   717
done
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   718
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   719
lemma CFinite_HFinite_Re:
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17300
diff changeset
   720
     "star_n X \<in> CFinite  
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17300
diff changeset
   721
      ==> star_n (%n. Re(X n)) \<in> HFinite"
14408
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   722
apply (auto simp add: CFinite_hcmod_iff hcmod HFinite_FreeUltrafilterNat_iff)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   723
apply (rule_tac x = u in exI, ultra)
20552
2c31dd358c21 generalized types of many constants to work over arbitrary vector spaces;
huffman
parents: 20432
diff changeset
   724
apply (case_tac "X x")
14408
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   725
apply (auto simp add: complex_mod numeral_2_eq_2 simp del: realpow_Suc)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   726
apply (rule ccontr, drule linorder_not_less [THEN iffD1])
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   727
apply (drule order_less_le_trans, assumption)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   728
apply (drule real_sqrt_ge_abs1 [THEN [2] order_less_le_trans]) 
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   729
apply (auto simp add: numeral_2_eq_2 [symmetric]) 
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   730
done
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   731
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   732
lemma CFinite_HFinite_Im:
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17300
diff changeset
   733
     "star_n X \<in> CFinite  
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17300
diff changeset
   734
      ==> star_n (%n. Im(X n)) \<in> HFinite"
14408
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   735
apply (auto simp add: CFinite_hcmod_iff hcmod HFinite_FreeUltrafilterNat_iff)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   736
apply (rule_tac x = u in exI, ultra)
20552
2c31dd358c21 generalized types of many constants to work over arbitrary vector spaces;
huffman
parents: 20432
diff changeset
   737
apply (case_tac "X x")
14408
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   738
apply (auto simp add: complex_mod simp del: realpow_Suc)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   739
apply (rule ccontr, drule linorder_not_less [THEN iffD1])
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   740
apply (drule order_less_le_trans, assumption)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   741
apply (drule real_sqrt_ge_abs2 [THEN [2] order_less_le_trans], auto) 
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   742
done
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   743
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   744
lemma HFinite_Re_Im_CFinite:
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17300
diff changeset
   745
     "[| star_n (%n. Re(X n)) \<in> HFinite;  
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17300
diff changeset
   746
         star_n (%n. Im(X n)) \<in> HFinite  
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17300
diff changeset
   747
      |] ==> star_n X \<in> CFinite"
14408
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   748
apply (auto simp add: CFinite_hcmod_iff hcmod HFinite_FreeUltrafilterNat_iff)
20552
2c31dd358c21 generalized types of many constants to work over arbitrary vector spaces;
huffman
parents: 20432
diff changeset
   749
apply (rename_tac u v)
14408
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   750
apply (rule_tac x = "2* (u + v) " in exI)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   751
apply ultra
20552
2c31dd358c21 generalized types of many constants to work over arbitrary vector spaces;
huffman
parents: 20432
diff changeset
   752
apply (case_tac "X x")
14408
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   753
apply (auto simp add: complex_mod numeral_2_eq_2 simp del: realpow_Suc)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   754
apply (subgoal_tac "0 < u")
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   755
 prefer 2 apply arith
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   756
apply (subgoal_tac "0 < v")
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   757
 prefer 2 apply arith
20552
2c31dd358c21 generalized types of many constants to work over arbitrary vector spaces;
huffman
parents: 20432
diff changeset
   758
apply (subgoal_tac "sqrt (abs (Re (X x)) ^ 2 + abs (Im (X x)) ^ 2) < 2*u + 2*v")
20432
07ec57376051 lin_arith_prover: splitting reverted because of performance loss
webertj
parents: 20256
diff changeset
   759
apply (rule_tac [2] lemma_sqrt_hcomplex_capprox, auto)
14408
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   760
apply (simp add: power2_eq_square)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   761
done
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   762
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   763
lemma CFinite_HFinite_iff:
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17300
diff changeset
   764
     "(star_n X \<in> CFinite) =  
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17300
diff changeset
   765
      (star_n (%n. Re(X n)) \<in> HFinite &  
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17300
diff changeset
   766
       star_n (%n. Im(X n)) \<in> HFinite)"
14408
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   767
by (blast intro: HFinite_Re_Im_CFinite CFinite_HFinite_Im CFinite_HFinite_Re)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   768
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   769
lemma SComplex_Re_SReal:
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17300
diff changeset
   770
     "star_n X \<in> SComplex  
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17300
diff changeset
   771
      ==> star_n (%n. Re(X n)) \<in> Reals"
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17300
diff changeset
   772
apply (auto simp add: SComplex_def SReal_def star_of_def star_n_eq_iff)
14408
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   773
apply (rule_tac x = "Re r" in exI, ultra)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   774
done
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   775
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   776
lemma SComplex_Im_SReal:
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17300
diff changeset
   777
     "star_n X \<in> SComplex  
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17300
diff changeset
   778
      ==> star_n (%n. Im(X n)) \<in> Reals"
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17300
diff changeset
   779
apply (auto simp add: SComplex_def SReal_def star_of_def star_n_eq_iff)
14408
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   780
apply (rule_tac x = "Im r" in exI, ultra)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   781
done
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   782
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   783
lemma Reals_Re_Im_SComplex:
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17300
diff changeset
   784
     "[| star_n (%n. Re(X n)) \<in> Reals;  
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17300
diff changeset
   785
         star_n (%n. Im(X n)) \<in> Reals  
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17300
diff changeset
   786
      |] ==> star_n X \<in> SComplex"
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17300
diff changeset
   787
apply (auto simp add: SComplex_def SReal_def star_of_def star_n_eq_iff)
14408
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   788
apply (rule_tac x = "Complex r ra" in exI, ultra)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   789
done
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   790
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   791
lemma SComplex_SReal_iff:
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17300
diff changeset
   792
     "(star_n X \<in> SComplex) =  
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17300
diff changeset
   793
      (star_n (%n. Re(X n)) \<in> Reals &  
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17300
diff changeset
   794
       star_n (%n. Im(X n)) \<in> Reals)"
14408
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   795
by (blast intro: SComplex_Re_SReal SComplex_Im_SReal Reals_Re_Im_SComplex)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   796
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   797
lemma CInfinitesimal_Infinitesimal_iff:
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17300
diff changeset
   798
     "(star_n X \<in> CInfinitesimal) =  
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17300
diff changeset
   799
      (star_n (%n. Re(X n)) \<in> Infinitesimal &  
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17300
diff changeset
   800
       star_n (%n. Im(X n)) \<in> Infinitesimal)"
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17300
diff changeset
   801
by (simp add: mem_cinfmal_iff mem_infmal_iff star_n_zero_num capprox_approx_iff)
14408
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   802
17300
5798fbf42a6a replace type hcomplex with complex star
huffman
parents: 17298
diff changeset
   803
lemma eq_Abs_star_EX:
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17300
diff changeset
   804
     "(\<exists>t. P t) = (\<exists>X. P (star_n X))"
17429
e8d6ed3aacfe merged Transfer.thy and StarType.thy into StarDef.thy; renamed Ifun2_of to starfun2; cleaned up
huffman
parents: 17373
diff changeset
   805
by (rule ex_star_eq)
14408
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   806
17300
5798fbf42a6a replace type hcomplex with complex star
huffman
parents: 17298
diff changeset
   807
lemma eq_Abs_star_Bex:
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17300
diff changeset
   808
     "(\<exists>t \<in> A. P t) = (\<exists>X. star_n X \<in> A & P (star_n X))"
17429
e8d6ed3aacfe merged Transfer.thy and StarType.thy into StarDef.thy; renamed Ifun2_of to starfun2; cleaned up
huffman
parents: 17373
diff changeset
   809
by (simp add: Bex_def ex_star_eq)
14408
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   810
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   811
(* Here we go - easy proof now!! *)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   812
lemma stc_part_Ex: "x:CFinite ==> \<exists>t \<in> SComplex. x @c= t"
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17300
diff changeset
   813
apply (cases x)
17300
5798fbf42a6a replace type hcomplex with complex star
huffman
parents: 17298
diff changeset
   814
apply (auto simp add: CFinite_HFinite_iff eq_Abs_star_Bex SComplex_SReal_iff capprox_approx_iff)
14408
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   815
apply (drule st_part_Ex, safe)+
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17300
diff changeset
   816
apply (rule_tac x = t in star_cases)
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17300
diff changeset
   817
apply (rule_tac x = ta in star_cases, auto)
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17300
diff changeset
   818
apply (rule_tac x = "%n. Complex (Xa n) (Xb n) " in exI)
14408
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   819
apply auto
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   820
done
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   821
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   822
lemma stc_part_Ex1: "x:CFinite ==> EX! t. t \<in> SComplex &  x @c= t"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   823
apply (drule stc_part_Ex, safe)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   824
apply (drule_tac [2] capprox_sym, drule_tac [2] capprox_sym, drule_tac [2] capprox_sym)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   825
apply (auto intro!: capprox_unique_complex)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   826
done
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   827
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   828
lemma CFinite_Int_CInfinite_empty: "CFinite Int CInfinite = {}"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   829
by (simp add: CFinite_def CInfinite_def, auto)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   830
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   831
lemma CFinite_not_CInfinite: "x \<in> CFinite ==> x \<notin> CInfinite"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   832
by (insert CFinite_Int_CInfinite_empty, blast)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   833
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   834
text{*Not sure this is a good idea!*}
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   835
declare CFinite_Int_CInfinite_empty [simp]
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   836
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   837
lemma not_CFinite_CInfinite: "x\<notin> CFinite ==> x \<in> CInfinite"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   838
by (auto intro: not_HFinite_HInfinite simp add: CFinite_hcmod_iff CInfinite_hcmod_iff)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   839
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   840
lemma CInfinite_CFinite_disj: "x \<in> CInfinite | x \<in> CFinite"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   841
by (blast intro: not_CFinite_CInfinite)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   842
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   843
lemma CInfinite_CFinite_iff: "(x \<in> CInfinite) = (x \<notin> CFinite)"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   844
by (blast dest: CFinite_not_CInfinite not_CFinite_CInfinite)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   845
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   846
lemma CFinite_CInfinite_iff: "(x \<in> CFinite) = (x \<notin> CInfinite)"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   847
by (simp add: CInfinite_CFinite_iff)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   848
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   849
lemma CInfinite_diff_CFinite_CInfinitesimal_disj:
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   850
     "x \<notin> CInfinitesimal ==> x \<in> CInfinite | x \<in> CFinite - CInfinitesimal"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   851
by (fast intro: not_CFinite_CInfinite)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   852
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   853
lemma CFinite_inverse:
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   854
     "[| x \<in> CFinite; x \<notin> CInfinitesimal |] ==> inverse x \<in> CFinite"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   855
apply (cut_tac x = "inverse x" in CInfinite_CFinite_disj)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   856
apply (auto dest!: CInfinite_inverse_CInfinitesimal)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   857
done
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   858
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   859
lemma CFinite_inverse2: "x \<in> CFinite - CInfinitesimal ==> inverse x \<in> CFinite"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   860
by (blast intro: CFinite_inverse)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   861
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   862
lemma CInfinitesimal_inverse_CFinite:
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   863
     "x \<notin> CInfinitesimal ==> inverse(x) \<in> CFinite"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   864
apply (drule CInfinite_diff_CFinite_CInfinitesimal_disj)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   865
apply (blast intro: CFinite_inverse CInfinite_inverse_CInfinitesimal CInfinitesimal_subset_CFinite [THEN subsetD])
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   866
done
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   867
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   868
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   869
lemma CFinite_not_CInfinitesimal_inverse:
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   870
     "x \<in> CFinite - CInfinitesimal ==> inverse x \<in> CFinite - CInfinitesimal"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   871
apply (auto intro: CInfinitesimal_inverse_CFinite)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   872
apply (drule CInfinitesimal_CFinite_mult2, assumption)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   873
apply (simp add: not_CInfinitesimal_not_zero)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   874
done
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   875
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   876
lemma capprox_inverse:
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   877
     "[| x @c= y; y \<in>  CFinite - CInfinitesimal |] ==> inverse x @c= inverse y"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   878
apply (frule CFinite_diff_CInfinitesimal_capprox, assumption)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   879
apply (frule not_CInfinitesimal_not_zero2)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   880
apply (frule_tac x = x in not_CInfinitesimal_not_zero2)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   881
apply (drule CFinite_inverse2)+
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   882
apply (drule capprox_mult2, assumption, auto)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   883
apply (drule_tac c = "inverse x" in capprox_mult1, assumption)
15013
34264f5e4691 new treatment of binary numerals
paulson
parents: 14653
diff changeset
   884
apply (auto intro: capprox_sym simp add: mult_assoc)
14408
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   885
done
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   886
15013
34264f5e4691 new treatment of binary numerals
paulson
parents: 14653
diff changeset
   887
lemmas hcomplex_of_complex_capprox_inverse =
34264f5e4691 new treatment of binary numerals
paulson
parents: 14653
diff changeset
   888
  hcomplex_of_complex_CFinite_diff_CInfinitesimal [THEN [2] capprox_inverse]
14408
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   889
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   890
lemma inverse_add_CInfinitesimal_capprox:
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   891
     "[| x \<in> CFinite - CInfinitesimal;  
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   892
         h \<in> CInfinitesimal |] ==> inverse(x + h) @c= inverse x"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   893
by (auto intro: capprox_inverse capprox_sym CInfinitesimal_add_capprox_self)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   894
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   895
lemma inverse_add_CInfinitesimal_capprox2:
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   896
     "[| x \<in> CFinite - CInfinitesimal;  
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   897
         h \<in> CInfinitesimal |] ==> inverse(h + x) @c= inverse x"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   898
apply (rule add_commute [THEN subst])
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   899
apply (blast intro: inverse_add_CInfinitesimal_capprox)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   900
done
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   901
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   902
lemma inverse_add_CInfinitesimal_approx_CInfinitesimal:
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   903
     "[| x \<in> CFinite - CInfinitesimal;  
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   904
         h \<in> CInfinitesimal |] ==> inverse(x + h) - inverse x @c= h"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   905
apply (rule capprox_trans2)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   906
apply (auto intro: inverse_add_CInfinitesimal_capprox 
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   907
       simp add: mem_cinfmal_iff diff_minus capprox_minus_iff [symmetric])
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   908
done
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   909
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   910
lemma CInfinitesimal_square_iff [iff]:
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   911
     "(x*x \<in> CInfinitesimal) = (x \<in> CInfinitesimal)"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   912
by (simp add: CInfinitesimal_hcmod_iff hcmod_mult)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   913
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   914
lemma capprox_CFinite_mult_cancel:
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   915
     "[| a \<in> CFinite-CInfinitesimal; a*w @c= a*z |] ==> w @c= z"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   916
apply safe
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   917
apply (frule CFinite_inverse, assumption)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   918
apply (drule not_CInfinitesimal_not_zero)
15013
34264f5e4691 new treatment of binary numerals
paulson
parents: 14653
diff changeset
   919
apply (auto dest: capprox_mult2 simp add: mult_assoc [symmetric])
14408
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   920
done
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   921
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   922
lemma capprox_CFinite_mult_cancel_iff1:
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   923
     "a \<in> CFinite-CInfinitesimal ==> (a * w @c= a * z) = (w @c= z)"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   924
by (auto intro: capprox_mult2 capprox_CFinite_mult_cancel)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   925
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   926
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   927
subsection{*Theorems About Monads*}
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   928
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   929
lemma capprox_cmonad_iff: "(x @c= y) = (cmonad(x)=cmonad(y))"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   930
apply (simp add: cmonad_def)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   931
apply (auto dest: capprox_sym elim!: capprox_trans equalityCE)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   932
done
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   933
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   934
lemma CInfinitesimal_cmonad_eq:
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   935
     "e \<in> CInfinitesimal ==> cmonad (x+e) = cmonad x"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   936
by (fast intro!: CInfinitesimal_add_capprox_self [THEN capprox_sym] capprox_cmonad_iff [THEN iffD1])
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   937
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   938
lemma mem_cmonad_iff: "(u \<in> cmonad x) = (-u \<in> cmonad (-x))"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   939
by (simp add: cmonad_def)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   940
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   941
lemma CInfinitesimal_cmonad_zero_iff: "(x:CInfinitesimal) = (x \<in> cmonad 0)"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   942
by (auto intro: capprox_sym simp add: mem_cinfmal_iff cmonad_def)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   943
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   944
lemma cmonad_zero_minus_iff: "(x \<in> cmonad 0) = (-x \<in> cmonad 0)"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   945
by (simp add: CInfinitesimal_cmonad_zero_iff [symmetric])
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   946
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   947
lemma cmonad_zero_hcmod_iff: "(x \<in> cmonad 0) = (hcmod x:monad 0)"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   948
by (simp add: CInfinitesimal_cmonad_zero_iff [symmetric] CInfinitesimal_hcmod_iff Infinitesimal_monad_zero_iff [symmetric])
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   949
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   950
lemma mem_cmonad_self [simp]: "x \<in> cmonad x"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   951
by (simp add: cmonad_def)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   952
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   953
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   954
subsection{*Theorems About Standard Part*}
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   955
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   956
lemma stc_capprox_self: "x \<in> CFinite ==> stc x @c= x"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   957
apply (simp add: stc_def)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   958
apply (frule stc_part_Ex, safe)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   959
apply (rule someI2)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   960
apply (auto intro: capprox_sym)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   961
done
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   962
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   963
lemma stc_SComplex: "x \<in> CFinite ==> stc x \<in> SComplex"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   964
apply (simp add: stc_def)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   965
apply (frule stc_part_Ex, safe)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   966
apply (rule someI2)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   967
apply (auto intro: capprox_sym)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   968
done
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   969
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   970
lemma stc_CFinite: "x \<in> CFinite ==> stc x \<in> CFinite"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   971
by (erule stc_SComplex [THEN SComplex_subset_CFinite [THEN subsetD]])
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   972
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   973
lemma stc_SComplex_eq [simp]: "x \<in> SComplex ==> stc x = x"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   974
apply (simp add: stc_def)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   975
apply (rule some_equality)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   976
apply (auto intro: SComplex_subset_CFinite [THEN subsetD])
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   977
apply (blast dest: SComplex_capprox_iff [THEN iffD1])
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   978
done
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   979
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   980
lemma stc_hcomplex_of_complex:
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   981
     "stc (hcomplex_of_complex x) = hcomplex_of_complex x"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   982
by auto
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   983
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   984
lemma stc_eq_capprox:
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   985
     "[| x \<in> CFinite; y \<in> CFinite; stc x = stc y |] ==> x @c= y"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   986
by (auto dest!: stc_capprox_self elim!: capprox_trans3)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   987
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   988
lemma capprox_stc_eq:
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   989
     "[| x \<in> CFinite; y \<in> CFinite; x @c= y |] ==> stc x = stc y"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   990
by (blast intro: capprox_trans capprox_trans2 SComplex_capprox_iff [THEN iffD1]
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   991
          dest: stc_capprox_self stc_SComplex)
13957
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   992
14408
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   993
lemma stc_eq_capprox_iff:
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   994
     "[| x \<in> CFinite; y \<in> CFinite|] ==> (x @c= y) = (stc x = stc y)"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   995
by (blast intro: capprox_stc_eq stc_eq_capprox)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   996
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   997
lemma stc_CInfinitesimal_add_SComplex:
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   998
     "[| x \<in> SComplex; e \<in> CInfinitesimal |] ==> stc(x + e) = x"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
   999
apply (frule stc_SComplex_eq [THEN subst])
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1000
prefer 2 apply assumption
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1001
apply (frule SComplex_subset_CFinite [THEN subsetD])
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1002
apply (frule CInfinitesimal_subset_CFinite [THEN subsetD])
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1003
apply (drule stc_SComplex_eq)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1004
apply (rule capprox_stc_eq)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1005
apply (auto intro: CFinite_add simp add: CInfinitesimal_add_capprox_self [THEN capprox_sym])
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1006
done
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1007
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1008
lemma stc_CInfinitesimal_add_SComplex2:
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1009
     "[| x \<in> SComplex; e \<in> CInfinitesimal |] ==> stc(e + x) = x"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1010
apply (rule add_commute [THEN subst])
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1011
apply (blast intro!: stc_CInfinitesimal_add_SComplex)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1012
done
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1013
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1014
lemma CFinite_stc_CInfinitesimal_add:
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1015
     "x \<in> CFinite ==> \<exists>e \<in> CInfinitesimal. x = stc(x) + e"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1016
by (blast dest!: stc_capprox_self [THEN capprox_sym] bex_CInfinitesimal_iff2 [THEN iffD2])
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1017
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1018
lemma stc_add:
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1019
     "[| x \<in> CFinite; y \<in> CFinite |] ==> stc (x + y) = stc(x) + stc(y)"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1020
apply (frule CFinite_stc_CInfinitesimal_add)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1021
apply (frule_tac x = y in CFinite_stc_CInfinitesimal_add, safe)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1022
apply (subgoal_tac "stc (x + y) = stc ((stc x + e) + (stc y + ea))")
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1023
apply (drule_tac [2] sym, drule_tac [2] sym)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1024
 prefer 2 apply simp 
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1025
apply (simp (no_asm_simp) add: add_ac)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1026
apply (drule stc_SComplex)+
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1027
apply (drule SComplex_add, assumption)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1028
apply (drule CInfinitesimal_add, assumption)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1029
apply (rule add_assoc [THEN subst])
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1030
apply (blast intro!: stc_CInfinitesimal_add_SComplex2)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1031
done
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1032
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1033
lemma stc_number_of [simp]: "stc (number_of w) = number_of w"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1034
by (rule SComplex_number_of [THEN stc_SComplex_eq])
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1035
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1036
lemma stc_zero [simp]: "stc 0 = 0"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1037
by simp
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1038
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1039
lemma stc_one [simp]: "stc 1 = 1"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1040
by simp
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1041
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1042
lemma stc_minus: "y \<in> CFinite ==> stc(-y) = -stc(y)"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1043
apply (frule CFinite_minus_iff [THEN iffD2])
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1044
apply (rule hcomplex_add_minus_eq_minus)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1045
apply (drule stc_add [symmetric], assumption)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1046
apply (simp add: add_commute)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1047
done
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1048
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1049
lemma stc_diff: 
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1050
     "[| x \<in> CFinite; y \<in> CFinite |] ==> stc (x-y) = stc(x) - stc(y)"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1051
apply (simp add: diff_minus)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1052
apply (frule_tac y1 = y in stc_minus [symmetric])
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1053
apply (drule_tac x1 = y in CFinite_minus_iff [THEN iffD2])
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1054
apply (auto intro: stc_add)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1055
done
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1056
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1057
lemma lemma_stc_mult:
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1058
     "[| x \<in> CFinite; y \<in> CFinite;  
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1059
         e \<in> CInfinitesimal;        
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1060
         ea: CInfinitesimal |]    
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1061
       ==> e*y + x*ea + e*ea: CInfinitesimal"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1062
apply (frule_tac x = e and y = y in CInfinitesimal_CFinite_mult)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1063
apply (frule_tac [2] x = ea and y = x in CInfinitesimal_CFinite_mult)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1064
apply (drule_tac [3] CInfinitesimal_mult)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1065
apply (auto intro: CInfinitesimal_add simp add: add_ac mult_ac)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1066
done
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1067
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1068
lemma stc_mult:
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1069
     "[| x \<in> CFinite; y \<in> CFinite |]  
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1070
               ==> stc (x * y) = stc(x) * stc(y)"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1071
apply (frule CFinite_stc_CInfinitesimal_add)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1072
apply (frule_tac x = y in CFinite_stc_CInfinitesimal_add, safe)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1073
apply (subgoal_tac "stc (x * y) = stc ((stc x + e) * (stc y + ea))")
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1074
apply (drule_tac [2] sym, drule_tac [2] sym)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1075
 prefer 2 apply simp 
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1076
apply (erule_tac V = "x = stc x + e" in thin_rl)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1077
apply (erule_tac V = "y = stc y + ea" in thin_rl)
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17300
diff changeset
  1078
apply (simp add: left_distrib right_distrib)
14408
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1079
apply (drule stc_SComplex)+
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1080
apply (simp (no_asm_use) add: add_assoc)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1081
apply (rule stc_CInfinitesimal_add_SComplex)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1082
apply (blast intro!: SComplex_mult)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1083
apply (drule SComplex_subset_CFinite [THEN subsetD])+
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1084
apply (rule add_assoc [THEN subst])
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1085
apply (blast intro!: lemma_stc_mult)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1086
done
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1087
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1088
lemma stc_CInfinitesimal: "x \<in> CInfinitesimal ==> stc x = 0"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1089
apply (rule stc_zero [THEN subst])
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1090
apply (rule capprox_stc_eq)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1091
apply (auto intro: CInfinitesimal_subset_CFinite [THEN subsetD]
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1092
                 simp add: mem_cinfmal_iff [symmetric])
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1093
done
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1094
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1095
lemma stc_not_CInfinitesimal: "stc(x) \<noteq> 0 ==> x \<notin> CInfinitesimal"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1096
by (fast intro: stc_CInfinitesimal)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1097
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1098
lemma stc_inverse:
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1099
     "[| x \<in> CFinite; stc x \<noteq> 0 |]  
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1100
      ==> stc(inverse x) = inverse (stc x)"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1101
apply (rule_tac c1 = "stc x" in hcomplex_mult_left_cancel [THEN iffD1])
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1102
apply (auto simp add: stc_mult [symmetric] stc_not_CInfinitesimal CFinite_inverse)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1103
apply (subst right_inverse, auto)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1104
done
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1105
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1106
lemma stc_divide [simp]:
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1107
     "[| x \<in> CFinite; y \<in> CFinite; stc y \<noteq> 0 |]  
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1108
      ==> stc(x/y) = (stc x) / (stc y)"
14430
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14408
diff changeset
  1109
by (simp add: divide_inverse stc_mult stc_not_CInfinitesimal CFinite_inverse stc_inverse)
14408
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1110
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1111
lemma stc_idempotent [simp]: "x \<in> CFinite ==> stc(stc(x)) = stc(x)"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1112
by (blast intro: stc_CFinite stc_capprox_self capprox_stc_eq)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1113
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1114
lemma CFinite_HFinite_hcomplex_of_hypreal:
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1115
     "z \<in> HFinite ==> hcomplex_of_hypreal z \<in> CFinite"
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17300
diff changeset
  1116
apply (cases z)
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17300
diff changeset
  1117
apply (simp add: hcomplex_of_hypreal CFinite_HFinite_iff star_n_zero_num [symmetric])
14408
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1118
done
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1119
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1120
lemma SComplex_SReal_hcomplex_of_hypreal:
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1121
     "x \<in> Reals ==>  hcomplex_of_hypreal x \<in> SComplex"
20557
81dd3679f92c complex_of_real abbreviates of_real::real=>complex;
huffman
parents: 20552
diff changeset
  1122
by (auto simp add: SReal_def SComplex_def hcomplex_of_hypreal_def
81dd3679f92c complex_of_real abbreviates of_real::real=>complex;
huffman
parents: 20552
diff changeset
  1123
         simp del: star_of_of_real)
14408
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1124
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1125
lemma stc_hcomplex_of_hypreal: 
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1126
 "z \<in> HFinite ==> stc(hcomplex_of_hypreal z) = hcomplex_of_hypreal (st z)"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1127
apply (simp add: st_def stc_def)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1128
apply (frule st_part_Ex, safe)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1129
apply (rule someI2)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1130
apply (auto intro: approx_sym)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1131
apply (drule CFinite_HFinite_hcomplex_of_hypreal)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1132
apply (frule stc_part_Ex, safe)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1133
apply (rule someI2)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1134
apply (auto intro: capprox_sym intro!: capprox_unique_complex dest: SComplex_SReal_hcomplex_of_hypreal)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1135
done
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1136
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1137
(*
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1138
Goal "x \<in> CFinite ==> hcmod(stc x) = st(hcmod x)"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1139
by (dtac stc_capprox_self 1)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1140
by (auto_tac (claset(),simpset() addsimps [bex_CInfinitesimal_iff2 RS sym]));
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1141
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1142
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1143
approx_hcmod_add_hcmod
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1144
*)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1145
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1146
lemma CInfinitesimal_hcnj_iff [simp]:
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1147
     "(hcnj z \<in> CInfinitesimal) = (z \<in> CInfinitesimal)"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1148
by (simp add: CInfinitesimal_hcmod_iff)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1149
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1150
lemma CInfinite_HInfinite_iff:
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17300
diff changeset
  1151
     "(star_n X \<in> CInfinite) =  
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17300
diff changeset
  1152
      (star_n (%n. Re(X n)) \<in> HInfinite |  
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17300
diff changeset
  1153
       star_n (%n. Im(X n)) \<in> HInfinite)"
14408
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1154
by (simp add: CInfinite_CFinite_iff HInfinite_HFinite_iff CFinite_HFinite_iff)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1155
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1156
text{*These theorems should probably be deleted*}
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1157
lemma hcomplex_split_CInfinitesimal_iff:
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1158
     "(hcomplex_of_hypreal x + iii * hcomplex_of_hypreal y \<in> CInfinitesimal) =  
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1159
      (x \<in> Infinitesimal & y \<in> Infinitesimal)"
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17300
diff changeset
  1160
apply (cases x, cases y)
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17300
diff changeset
  1161
apply (simp add: iii_def star_of_def star_n_add star_n_mult hcomplex_of_hypreal CInfinitesimal_Infinitesimal_iff)
14408
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1162
done
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1163
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1164
lemma hcomplex_split_CFinite_iff:
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1165
     "(hcomplex_of_hypreal x + iii * hcomplex_of_hypreal y \<in> CFinite) =  
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1166
      (x \<in> HFinite & y \<in> HFinite)"
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17300
diff changeset
  1167
apply (cases x, cases y)
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17300
diff changeset
  1168
apply (simp add: iii_def star_of_def star_n_add star_n_mult hcomplex_of_hypreal CFinite_HFinite_iff)
14408
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1169
done
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1170
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1171
lemma hcomplex_split_SComplex_iff:
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1172
     "(hcomplex_of_hypreal x + iii * hcomplex_of_hypreal y \<in> SComplex) =  
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1173
      (x \<in> Reals & y \<in> Reals)"
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17300
diff changeset
  1174
apply (cases x, cases y)
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17300
diff changeset
  1175
apply (simp add: iii_def star_of_def star_n_add star_n_mult hcomplex_of_hypreal SComplex_SReal_iff)
14408
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1176
done
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1177
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1178
lemma hcomplex_split_CInfinite_iff:
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1179
     "(hcomplex_of_hypreal x + iii * hcomplex_of_hypreal y \<in> CInfinite) =  
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1180
      (x \<in> HInfinite | y \<in> HInfinite)"
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17300
diff changeset
  1181
apply (cases x, cases y)
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17300
diff changeset
  1182
apply (simp add: iii_def star_of_def star_n_add star_n_mult hcomplex_of_hypreal CInfinite_HInfinite_iff)
14408
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1183
done
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1184
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1185
lemma hcomplex_split_capprox_iff:
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1186
     "(hcomplex_of_hypreal x + iii * hcomplex_of_hypreal y @c=  
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1187
       hcomplex_of_hypreal x' + iii * hcomplex_of_hypreal y') =  
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1188
      (x @= x' & y @= y')"
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17300
diff changeset
  1189
apply (cases x, cases y, cases x', cases y')
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17300
diff changeset
  1190
apply (simp add: iii_def star_of_def star_n_add star_n_mult hcomplex_of_hypreal capprox_approx_iff)
14408
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1191
done
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1192
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1193
lemma complex_seq_to_hcomplex_CInfinitesimal:
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1194
     "\<forall>n. cmod (X n - x) < inverse (real (Suc n)) ==>  
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17300
diff changeset
  1195
      star_n X - hcomplex_of_complex x \<in> CInfinitesimal"
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17300
diff changeset
  1196
apply (simp add: star_n_diff CInfinitesimal_hcmod_iff star_of_def Infinitesimal_FreeUltrafilterNat_iff hcmod)
14408
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1197
apply (auto dest: FreeUltrafilterNat_inverse_real_of_posnat FreeUltrafilterNat_all FreeUltrafilterNat_Int intro: order_less_trans FreeUltrafilterNat_subset)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1198
done
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1199
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1200
lemma CInfinitesimal_hcomplex_of_hypreal_epsilon [simp]:
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1201
     "hcomplex_of_hypreal epsilon \<in> CInfinitesimal"
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1202
by (simp add: CInfinitesimal_hcmod_iff)
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1203
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1204
lemma hcomplex_of_complex_approx_zero_iff [simp]:
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1205
     "(hcomplex_of_complex z @c= 0) = (z = 0)"
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17300
diff changeset
  1206
by (simp add: star_of_zero [symmetric] del: star_of_zero)
14408
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1207
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1208
lemma hcomplex_of_complex_approx_zero_iff2 [simp]:
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1209
     "(0 @c= hcomplex_of_complex z) = (z = 0)"
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17300
diff changeset
  1210
by (simp add: star_of_zero [symmetric] del: star_of_zero)
14408
0cc42bb96330 converted HOL/Complex/NSCA to Isar script
paulson
parents: 14387
diff changeset
  1211
13957
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1212
end