src/HOL/Hyperreal/NSA.thy
author huffman
Thu, 14 Sep 2006 21:42:21 +0200
changeset 20541 f614c619b1e1
parent 20485 3078fd2eec7b
child 20552 2c31dd358c21
permissions -rw-r--r--
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
10751
a81ea5d3dd41 separation of HOL-Hyperreal from HOL-Real
paulson
parents:
diff changeset
     1
(*  Title       : NSA.thy
a81ea5d3dd41 separation of HOL-Hyperreal from HOL-Real
paulson
parents:
diff changeset
     2
    Author      : Jacques D. Fleuriot
a81ea5d3dd41 separation of HOL-Hyperreal from HOL-Real
paulson
parents:
diff changeset
     3
    Copyright   : 1998  University of Cambridge
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
     4
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
     5
Converted to Isar and polished by lcp
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
     6
*)
10751
a81ea5d3dd41 separation of HOL-Hyperreal from HOL-Real
paulson
parents:
diff changeset
     7
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
     8
header{*Infinite Numbers, Infinitesimals, Infinitely Close Relation*}
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
     9
15131
c69542757a4d New theory header syntax.
nipkow
parents: 15003
diff changeset
    10
theory NSA
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17298
diff changeset
    11
imports HyperArith "../Real/RComplete"
15131
c69542757a4d New theory header syntax.
nipkow
parents: 15003
diff changeset
    12
begin
10751
a81ea5d3dd41 separation of HOL-Hyperreal from HOL-Real
paulson
parents:
diff changeset
    13
19765
dfe940911617 misc cleanup;
wenzelm
parents: 17431
diff changeset
    14
definition
10751
a81ea5d3dd41 separation of HOL-Hyperreal from HOL-Real
paulson
parents:
diff changeset
    15
20541
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
    16
  hnorm :: "'a::norm star \<Rightarrow> real star"
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
    17
  "hnorm = *f* norm"
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
    18
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
    19
  Infinitesimal  :: "('a::real_normed_vector) star set"
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
    20
  "Infinitesimal = {x. \<forall>r \<in> Reals. 0 < r --> hnorm x < r}"
10751
a81ea5d3dd41 separation of HOL-Hyperreal from HOL-Real
paulson
parents:
diff changeset
    21
20541
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
    22
  HFinite :: "('a::real_normed_vector) star set"
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
    23
  "HFinite = {x. \<exists>r \<in> Reals. hnorm x < r}"
10751
a81ea5d3dd41 separation of HOL-Hyperreal from HOL-Real
paulson
parents:
diff changeset
    24
20541
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
    25
  HInfinite :: "('a::real_normed_vector) star set"
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
    26
  "HInfinite = {x. \<forall>r \<in> Reals. r < hnorm x}"
10751
a81ea5d3dd41 separation of HOL-Hyperreal from HOL-Real
paulson
parents:
diff changeset
    27
14653
0848ab6fe5fc constdefs: proper order;
wenzelm
parents: 14565
diff changeset
    28
  approx :: "[hypreal, hypreal] => bool"    (infixl "@=" 50)
15229
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15140
diff changeset
    29
    --{*the `infinitely close' relation*}
19765
dfe940911617 misc cleanup;
wenzelm
parents: 17431
diff changeset
    30
  "(x @= y) = ((x + -y) \<in> Infinitesimal)"
14653
0848ab6fe5fc constdefs: proper order;
wenzelm
parents: 14565
diff changeset
    31
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
    32
  st        :: "hypreal => hypreal"
15229
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15140
diff changeset
    33
    --{*the standard part of a hyperreal*}
19765
dfe940911617 misc cleanup;
wenzelm
parents: 17431
diff changeset
    34
  "st = (%x. @r. x \<in> HFinite & r \<in> Reals & r @= x)"
10751
a81ea5d3dd41 separation of HOL-Hyperreal from HOL-Real
paulson
parents:
diff changeset
    35
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
    36
  monad     :: "hypreal => hypreal set"
19765
dfe940911617 misc cleanup;
wenzelm
parents: 17431
diff changeset
    37
  "monad x = {y. x @= y}"
10751
a81ea5d3dd41 separation of HOL-Hyperreal from HOL-Real
paulson
parents:
diff changeset
    38
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
    39
  galaxy    :: "hypreal => hypreal set"
19765
dfe940911617 misc cleanup;
wenzelm
parents: 17431
diff changeset
    40
  "galaxy x = {y. (x + -y) \<in> HFinite}"
dfe940911617 misc cleanup;
wenzelm
parents: 17431
diff changeset
    41
dfe940911617 misc cleanup;
wenzelm
parents: 17431
diff changeset
    42
const_syntax (xsymbols)
dfe940911617 misc cleanup;
wenzelm
parents: 17431
diff changeset
    43
  approx  (infixl "\<approx>" 50)
dfe940911617 misc cleanup;
wenzelm
parents: 17431
diff changeset
    44
dfe940911617 misc cleanup;
wenzelm
parents: 17431
diff changeset
    45
const_syntax (HTML output)
dfe940911617 misc cleanup;
wenzelm
parents: 17431
diff changeset
    46
  approx  (infixl "\<approx>" 50)
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
    47
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
    48
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14371
diff changeset
    49
defs (overloaded)
19765
dfe940911617 misc cleanup;
wenzelm
parents: 17431
diff changeset
    50
  SReal_def:      "Reals == {x. \<exists>r. x = hypreal_of_real r}"
dfe940911617 misc cleanup;
wenzelm
parents: 17431
diff changeset
    51
    --{*the standard real numbers as a subset of the hyperreals*}
14565
c6dc17aab88a use more symbols in HTML output
kleing
parents: 14477
diff changeset
    52
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
    53
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
    54
20541
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
    55
subsection{*Nonstandard extension of the norm function*}
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
    56
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
    57
declare hnorm_def [transfer_unfold]
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
    58
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
    59
lemma hnorm_ge_zero [simp]:
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
    60
  "\<And>x::'a::real_normed_vector star. 0 \<le> hnorm x"
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
    61
by transfer (rule norm_ge_zero)
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
    62
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
    63
lemma hnorm_eq_zero [simp]:
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
    64
  "\<And>x::'a::real_normed_vector star. (hnorm x = 0) = (x = 0)"
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
    65
by transfer (rule norm_eq_zero)
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
    66
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
    67
lemma hnorm_triangle_ineq:
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
    68
  "\<And>x y::'a::real_normed_vector star. hnorm (x + y) \<le> hnorm x + hnorm y"
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
    69
by transfer (rule norm_triangle_ineq)
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
    70
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
    71
lemma hnorm_scaleR:
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
    72
  "\<And>a (x::'a::real_normed_vector star).
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
    73
   hnorm (( *f2* scaleR) a x) = \<bar>a\<bar> * hnorm x"
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
    74
by transfer (rule norm_scaleR)
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
    75
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
    76
lemma hnorm_mult_ineq:
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
    77
  "\<And>x y::'a::real_normed_algebra star. hnorm (x * y) \<le> hnorm x * hnorm y"
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
    78
by transfer (rule norm_mult_ineq)
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
    79
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
    80
lemma hnorm_mult:
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
    81
  "\<And>x y::'a::real_normed_div_algebra star. hnorm (x * y) = hnorm x * hnorm y"
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
    82
by transfer (rule norm_mult)
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
    83
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
    84
lemma hnorm_one [simp]:
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
    85
  "hnorm (1\<Colon>'a::real_normed_div_algebra star) = 1"
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
    86
by transfer (rule norm_one)
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
    87
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
    88
lemma hnorm_zero [simp]:
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
    89
  "hnorm (0\<Colon>'a::real_normed_vector star) = 0"
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
    90
by transfer (rule norm_zero)
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
    91
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
    92
lemma zero_less_hnorm_iff [simp]:
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
    93
  "\<And>x::'a::real_normed_vector star. (0 < hnorm x) = (x \<noteq> 0)"
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
    94
by transfer (rule zero_less_norm_iff)
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
    95
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
    96
lemma hnorm_minus_cancel [simp]:
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
    97
  "\<And>x::'a::real_normed_vector star. hnorm (- x) = hnorm x"
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
    98
by transfer (rule norm_minus_cancel)
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
    99
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   100
lemma hnorm_minus_commute:
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   101
  "\<And>a b::'a::real_normed_vector star. hnorm (a - b) = hnorm (b - a)"
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   102
by transfer (rule norm_minus_commute)
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   103
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   104
lemma hnorm_triangle_ineq2:
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   105
  "\<And>a b::'a::real_normed_vector star. hnorm a - hnorm b \<le> hnorm (a - b)"
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   106
by transfer (rule norm_triangle_ineq2)
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   107
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   108
lemma hnorm_triangle_ineq4:
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   109
  "\<And>a b::'a::real_normed_vector star. hnorm (a - b) \<le> hnorm a + hnorm b"
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   110
by transfer (rule norm_triangle_ineq4)
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   111
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   112
lemma nonzero_hnorm_inverse:
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   113
  "\<And>a::'a::real_normed_div_algebra star.
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   114
   a \<noteq> 0 \<Longrightarrow> hnorm (inverse a) = inverse (hnorm a)"
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   115
by transfer (rule nonzero_norm_inverse)
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   116
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   117
lemma hnorm_inverse:
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   118
  "\<And>a::'a::{real_normed_div_algebra,division_by_zero} star.
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   119
   hnorm (inverse a) = inverse (hnorm a)"
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   120
by transfer (rule norm_inverse)
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   121
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   122
lemma hypreal_hnorm_def [simp]:
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   123
  "\<And>r::hypreal. hnorm r \<equiv> \<bar>r\<bar>"
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   124
by transfer (rule real_norm_def)
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   125
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   126
lemma hnorm_add_less:
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   127
  fixes x y :: "'a::real_normed_vector star"
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   128
  shows "\<lbrakk>hnorm x < r; hnorm y < s\<rbrakk> \<Longrightarrow> hnorm (x + y) < r + s"
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   129
by (rule order_le_less_trans [OF hnorm_triangle_ineq add_strict_mono])
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   130
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   131
lemma hnorm_mult_less:
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   132
  fixes x y :: "'a::real_normed_algebra star"
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   133
  shows "\<lbrakk>hnorm x < r; hnorm y < s\<rbrakk> \<Longrightarrow> hnorm (x * y) < r * s"
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   134
apply (rule order_le_less_trans [OF hnorm_mult_ineq])
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   135
apply (simp add: mult_strict_mono')
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   136
done
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   137
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   138
15229
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15140
diff changeset
   139
subsection{*Closure Laws for the Standard Reals*}
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   140
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14371
diff changeset
   141
lemma SReal_add [simp]:
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14371
diff changeset
   142
     "[| (x::hypreal) \<in> Reals; y \<in> Reals |] ==> x + y \<in> Reals"
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   143
apply (auto simp add: SReal_def)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   144
apply (rule_tac x = "r + ra" in exI, simp)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   145
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   146
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   147
lemma SReal_mult: "[| (x::hypreal) \<in> Reals; y \<in> Reals |] ==> x * y \<in> Reals"
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   148
apply (simp add: SReal_def, safe)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   149
apply (rule_tac x = "r * ra" in exI)
15539
333a88244569 comprehensive cleanup, replacing sumr by setsum
nipkow
parents: 15531
diff changeset
   150
apply (simp (no_asm))
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   151
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   152
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   153
lemma SReal_inverse: "(x::hypreal) \<in> Reals ==> inverse x \<in> Reals"
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   154
apply (simp add: SReal_def)
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17298
diff changeset
   155
apply (blast intro: star_of_inverse [symmetric])
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   156
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   157
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   158
lemma SReal_divide: "[| (x::hypreal) \<in> Reals;  y \<in> Reals |] ==> x/y \<in> Reals"
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17298
diff changeset
   159
by (simp (no_asm_simp) add: SReal_mult SReal_inverse divide_inverse)
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   160
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   161
lemma SReal_minus: "(x::hypreal) \<in> Reals ==> -x \<in> Reals"
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   162
apply (simp add: SReal_def)
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17298
diff changeset
   163
apply (blast intro: star_of_minus [symmetric])
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   164
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   165
15229
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15140
diff changeset
   166
lemma SReal_minus_iff [simp]: "(-x \<in> Reals) = ((x::hypreal) \<in> Reals)"
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   167
apply auto
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   168
apply (erule_tac [2] SReal_minus)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   169
apply (drule SReal_minus, auto)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   170
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   171
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
   172
lemma SReal_add_cancel:
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
   173
     "[| (x::hypreal) + y \<in> Reals; y \<in> Reals |] ==> x \<in> Reals"
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   174
apply (drule_tac x = y in SReal_minus)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   175
apply (drule SReal_add, assumption, auto)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   176
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   177
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   178
lemma SReal_hrabs: "(x::hypreal) \<in> Reals ==> abs x \<in> Reals"
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17298
diff changeset
   179
apply (auto simp add: SReal_def)
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17298
diff changeset
   180
apply (rule_tac x="abs r" in exI)
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17298
diff changeset
   181
apply simp
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   182
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   183
15229
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15140
diff changeset
   184
lemma SReal_hypreal_of_real [simp]: "hypreal_of_real x \<in> Reals"
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   185
by (simp add: SReal_def)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   186
15229
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15140
diff changeset
   187
lemma SReal_number_of [simp]: "(number_of w ::hypreal) \<in> Reals"
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17298
diff changeset
   188
apply (simp only: star_of_number_of [symmetric])
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   189
apply (rule SReal_hypreal_of_real)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   190
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   191
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   192
(** As always with numerals, 0 and 1 are special cases **)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   193
15229
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15140
diff changeset
   194
lemma Reals_0 [simp]: "(0::hypreal) \<in> Reals"
14387
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   195
apply (subst numeral_0_eq_0 [symmetric])
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   196
apply (rule SReal_number_of)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   197
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   198
15229
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15140
diff changeset
   199
lemma Reals_1 [simp]: "(1::hypreal) \<in> Reals"
14387
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   200
apply (subst numeral_1_eq_1 [symmetric])
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   201
apply (rule SReal_number_of)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   202
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   203
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   204
lemma SReal_divide_number_of: "r \<in> Reals ==> r/(number_of w::hypreal) \<in> Reals"
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17298
diff changeset
   205
apply (simp only: divide_inverse)
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   206
apply (blast intro!: SReal_number_of SReal_mult SReal_inverse)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   207
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   208
15229
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15140
diff changeset
   209
text{*epsilon is not in Reals because it is an infinitesimal*}
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   210
lemma SReal_epsilon_not_mem: "epsilon \<notin> Reals"
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   211
apply (simp add: SReal_def)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   212
apply (auto simp add: hypreal_of_real_not_eq_epsilon [THEN not_sym])
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   213
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   214
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   215
lemma SReal_omega_not_mem: "omega \<notin> Reals"
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   216
apply (simp add: SReal_def)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   217
apply (auto simp add: hypreal_of_real_not_eq_omega [THEN not_sym])
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   218
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   219
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   220
lemma SReal_UNIV_real: "{x. hypreal_of_real x \<in> Reals} = (UNIV::real set)"
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   221
by (simp add: SReal_def)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   222
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   223
lemma SReal_iff: "(x \<in> Reals) = (\<exists>y. x = hypreal_of_real y)"
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   224
by (simp add: SReal_def)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   225
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   226
lemma hypreal_of_real_image: "hypreal_of_real `(UNIV::real set) = Reals"
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   227
by (auto simp add: SReal_def)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   228
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   229
lemma inv_hypreal_of_real_image: "inv hypreal_of_real ` Reals = UNIV"
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   230
apply (auto simp add: SReal_def)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   231
apply (rule inj_hypreal_of_real [THEN inv_f_f, THEN subst], blast)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   232
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   233
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   234
lemma SReal_hypreal_of_real_image:
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
   235
      "[| \<exists>x. x: P; P \<subseteq> Reals |] ==> \<exists>Q. P = hypreal_of_real ` Q"
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   236
apply (simp add: SReal_def, blast)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   237
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   238
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
   239
lemma SReal_dense:
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
   240
     "[| (x::hypreal) \<in> Reals; y \<in> Reals;  x<y |] ==> \<exists>r \<in> Reals. x<r & r<y"
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   241
apply (auto simp add: SReal_iff)
14477
cc61fd03e589 conversion of Hyperreal/Lim to new-style
paulson
parents: 14468
diff changeset
   242
apply (drule dense, safe)
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   243
apply (rule_tac x = "hypreal_of_real r" in bexI, auto)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   244
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   245
15229
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15140
diff changeset
   246
text{*Completeness of Reals, but both lemmas are unused.*}
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15140
diff changeset
   247
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
   248
lemma SReal_sup_lemma:
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
   249
     "P \<subseteq> Reals ==> ((\<exists>x \<in> P. y < x) =
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   250
      (\<exists>X. hypreal_of_real X \<in> P & y < hypreal_of_real X))"
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   251
by (blast dest!: SReal_iff [THEN iffD1])
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   252
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   253
lemma SReal_sup_lemma2:
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
   254
     "[| P \<subseteq> Reals; \<exists>x. x \<in> P; \<exists>y \<in> Reals. \<forall>x \<in> P. x < y |]
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   255
      ==> (\<exists>X. X \<in> {w. hypreal_of_real w \<in> P}) &
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   256
          (\<exists>Y. \<forall>X \<in> {w. hypreal_of_real w \<in> P}. X < Y)"
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   257
apply (rule conjI)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   258
apply (fast dest!: SReal_iff [THEN iffD1])
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   259
apply (auto, frule subsetD, assumption)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   260
apply (drule SReal_iff [THEN iffD1])
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   261
apply (auto, rule_tac x = ya in exI, auto)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   262
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   263
15229
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15140
diff changeset
   264
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15140
diff changeset
   265
subsection{*Lifting of the Ub and Lub Properties*}
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15140
diff changeset
   266
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   267
lemma hypreal_of_real_isUb_iff:
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   268
      "(isUb (Reals) (hypreal_of_real ` Q) (hypreal_of_real Y)) =
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   269
       (isUb (UNIV :: real set) Q Y)"
15229
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15140
diff changeset
   270
by (simp add: isUb_def setle_def)
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   271
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   272
lemma hypreal_of_real_isLub1:
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   273
     "isLub Reals (hypreal_of_real ` Q) (hypreal_of_real Y)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   274
      ==> isLub (UNIV :: real set) Q Y"
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   275
apply (simp add: isLub_def leastP_def)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   276
apply (auto intro: hypreal_of_real_isUb_iff [THEN iffD2]
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   277
            simp add: hypreal_of_real_isUb_iff setge_def)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   278
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   279
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   280
lemma hypreal_of_real_isLub2:
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   281
      "isLub (UNIV :: real set) Q Y
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   282
       ==> isLub Reals (hypreal_of_real ` Q) (hypreal_of_real Y)"
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   283
apply (simp add: isLub_def leastP_def)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   284
apply (auto simp add: hypreal_of_real_isUb_iff setge_def)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   285
apply (frule_tac x2 = x in isUbD2a [THEN SReal_iff [THEN iffD1], THEN exE])
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   286
 prefer 2 apply assumption
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   287
apply (drule_tac x = xa in spec)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   288
apply (auto simp add: hypreal_of_real_isUb_iff)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   289
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   290
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
   291
lemma hypreal_of_real_isLub_iff:
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
   292
     "(isLub Reals (hypreal_of_real ` Q) (hypreal_of_real Y)) =
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   293
      (isLub (UNIV :: real set) Q Y)"
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
   294
by (blast intro: hypreal_of_real_isLub1 hypreal_of_real_isLub2)
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   295
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   296
lemma lemma_isUb_hypreal_of_real:
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   297
     "isUb Reals P Y ==> \<exists>Yo. isUb Reals P (hypreal_of_real Yo)"
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   298
by (auto simp add: SReal_iff isUb_def)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   299
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   300
lemma lemma_isLub_hypreal_of_real:
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   301
     "isLub Reals P Y ==> \<exists>Yo. isLub Reals P (hypreal_of_real Yo)"
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   302
by (auto simp add: isLub_def leastP_def isUb_def SReal_iff)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   303
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   304
lemma lemma_isLub_hypreal_of_real2:
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   305
     "\<exists>Yo. isLub Reals P (hypreal_of_real Yo) ==> \<exists>Y. isLub Reals P Y"
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   306
by (auto simp add: isLub_def leastP_def isUb_def)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   307
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
   308
lemma SReal_complete:
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
   309
     "[| P \<subseteq> Reals;  \<exists>x. x \<in> P;  \<exists>Y. isUb Reals P Y |]
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   310
      ==> \<exists>t::hypreal. isLub Reals P t"
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   311
apply (frule SReal_hypreal_of_real_image)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   312
apply (auto, drule lemma_isUb_hypreal_of_real)
15229
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15140
diff changeset
   313
apply (auto intro!: reals_complete lemma_isLub_hypreal_of_real2
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15140
diff changeset
   314
            simp add: hypreal_of_real_isLub_iff hypreal_of_real_isUb_iff)
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   315
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   316
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
   317
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
   318
subsection{* Set of Finite Elements is a Subring of the Extended Reals*}
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
   319
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   320
lemma HFinite_add: "[|x \<in> HFinite; y \<in> HFinite|] ==> (x+y) \<in> HFinite"
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   321
apply (simp add: HFinite_def)
20541
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   322
apply (blast intro!: SReal_add hnorm_add_less)
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   323
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   324
20541
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   325
lemma HFinite_mult:
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   326
  fixes x y :: "'a::real_normed_algebra star"
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   327
  shows "[|x \<in> HFinite; y \<in> HFinite|] ==> x*y \<in> HFinite"
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   328
apply (simp add: HFinite_def)
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   329
apply (blast intro!: SReal_mult hnorm_mult_less)
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   330
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   331
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   332
lemma HFinite_minus_iff: "(-x \<in> HFinite) = (x \<in> HFinite)"
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   333
by (simp add: HFinite_def)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   334
20541
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   335
lemma HFinite_star_of [simp]: "star_of x \<in> HFinite"
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   336
apply (simp add: HFinite_def)
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   337
apply (rule_tac x="star_of (norm x) + 1" in bexI)
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   338
apply (transfer, simp)
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   339
apply (blast intro: SReal_add SReal_hypreal_of_real Reals_1)
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   340
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   341
20541
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   342
lemma SReal_subset_HFinite: "(Reals::hypreal set) \<subseteq> HFinite"
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   343
by (auto simp add: SReal_def)
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   344
20541
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   345
lemma HFinite_hypreal_of_real: "hypreal_of_real x \<in> HFinite"
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   346
by (rule HFinite_star_of)
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   347
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   348
lemma HFiniteD: "x \<in> HFinite ==> \<exists>t \<in> Reals. hnorm x < t"
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   349
by (simp add: HFinite_def)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   350
20541
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   351
lemma HFinite_hrabs_iff [iff]: "(abs (x::hypreal) \<in> HFinite) = (x \<in> HFinite)"
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   352
by (simp add: HFinite_def)
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   353
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   354
lemma HFinite_hnorm_iff [iff]: "(hnorm x \<in> HFinite) = (x \<in> HFinite)"
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   355
by (simp add: HFinite_def)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   356
15229
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15140
diff changeset
   357
lemma HFinite_number_of [simp]: "number_of w \<in> HFinite"
20541
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   358
by (unfold star_number_def, rule HFinite_star_of)
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   359
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   360
(** As always with numerals, 0 and 1 are special cases **)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   361
15229
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15140
diff changeset
   362
lemma HFinite_0 [simp]: "0 \<in> HFinite"
20541
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   363
by (unfold star_zero_def, rule HFinite_star_of)
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   364
15229
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15140
diff changeset
   365
lemma HFinite_1 [simp]: "1 \<in> HFinite"
20541
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   366
by (unfold star_one_def, rule HFinite_star_of)
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   367
20541
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   368
lemma HFinite_bounded:
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   369
  "[|(x::hypreal) \<in> HFinite; y \<le> x; 0 \<le> y |] ==> y \<in> HFinite"
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
   370
apply (case_tac "x \<le> 0")
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   371
apply (drule_tac y = x in order_trans)
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17298
diff changeset
   372
apply (drule_tac [2] order_antisym)
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   373
apply (auto simp add: linorder_not_le)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   374
apply (auto intro: order_le_less_trans simp add: abs_if HFinite_def)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   375
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   376
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
   377
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
   378
subsection{* Set of Infinitesimals is a Subring of the Hyperreals*}
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
   379
20407
93a34d5d1dc5 speed up some proofs
huffman
parents: 20254
diff changeset
   380
lemma InfinitesimalI:
20541
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   381
  "(\<And>r. \<lbrakk>r \<in> \<real>; 0 < r\<rbrakk> \<Longrightarrow> hnorm x < r) \<Longrightarrow> x \<in> Infinitesimal"
20407
93a34d5d1dc5 speed up some proofs
huffman
parents: 20254
diff changeset
   382
by (simp add: Infinitesimal_def)
93a34d5d1dc5 speed up some proofs
huffman
parents: 20254
diff changeset
   383
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   384
lemma InfinitesimalD:
20541
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   385
      "x \<in> Infinitesimal ==> \<forall>r \<in> Reals. 0 < r --> hnorm x < r"
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
   386
by (simp add: Infinitesimal_def)
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   387
15229
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15140
diff changeset
   388
lemma Infinitesimal_zero [iff]: "0 \<in> Infinitesimal"
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   389
by (simp add: Infinitesimal_def)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   390
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   391
lemma hypreal_sum_of_halves: "x/(2::hypreal) + x/(2::hypreal) = x"
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   392
by auto
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   393
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   394
lemma Infinitesimal_add:
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   395
     "[| x \<in> Infinitesimal; y \<in> Infinitesimal |] ==> (x+y) \<in> Infinitesimal"
20407
93a34d5d1dc5 speed up some proofs
huffman
parents: 20254
diff changeset
   396
apply (rule InfinitesimalI)
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   397
apply (rule hypreal_sum_of_halves [THEN subst])
14477
cc61fd03e589 conversion of Hyperreal/Lim to new-style
paulson
parents: 14468
diff changeset
   398
apply (drule half_gt_zero)
20541
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   399
apply (blast intro: hnorm_add_less SReal_divide_number_of dest: InfinitesimalD)
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   400
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   401
15229
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15140
diff changeset
   402
lemma Infinitesimal_minus_iff [simp]: "(-x:Infinitesimal) = (x:Infinitesimal)"
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   403
by (simp add: Infinitesimal_def)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   404
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
   405
lemma Infinitesimal_diff:
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
   406
     "[| x \<in> Infinitesimal;  y \<in> Infinitesimal |] ==> x-y \<in> Infinitesimal"
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17298
diff changeset
   407
by (simp add: diff_def Infinitesimal_add)
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   408
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   409
lemma Infinitesimal_mult:
20541
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   410
  fixes x y :: "'a::real_normed_algebra star"
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   411
  shows "[|x \<in> Infinitesimal; y \<in> Infinitesimal|] ==> (x * y) \<in> Infinitesimal"
20407
93a34d5d1dc5 speed up some proofs
huffman
parents: 20254
diff changeset
   412
apply (rule InfinitesimalI)
93a34d5d1dc5 speed up some proofs
huffman
parents: 20254
diff changeset
   413
apply (case_tac "y = 0", simp)
20541
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   414
apply (subgoal_tac "hnorm (x * y) < 1 * r", simp only: mult_1)
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   415
apply (rule hnorm_mult_less)
20407
93a34d5d1dc5 speed up some proofs
huffman
parents: 20254
diff changeset
   416
apply (simp_all add: InfinitesimalD)
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   417
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   418
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
   419
lemma Infinitesimal_HFinite_mult:
20541
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   420
  fixes x y :: "'a::real_normed_algebra star"
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   421
  shows "[| x \<in> Infinitesimal; y \<in> HFinite |] ==> (x * y) \<in> Infinitesimal"
20407
93a34d5d1dc5 speed up some proofs
huffman
parents: 20254
diff changeset
   422
apply (rule InfinitesimalI)
93a34d5d1dc5 speed up some proofs
huffman
parents: 20254
diff changeset
   423
apply (drule HFiniteD, clarify)
20541
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   424
apply (subgoal_tac "0 < t")
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   425
apply (subgoal_tac "hnorm (x * y) < (r / t) * t", simp)
20407
93a34d5d1dc5 speed up some proofs
huffman
parents: 20254
diff changeset
   426
apply (subgoal_tac "0 < r / t")
20541
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   427
apply (rule hnorm_mult_less)
20407
93a34d5d1dc5 speed up some proofs
huffman
parents: 20254
diff changeset
   428
apply (simp add: InfinitesimalD SReal_divide)
20541
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   429
apply assumption
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   430
apply (simp only: divide_pos_pos)
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   431
apply (erule order_le_less_trans [OF hnorm_ge_zero])
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   432
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   433
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
   434
lemma Infinitesimal_HFinite_mult2:
20541
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   435
  fixes x y :: "'a::real_normed_algebra star"
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   436
  shows "[| x \<in> Infinitesimal; y \<in> HFinite |] ==> (y * x) \<in> Infinitesimal"
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   437
apply (rule InfinitesimalI)
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   438
apply (drule HFiniteD, clarify)
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   439
apply (subgoal_tac "0 < t")
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   440
apply (subgoal_tac "hnorm (y * x) < t * (r / t)", simp)
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   441
apply (subgoal_tac "0 < r / t")
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   442
apply (rule hnorm_mult_less)
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   443
apply assumption
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   444
apply (simp add: InfinitesimalD SReal_divide)
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   445
apply (simp only: divide_pos_pos)
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   446
apply (erule order_le_less_trans [OF hnorm_ge_zero])
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   447
done
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   448
20541
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   449
lemma Compl_HFinite: "- HFinite = HInfinite"
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   450
apply (auto simp add: HInfinite_def HFinite_def linorder_not_less)
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   451
apply (rule_tac y="r + 1" in order_less_le_trans, simp)
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   452
apply (simp add: SReal_add Reals_1)
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   453
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   454
20541
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   455
lemma HInfinite_inverse_Infinitesimal:
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   456
  fixes x :: "'a::real_normed_div_algebra star"
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   457
  shows "x \<in> HInfinite ==> inverse x \<in> Infinitesimal"
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   458
apply (rule InfinitesimalI)
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   459
apply (subgoal_tac "x \<noteq> 0")
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   460
apply (rule inverse_less_imp_less)
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   461
apply (simp add: nonzero_hnorm_inverse)
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   462
apply (simp add: HInfinite_def SReal_inverse)
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   463
apply assumption
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   464
apply (clarify, simp add: Compl_HFinite [symmetric])
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   465
done
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   466
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   467
lemma HInfiniteI: "(\<And>r. r \<in> \<real> \<Longrightarrow> r < hnorm x) \<Longrightarrow> x \<in> HInfinite"
20407
93a34d5d1dc5 speed up some proofs
huffman
parents: 20254
diff changeset
   468
by (simp add: HInfinite_def)
93a34d5d1dc5 speed up some proofs
huffman
parents: 20254
diff changeset
   469
20541
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   470
lemma HInfiniteD: "\<lbrakk>x \<in> HInfinite; r \<in> \<real>\<rbrakk> \<Longrightarrow> r < hnorm x"
20407
93a34d5d1dc5 speed up some proofs
huffman
parents: 20254
diff changeset
   471
by (simp add: HInfinite_def)
93a34d5d1dc5 speed up some proofs
huffman
parents: 20254
diff changeset
   472
20541
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   473
lemma HInfinite_mult:
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   474
  fixes x y :: "'a::real_normed_div_algebra star"
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   475
  shows "[|x \<in> HInfinite; y \<in> HInfinite|] ==> (x*y) \<in> HInfinite"
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   476
apply (rule HInfiniteI, simp only: hnorm_mult)
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   477
apply (subgoal_tac "r * 1 < hnorm x * hnorm y", simp only: mult_1)
20407
93a34d5d1dc5 speed up some proofs
huffman
parents: 20254
diff changeset
   478
apply (case_tac "x = 0", simp add: HInfinite_def)
93a34d5d1dc5 speed up some proofs
huffman
parents: 20254
diff changeset
   479
apply (rule mult_strict_mono)
93a34d5d1dc5 speed up some proofs
huffman
parents: 20254
diff changeset
   480
apply (simp_all add: HInfiniteD)
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   481
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   482
15229
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15140
diff changeset
   483
lemma hypreal_add_zero_less_le_mono: "[|r < x; (0::hypreal) \<le> y|] ==> r < x+y"
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15140
diff changeset
   484
by (auto dest: add_less_le_mono)
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15140
diff changeset
   485
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   486
lemma HInfinite_add_ge_zero:
20541
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   487
     "[|(x::hypreal) \<in> HInfinite; 0 \<le> y; 0 \<le> x|] ==> (x + y): HInfinite"
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   488
by (auto intro!: hypreal_add_zero_less_le_mono 
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17298
diff changeset
   489
       simp add: abs_if add_commute add_nonneg_nonneg HInfinite_def)
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   490
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
   491
lemma HInfinite_add_ge_zero2:
20541
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   492
     "[|(x::hypreal) \<in> HInfinite; 0 \<le> y; 0 \<le> x|] ==> (y + x): HInfinite"
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17298
diff changeset
   493
by (auto intro!: HInfinite_add_ge_zero simp add: add_commute)
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   494
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
   495
lemma HInfinite_add_gt_zero:
20541
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   496
     "[|(x::hypreal) \<in> HInfinite; 0 < y; 0 < x|] ==> (x + y): HInfinite"
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   497
by (blast intro: HInfinite_add_ge_zero order_less_imp_le)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   498
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   499
lemma HInfinite_minus_iff: "(-x \<in> HInfinite) = (x \<in> HInfinite)"
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   500
by (simp add: HInfinite_def)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   501
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
   502
lemma HInfinite_add_le_zero:
20541
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   503
     "[|(x::hypreal) \<in> HInfinite; y \<le> 0; x \<le> 0|] ==> (x + y): HInfinite"
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   504
apply (drule HInfinite_minus_iff [THEN iffD2])
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   505
apply (rule HInfinite_minus_iff [THEN iffD1])
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   506
apply (auto intro: HInfinite_add_ge_zero)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   507
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   508
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
   509
lemma HInfinite_add_lt_zero:
20541
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   510
     "[|(x::hypreal) \<in> HInfinite; y < 0; x < 0|] ==> (x + y): HInfinite"
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   511
by (blast intro: HInfinite_add_le_zero order_less_imp_le)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   512
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
   513
lemma HFinite_sum_squares:
20541
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   514
  fixes a b c :: "'a::real_normed_algebra star"
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   515
  shows "[|a: HFinite; b: HFinite; c: HFinite|]
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   516
      ==> a*a + b*b + c*c \<in> HFinite"
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
   517
by (auto intro: HFinite_mult HFinite_add)
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   518
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   519
lemma not_Infinitesimal_not_zero: "x \<notin> Infinitesimal ==> x \<noteq> 0"
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   520
by auto
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   521
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   522
lemma not_Infinitesimal_not_zero2: "x \<in> HFinite - Infinitesimal ==> x \<noteq> 0"
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   523
by auto
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   524
15229
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15140
diff changeset
   525
lemma Infinitesimal_hrabs_iff [iff]:
20541
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   526
     "(abs (x::hypreal) \<in> Infinitesimal) = (x \<in> Infinitesimal)"
15003
6145dd7538d7 replaced monomorphic abs definitions by abs_if
paulson
parents: 14653
diff changeset
   527
by (auto simp add: abs_if)
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   528
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
   529
lemma HFinite_diff_Infinitesimal_hrabs:
20541
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   530
  "(x::hypreal) \<in> HFinite - Infinitesimal ==> abs x \<in> HFinite - Infinitesimal"
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   531
by blast
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   532
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   533
lemma hrabs_less_Infinitesimal:
20541
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   534
      "[| e \<in> Infinitesimal; abs (x::hypreal) < e |] ==> x \<in> Infinitesimal"
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
   535
by (auto simp add: Infinitesimal_def abs_less_iff)
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   536
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
   537
lemma hrabs_le_Infinitesimal:
20541
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   538
     "[| e \<in> Infinitesimal; abs (x::hypreal) \<le> e |] ==> x \<in> Infinitesimal"
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   539
by (blast dest: order_le_imp_less_or_eq intro: hrabs_less_Infinitesimal)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   540
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   541
lemma Infinitesimal_interval:
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   542
      "[| e \<in> Infinitesimal; e' \<in> Infinitesimal; e' < x ; x < e |] 
20541
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   543
       ==> (x::hypreal) \<in> Infinitesimal"
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
   544
by (auto simp add: Infinitesimal_def abs_less_iff)
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   545
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
   546
lemma Infinitesimal_interval2:
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
   547
     "[| e \<in> Infinitesimal; e' \<in> Infinitesimal;
20541
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   548
         e' \<le> x ; x \<le> e |] ==> (x::hypreal) \<in> Infinitesimal"
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
   549
by (auto intro: Infinitesimal_interval simp add: order_le_less)
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   550
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   551
lemma not_Infinitesimal_mult:
20541
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   552
  fixes x y :: "'a::real_normed_div_algebra star"
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   553
  shows "[| x \<notin> Infinitesimal;  y \<notin> Infinitesimal|] ==> (x*y) \<notin>Infinitesimal"
20407
93a34d5d1dc5 speed up some proofs
huffman
parents: 20254
diff changeset
   554
apply (unfold Infinitesimal_def, clarify, rename_tac r s)
20541
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   555
apply (simp only: linorder_not_less hnorm_mult)
20407
93a34d5d1dc5 speed up some proofs
huffman
parents: 20254
diff changeset
   556
apply (drule_tac x = "r * s" in bspec)
93a34d5d1dc5 speed up some proofs
huffman
parents: 20254
diff changeset
   557
apply (fast intro: SReal_mult)
93a34d5d1dc5 speed up some proofs
huffman
parents: 20254
diff changeset
   558
apply (drule mp, blast intro: mult_pos_pos)
20541
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   559
apply (drule_tac c = s and d = "hnorm y" and a = r and b = "hnorm x" in mult_mono)
20407
93a34d5d1dc5 speed up some proofs
huffman
parents: 20254
diff changeset
   560
apply (simp_all (no_asm_simp))
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   561
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   562
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
   563
lemma Infinitesimal_mult_disj:
20541
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   564
  fixes x y :: "'a::real_normed_div_algebra star"
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   565
  shows "x*y \<in> Infinitesimal ==> x \<in> Infinitesimal | y \<in> Infinitesimal"
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   566
apply (rule ccontr)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   567
apply (drule de_Morgan_disj [THEN iffD1])
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   568
apply (fast dest: not_Infinitesimal_mult)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   569
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   570
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   571
lemma HFinite_Infinitesimal_not_zero: "x \<in> HFinite-Infinitesimal ==> x \<noteq> 0"
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   572
by blast
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   573
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
   574
lemma HFinite_Infinitesimal_diff_mult:
20541
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   575
  fixes x y :: "'a::real_normed_div_algebra star"
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   576
  shows "[| x \<in> HFinite - Infinitesimal;
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   577
                   y \<in> HFinite - Infinitesimal
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   578
                |] ==> (x*y) \<in> HFinite - Infinitesimal"
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   579
apply clarify
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   580
apply (blast dest: HFinite_mult not_Infinitesimal_mult)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   581
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   582
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   583
lemma Infinitesimal_subset_HFinite:
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
   584
      "Infinitesimal \<subseteq> HFinite"
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   585
apply (simp add: Infinitesimal_def HFinite_def, auto)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   586
apply (rule_tac x = 1 in bexI, auto)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   587
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   588
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
   589
lemma Infinitesimal_hypreal_of_real_mult:
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
   590
     "x \<in> Infinitesimal ==> x * hypreal_of_real r \<in> Infinitesimal"
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   591
by (erule HFinite_hypreal_of_real [THEN [2] Infinitesimal_HFinite_mult])
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   592
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
   593
lemma Infinitesimal_hypreal_of_real_mult2:
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
   594
     "x \<in> Infinitesimal ==> hypreal_of_real r * x \<in> Infinitesimal"
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   595
by (erule HFinite_hypreal_of_real [THEN [2] Infinitesimal_HFinite_mult2])
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   596
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
   597
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
   598
subsection{*The Infinitely Close Relation*}
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   599
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   600
lemma mem_infmal_iff: "(x \<in> Infinitesimal) = (x @= 0)"
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   601
by (simp add: Infinitesimal_def approx_def)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   602
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   603
lemma approx_minus_iff: " (x @= y) = (x + -y @= 0)"
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   604
by (simp add: approx_def)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   605
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   606
lemma approx_minus_iff2: " (x @= y) = (-y + x @= 0)"
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17298
diff changeset
   607
by (simp add: approx_def add_commute)
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   608
15229
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15140
diff changeset
   609
lemma approx_refl [iff]: "x @= x"
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   610
by (simp add: approx_def Infinitesimal_def)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   611
14477
cc61fd03e589 conversion of Hyperreal/Lim to new-style
paulson
parents: 14468
diff changeset
   612
lemma hypreal_minus_distrib1: "-(y + -(x::hypreal)) = x + -y"
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17298
diff changeset
   613
by (simp add: add_commute)
14477
cc61fd03e589 conversion of Hyperreal/Lim to new-style
paulson
parents: 14468
diff changeset
   614
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   615
lemma approx_sym: "x @= y ==> y @= x"
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   616
apply (simp add: approx_def)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   617
apply (rule hypreal_minus_distrib1 [THEN subst])
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   618
apply (erule Infinitesimal_minus_iff [THEN iffD2])
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   619
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   620
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   621
lemma approx_trans: "[| x @= y; y @= z |] ==> x @= z"
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   622
apply (simp add: approx_def)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   623
apply (drule Infinitesimal_add, assumption, auto)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   624
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   625
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   626
lemma approx_trans2: "[| r @= x; s @= x |] ==> r @= s"
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   627
by (blast intro: approx_sym approx_trans)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   628
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   629
lemma approx_trans3: "[| x @= r; x @= s|] ==> r @= s"
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   630
by (blast intro: approx_sym approx_trans)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   631
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   632
lemma number_of_approx_reorient: "(number_of w @= x) = (x @= number_of w)"
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   633
by (blast intro: approx_sym)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   634
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   635
lemma zero_approx_reorient: "(0 @= x) = (x @= 0)"
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   636
by (blast intro: approx_sym)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   637
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   638
lemma one_approx_reorient: "(1 @= x) = (x @= 1)"
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   639
by (blast intro: approx_sym)
10751
a81ea5d3dd41 separation of HOL-Hyperreal from HOL-Real
paulson
parents:
diff changeset
   640
a81ea5d3dd41 separation of HOL-Hyperreal from HOL-Real
paulson
parents:
diff changeset
   641
19765
dfe940911617 misc cleanup;
wenzelm
parents: 17431
diff changeset
   642
ML {*
dfe940911617 misc cleanup;
wenzelm
parents: 17431
diff changeset
   643
local
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   644
(*** re-orientation, following HOL/Integ/Bin.ML
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   645
     We re-orient x @=y where x is 0, 1 or a numeral, unless y is as well!
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   646
 ***)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   647
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   648
(*reorientation simprules using ==, for the following simproc*)
19765
dfe940911617 misc cleanup;
wenzelm
parents: 17431
diff changeset
   649
val meta_zero_approx_reorient = thm "zero_approx_reorient" RS eq_reflection;
dfe940911617 misc cleanup;
wenzelm
parents: 17431
diff changeset
   650
val meta_one_approx_reorient = thm "one_approx_reorient" RS eq_reflection;
dfe940911617 misc cleanup;
wenzelm
parents: 17431
diff changeset
   651
val meta_number_of_approx_reorient = thm "number_of_approx_reorient" RS eq_reflection
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   652
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   653
(*reorientation simplification procedure: reorients (polymorphic)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   654
  0 = x, 1 = x, nnn = x provided x isn't 0, 1 or a numeral.*)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   655
fun reorient_proc sg _ (_ $ t $ u) =
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   656
  case u of
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15251
diff changeset
   657
      Const("0", _) => NONE
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15251
diff changeset
   658
    | Const("1", _) => NONE
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15251
diff changeset
   659
    | Const("Numeral.number_of", _) $ _ => NONE
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15251
diff changeset
   660
    | _ => SOME (case t of
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   661
                Const("0", _) => meta_zero_approx_reorient
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   662
              | Const("1", _) => meta_one_approx_reorient
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   663
              | Const("Numeral.number_of", _) $ _ =>
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   664
                                 meta_number_of_approx_reorient);
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   665
19765
dfe940911617 misc cleanup;
wenzelm
parents: 17431
diff changeset
   666
in
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   667
val approx_reorient_simproc =
20485
3078fd2eec7b got rid of Numeral.bin type
haftmann
parents: 20432
diff changeset
   668
  Int_Numeral_Base_Simprocs.prep_simproc
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   669
    ("reorient_simproc", ["0@=x", "1@=x", "number_of w @= x"], reorient_proc);
19765
dfe940911617 misc cleanup;
wenzelm
parents: 17431
diff changeset
   670
end;
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   671
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   672
Addsimprocs [approx_reorient_simproc];
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   673
*}
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   674
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   675
lemma Infinitesimal_approx_minus: "(x-y \<in> Infinitesimal) = (x @= y)"
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17298
diff changeset
   676
by (auto simp add: diff_def approx_minus_iff [symmetric] mem_infmal_iff)
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   677
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   678
lemma approx_monad_iff: "(x @= y) = (monad(x)=monad(y))"
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   679
apply (simp add: monad_def)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   680
apply (auto dest: approx_sym elim!: approx_trans equalityCE)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   681
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   682
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
   683
lemma Infinitesimal_approx:
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
   684
     "[| x \<in> Infinitesimal; y \<in> Infinitesimal |] ==> x @= y"
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   685
apply (simp add: mem_infmal_iff)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   686
apply (blast intro: approx_trans approx_sym)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   687
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   688
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   689
lemma approx_add: "[| a @= b; c @= d |] ==> a+c @= b+d"
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   690
proof (unfold approx_def)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   691
  assume inf: "a + - b \<in> Infinitesimal" "c + - d \<in> Infinitesimal"
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   692
  have "a + c + - (b + d) = (a + - b) + (c + - d)" by arith
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   693
  also have "... \<in> Infinitesimal" using inf by (rule Infinitesimal_add)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   694
  finally show "a + c + - (b + d) \<in> Infinitesimal" .
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   695
qed
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   696
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   697
lemma approx_minus: "a @= b ==> -a @= -b"
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   698
apply (rule approx_minus_iff [THEN iffD2, THEN approx_sym])
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   699
apply (drule approx_minus_iff [THEN iffD1])
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17298
diff changeset
   700
apply (simp (no_asm) add: add_commute)
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   701
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   702
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   703
lemma approx_minus2: "-a @= -b ==> a @= b"
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   704
by (auto dest: approx_minus)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   705
15229
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15140
diff changeset
   706
lemma approx_minus_cancel [simp]: "(-a @= -b) = (a @= b)"
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   707
by (blast intro: approx_minus approx_minus2)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   708
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   709
lemma approx_add_minus: "[| a @= b; c @= d |] ==> a + -c @= b + -d"
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   710
by (blast intro!: approx_add approx_minus)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   711
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   712
lemma approx_mult1: "[| a @= b; c: HFinite|] ==> a*c @= b*c"
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   713
by (simp add: approx_def Infinitesimal_HFinite_mult minus_mult_left 
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   714
              left_distrib [symmetric] 
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   715
         del: minus_mult_left [symmetric])
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   716
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   717
lemma approx_mult2: "[|a @= b; c: HFinite|] ==> c*a @= c*b"
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17298
diff changeset
   718
by (simp add: approx_mult1 mult_commute)
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   719
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   720
lemma approx_mult_subst: "[|u @= v*x; x @= y; v \<in> HFinite|] ==> u @= v*y"
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   721
by (blast intro: approx_mult2 approx_trans)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   722
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   723
lemma approx_mult_subst2: "[| u @= x*v; x @= y; v \<in> HFinite |] ==> u @= y*v"
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   724
by (blast intro: approx_mult1 approx_trans)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   725
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
   726
lemma approx_mult_subst_SReal:
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
   727
     "[| u @= x*hypreal_of_real v; x @= y |] ==> u @= y*hypreal_of_real v"
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   728
by (auto intro: approx_mult_subst2)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   729
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   730
lemma approx_eq_imp: "a = b ==> a @= b"
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   731
by (simp add: approx_def)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   732
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   733
lemma Infinitesimal_minus_approx: "x \<in> Infinitesimal ==> -x @= x"
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   734
by (blast intro: Infinitesimal_minus_iff [THEN iffD2] 
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   735
                    mem_infmal_iff [THEN iffD1] approx_trans2)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   736
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   737
lemma bex_Infinitesimal_iff: "(\<exists>y \<in> Infinitesimal. x + -z = y) = (x @= z)"
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   738
by (simp add: approx_def)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   739
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   740
lemma bex_Infinitesimal_iff2: "(\<exists>y \<in> Infinitesimal. x = z + y) = (x @= z)"
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   741
by (force simp add: bex_Infinitesimal_iff [symmetric])
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   742
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   743
lemma Infinitesimal_add_approx: "[| y \<in> Infinitesimal; x + y = z |] ==> x @= z"
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   744
apply (rule bex_Infinitesimal_iff [THEN iffD1])
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   745
apply (drule Infinitesimal_minus_iff [THEN iffD2])
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17298
diff changeset
   746
apply (auto simp add: add_assoc [symmetric])
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   747
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   748
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   749
lemma Infinitesimal_add_approx_self: "y \<in> Infinitesimal ==> x @= x + y"
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   750
apply (rule bex_Infinitesimal_iff [THEN iffD1])
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   751
apply (drule Infinitesimal_minus_iff [THEN iffD2])
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17298
diff changeset
   752
apply (auto simp add: add_assoc [symmetric])
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   753
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   754
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   755
lemma Infinitesimal_add_approx_self2: "y \<in> Infinitesimal ==> x @= y + x"
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17298
diff changeset
   756
by (auto dest: Infinitesimal_add_approx_self simp add: add_commute)
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   757
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   758
lemma Infinitesimal_add_minus_approx_self: "y \<in> Infinitesimal ==> x @= x + -y"
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   759
by (blast intro!: Infinitesimal_add_approx_self Infinitesimal_minus_iff [THEN iffD2])
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   760
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   761
lemma Infinitesimal_add_cancel: "[| y \<in> Infinitesimal; x+y @= z|] ==> x @= z"
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   762
apply (drule_tac x = x in Infinitesimal_add_approx_self [THEN approx_sym])
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   763
apply (erule approx_trans3 [THEN approx_sym], assumption)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   764
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   765
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
   766
lemma Infinitesimal_add_right_cancel:
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
   767
     "[| y \<in> Infinitesimal; x @= z + y|] ==> x @= z"
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   768
apply (drule_tac x = z in Infinitesimal_add_approx_self2 [THEN approx_sym])
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   769
apply (erule approx_trans3 [THEN approx_sym])
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17298
diff changeset
   770
apply (simp add: add_commute)
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   771
apply (erule approx_sym)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   772
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   773
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   774
lemma approx_add_left_cancel: "d + b  @= d + c ==> b @= c"
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   775
apply (drule approx_minus_iff [THEN iffD1])
15539
333a88244569 comprehensive cleanup, replacing sumr by setsum
nipkow
parents: 15531
diff changeset
   776
apply (simp add: approx_minus_iff [symmetric] add_ac)
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   777
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   778
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   779
lemma approx_add_right_cancel: "b + d @= c + d ==> b @= c"
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   780
apply (rule approx_add_left_cancel)
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17298
diff changeset
   781
apply (simp add: add_commute)
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   782
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   783
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   784
lemma approx_add_mono1: "b @= c ==> d + b @= d + c"
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   785
apply (rule approx_minus_iff [THEN iffD2])
15539
333a88244569 comprehensive cleanup, replacing sumr by setsum
nipkow
parents: 15531
diff changeset
   786
apply (simp add: approx_minus_iff [symmetric] add_ac)
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   787
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   788
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   789
lemma approx_add_mono2: "b @= c ==> b + a @= c + a"
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17298
diff changeset
   790
by (simp add: add_commute approx_add_mono1)
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   791
15229
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15140
diff changeset
   792
lemma approx_add_left_iff [simp]: "(a + b @= a + c) = (b @= c)"
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   793
by (fast elim: approx_add_left_cancel approx_add_mono1)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   794
15229
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15140
diff changeset
   795
lemma approx_add_right_iff [simp]: "(b + a @= c + a) = (b @= c)"
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17298
diff changeset
   796
by (simp add: add_commute)
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   797
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   798
lemma approx_HFinite: "[| x \<in> HFinite; x @= y |] ==> y \<in> HFinite"
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   799
apply (drule bex_Infinitesimal_iff2 [THEN iffD2], safe)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   800
apply (drule Infinitesimal_subset_HFinite [THEN subsetD, THEN HFinite_minus_iff [THEN iffD2]])
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   801
apply (drule HFinite_add)
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17298
diff changeset
   802
apply (auto simp add: add_assoc)
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   803
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   804
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   805
lemma approx_hypreal_of_real_HFinite: "x @= hypreal_of_real D ==> x \<in> HFinite"
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   806
by (rule approx_sym [THEN [2] approx_HFinite], auto)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   807
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
   808
lemma approx_mult_HFinite:
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
   809
     "[|a @= b; c @= d; b: HFinite; d: HFinite|] ==> a*c @= b*d"
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   810
apply (rule approx_trans)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   811
apply (rule_tac [2] approx_mult2)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   812
apply (rule approx_mult1)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   813
prefer 2 apply (blast intro: approx_HFinite approx_sym, auto)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   814
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   815
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
   816
lemma approx_mult_hypreal_of_real:
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
   817
     "[|a @= hypreal_of_real b; c @= hypreal_of_real d |]
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   818
      ==> a*c @= hypreal_of_real b*hypreal_of_real d"
15229
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15140
diff changeset
   819
by (blast intro!: approx_mult_HFinite approx_hypreal_of_real_HFinite 
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15140
diff changeset
   820
                  HFinite_hypreal_of_real)
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   821
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
   822
lemma approx_SReal_mult_cancel_zero:
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
   823
     "[| a \<in> Reals; a \<noteq> 0; a*x @= 0 |] ==> x @= 0"
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   824
apply (drule SReal_inverse [THEN SReal_subset_HFinite [THEN subsetD]])
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17298
diff changeset
   825
apply (auto dest: approx_mult2 simp add: mult_assoc [symmetric])
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   826
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   827
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   828
lemma approx_mult_SReal1: "[| a \<in> Reals; x @= 0 |] ==> x*a @= 0"
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   829
by (auto dest: SReal_subset_HFinite [THEN subsetD] approx_mult1)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   830
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   831
lemma approx_mult_SReal2: "[| a \<in> Reals; x @= 0 |] ==> a*x @= 0"
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   832
by (auto dest: SReal_subset_HFinite [THEN subsetD] approx_mult2)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   833
15229
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15140
diff changeset
   834
lemma approx_mult_SReal_zero_cancel_iff [simp]:
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
   835
     "[|a \<in> Reals; a \<noteq> 0 |] ==> (a*x @= 0) = (x @= 0)"
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   836
by (blast intro: approx_SReal_mult_cancel_zero approx_mult_SReal2)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   837
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
   838
lemma approx_SReal_mult_cancel:
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
   839
     "[| a \<in> Reals; a \<noteq> 0; a* w @= a*z |] ==> w @= z"
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   840
apply (drule SReal_inverse [THEN SReal_subset_HFinite [THEN subsetD]])
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17298
diff changeset
   841
apply (auto dest: approx_mult2 simp add: mult_assoc [symmetric])
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   842
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   843
15229
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15140
diff changeset
   844
lemma approx_SReal_mult_cancel_iff1 [simp]:
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
   845
     "[| a \<in> Reals; a \<noteq> 0|] ==> (a* w @= a*z) = (w @= z)"
15229
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15140
diff changeset
   846
by (auto intro!: approx_mult2 SReal_subset_HFinite [THEN subsetD]
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15140
diff changeset
   847
         intro: approx_SReal_mult_cancel)
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   848
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
   849
lemma approx_le_bound: "[| z \<le> f; f @= g; g \<le> z |] ==> f @= z"
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   850
apply (simp add: bex_Infinitesimal_iff2 [symmetric], auto)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   851
apply (rule_tac x = "g+y-z" in bexI)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   852
apply (simp (no_asm))
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   853
apply (rule Infinitesimal_interval2)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   854
apply (rule_tac [2] Infinitesimal_zero, auto)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   855
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   856
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
   857
15539
333a88244569 comprehensive cleanup, replacing sumr by setsum
nipkow
parents: 15531
diff changeset
   858
subsection{* Zero is the Only Infinitesimal that is also a Real*}
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   859
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   860
lemma Infinitesimal_less_SReal:
20541
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   861
     "[| (x::hypreal) \<in> Reals; y \<in> Infinitesimal; 0 < x |] ==> y < x"
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   862
apply (simp add: Infinitesimal_def)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   863
apply (rule abs_ge_self [THEN order_le_less_trans], auto)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   864
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   865
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
   866
lemma Infinitesimal_less_SReal2:
20541
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   867
     "(y::hypreal) \<in> Infinitesimal ==> \<forall>r \<in> Reals. 0 < r --> y < r"
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   868
by (blast intro: Infinitesimal_less_SReal)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   869
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   870
lemma SReal_not_Infinitesimal:
20541
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   871
     "[| 0 < y;  (y::hypreal) \<in> Reals|] ==> y \<notin> Infinitesimal"
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   872
apply (simp add: Infinitesimal_def)
15003
6145dd7538d7 replaced monomorphic abs definitions by abs_if
paulson
parents: 14653
diff changeset
   873
apply (auto simp add: abs_if)
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   874
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   875
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
   876
lemma SReal_minus_not_Infinitesimal:
20541
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   877
     "[| y < 0;  (y::hypreal) \<in> Reals |] ==> y \<notin> Infinitesimal"
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   878
apply (subst Infinitesimal_minus_iff [symmetric])
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   879
apply (rule SReal_not_Infinitesimal, auto)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   880
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   881
20541
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   882
lemma SReal_Int_Infinitesimal_zero: "Reals Int Infinitesimal = {0::hypreal}"
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   883
apply auto
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   884
apply (cut_tac x = x and y = 0 in linorder_less_linear)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   885
apply (blast dest: SReal_not_Infinitesimal SReal_minus_not_Infinitesimal)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   886
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   887
20541
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   888
lemma SReal_Infinitesimal_zero:
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   889
  "[| (x::hypreal) \<in> Reals; x \<in> Infinitesimal|] ==> x = 0"
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   890
by (cut_tac SReal_Int_Infinitesimal_zero, blast)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   891
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
   892
lemma SReal_HFinite_diff_Infinitesimal:
20541
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   893
     "[| (x::hypreal) \<in> Reals; x \<noteq> 0 |] ==> x \<in> HFinite - Infinitesimal"
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   894
by (auto dest: SReal_Infinitesimal_zero SReal_subset_HFinite [THEN subsetD])
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   895
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
   896
lemma hypreal_of_real_HFinite_diff_Infinitesimal:
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
   897
     "hypreal_of_real x \<noteq> 0 ==> hypreal_of_real x \<in> HFinite - Infinitesimal"
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   898
by (rule SReal_HFinite_diff_Infinitesimal, auto)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   899
20541
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   900
lemma star_of_Infinitesimal_iff_0 [iff]:
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   901
  "(star_of x \<in> Infinitesimal) = (x = 0)"
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   902
apply (auto simp add: Infinitesimal_def)
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   903
apply (drule_tac x="hnorm (star_of x)" in bspec)
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   904
apply (simp add: hnorm_def)
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   905
apply simp
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   906
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   907
20541
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   908
lemma hypreal_of_real_Infinitesimal_iff_0:
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   909
     "(hypreal_of_real x \<in> Infinitesimal) = (x=0)"
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   910
by (rule star_of_Infinitesimal_iff_0)
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   911
15229
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15140
diff changeset
   912
lemma number_of_not_Infinitesimal [simp]:
20541
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   913
     "number_of w \<noteq> (0::hypreal) ==> (number_of w :: hypreal) \<notin> Infinitesimal"
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   914
by (fast dest: SReal_number_of [THEN SReal_Infinitesimal_zero])
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   915
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   916
(*again: 1 is a special case, but not 0 this time*)
20541
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   917
lemma one_not_Infinitesimal [simp]:
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   918
  "(1::'a::{real_normed_vector,axclass_0_neq_1} star) \<notin> Infinitesimal"
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   919
apply (simp only: star_one_def star_of_Infinitesimal_iff_0)
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   920
apply simp
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   921
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   922
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   923
lemma approx_SReal_not_zero: "[| y \<in> Reals; x @= y; y\<noteq> 0 |] ==> x \<noteq> 0"
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   924
apply (cut_tac x = 0 and y = y in linorder_less_linear, simp)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   925
apply (blast dest: approx_sym [THEN mem_infmal_iff [THEN iffD2]] SReal_not_Infinitesimal SReal_minus_not_Infinitesimal)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   926
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   927
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
   928
lemma HFinite_diff_Infinitesimal_approx:
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
   929
     "[| x @= y; y \<in> HFinite - Infinitesimal |]
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   930
      ==> x \<in> HFinite - Infinitesimal"
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   931
apply (auto intro: approx_sym [THEN [2] approx_HFinite]
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   932
            simp add: mem_infmal_iff)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   933
apply (drule approx_trans3, assumption)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   934
apply (blast dest: approx_sym)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   935
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   936
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   937
(*The premise y\<noteq>0 is essential; otherwise x/y =0 and we lose the
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   938
  HFinite premise.*)
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
   939
lemma Infinitesimal_ratio:
20541
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   940
  fixes x y :: "'a::{real_normed_div_algebra,field} star"
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   941
  shows "[| y \<noteq> 0;  y \<in> Infinitesimal;  x/y \<in> HFinite |]
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   942
         ==> x \<in> Infinitesimal"
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   943
apply (drule Infinitesimal_HFinite_mult2, assumption)
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17298
diff changeset
   944
apply (simp add: divide_inverse mult_assoc)
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   945
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   946
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
   947
lemma Infinitesimal_SReal_divide: 
20541
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
   948
  "[| (x::hypreal) \<in> Infinitesimal; y \<in> Reals |] ==> x/y \<in> Infinitesimal"
14430
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14420
diff changeset
   949
apply (simp add: divide_inverse)
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
   950
apply (auto intro!: Infinitesimal_HFinite_mult 
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
   951
            dest!: SReal_inverse [THEN SReal_subset_HFinite [THEN subsetD]])
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
   952
done
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
   953
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   954
(*------------------------------------------------------------------
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   955
       Standard Part Theorem: Every finite x: R* is infinitely
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   956
       close to a unique real number (i.e a member of Reals)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   957
 ------------------------------------------------------------------*)
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
   958
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
   959
subsection{* Uniqueness: Two Infinitely Close Reals are Equal*}
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   960
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   961
lemma SReal_approx_iff: "[|x \<in> Reals; y \<in> Reals|] ==> (x @= y) = (x = y)"
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   962
apply auto
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   963
apply (simp add: approx_def)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   964
apply (drule_tac x = y in SReal_minus)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   965
apply (drule SReal_add, assumption)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   966
apply (drule SReal_Infinitesimal_zero, assumption)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   967
apply (drule sym)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   968
apply (simp add: hypreal_eq_minus_iff [symmetric])
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   969
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   970
15229
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15140
diff changeset
   971
lemma number_of_approx_iff [simp]:
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
   972
     "(number_of v @= number_of w) = (number_of v = (number_of w :: hypreal))"
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   973
by (auto simp add: SReal_approx_iff)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   974
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   975
(*And also for 0 @= #nn and 1 @= #nn, #nn @= 0 and #nn @= 1.*)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   976
lemma [simp]: "(0 @= number_of w) = ((number_of w :: hypreal) = 0)"
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   977
              "(number_of w @= 0) = ((number_of w :: hypreal) = 0)"
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   978
              "(1 @= number_of w) = ((number_of w :: hypreal) = 1)"
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   979
              "(number_of w @= 1) = ((number_of w :: hypreal) = 1)"
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   980
              "~ (0 @= 1)" "~ (1 @= 0)"
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   981
by (auto simp only: SReal_number_of SReal_approx_iff Reals_0 Reals_1)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   982
15229
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15140
diff changeset
   983
lemma hypreal_of_real_approx_iff [simp]:
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
   984
     "(hypreal_of_real k @= hypreal_of_real m) = (k = m)"
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   985
apply auto
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   986
apply (rule inj_hypreal_of_real [THEN injD])
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   987
apply (rule SReal_approx_iff [THEN iffD1], auto)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   988
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   989
15229
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15140
diff changeset
   990
lemma hypreal_of_real_approx_number_of_iff [simp]:
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
   991
     "(hypreal_of_real k @= number_of w) = (k = number_of w)"
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   992
by (subst hypreal_of_real_approx_iff [symmetric], auto)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   993
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   994
(*And also for 0 and 1.*)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   995
(*And also for 0 @= #nn and 1 @= #nn, #nn @= 0 and #nn @= 1.*)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   996
lemma [simp]: "(hypreal_of_real k @= 0) = (k = 0)"
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   997
              "(hypreal_of_real k @= 1) = (k = 1)"
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   998
  by (simp_all add:  hypreal_of_real_approx_iff [symmetric])
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
   999
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  1000
lemma approx_unique_real:
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  1001
     "[| r \<in> Reals; s \<in> Reals; r @= x; s @= x|] ==> r = s"
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1002
by (blast intro: SReal_approx_iff [THEN iffD1] approx_trans2)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1003
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  1004
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  1005
subsection{* Existence of Unique Real Infinitely Close*}
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  1006
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1007
(* lemma about lubs *)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1008
lemma hypreal_isLub_unique:
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1009
     "[| isLub R S x; isLub R S y |] ==> x = (y::hypreal)"
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1010
apply (frule isLub_isUb)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1011
apply (frule_tac x = y in isLub_isUb)
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17298
diff changeset
  1012
apply (blast intro!: order_antisym dest!: isLub_le_isUb)
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1013
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1014
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  1015
lemma lemma_st_part_ub:
20541
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
  1016
     "(x::hypreal) \<in> HFinite ==> \<exists>u. isUb Reals {s. s \<in> Reals & s < x} u"
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1017
apply (drule HFiniteD, safe)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1018
apply (rule exI, rule isUbI)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1019
apply (auto intro: setleI isUbI simp add: abs_less_iff)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1020
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1021
20541
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
  1022
lemma lemma_st_part_nonempty:
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
  1023
  "(x::hypreal) \<in> HFinite ==> \<exists>y. y \<in> {s. s \<in> Reals & s < x}"
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1024
apply (drule HFiniteD, safe)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1025
apply (drule SReal_minus)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1026
apply (rule_tac x = "-t" in exI)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1027
apply (auto simp add: abs_less_iff)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1028
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1029
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  1030
lemma lemma_st_part_subset: "{s. s \<in> Reals & s < x} \<subseteq> Reals"
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1031
by auto
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1032
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  1033
lemma lemma_st_part_lub:
20541
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
  1034
     "(x::hypreal) \<in> HFinite ==> \<exists>t. isLub Reals {s. s \<in> Reals & s < x} t"
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1035
by (blast intro!: SReal_complete lemma_st_part_ub lemma_st_part_nonempty lemma_st_part_subset)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1036
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  1037
lemma lemma_hypreal_le_left_cancel: "((t::hypreal) + r \<le> t) = (r \<le> 0)"
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1038
apply safe
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1039
apply (drule_tac c = "-t" in add_left_mono)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1040
apply (drule_tac [2] c = t in add_left_mono)
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17298
diff changeset
  1041
apply (auto simp add: add_assoc [symmetric])
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1042
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1043
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  1044
lemma lemma_st_part_le1:
20541
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
  1045
     "[| (x::hypreal) \<in> HFinite;  isLub Reals {s. s \<in> Reals & s < x} t;
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  1046
         r \<in> Reals;  0 < r |] ==> x \<le> t + r"
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1047
apply (frule isLubD1a)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1048
apply (rule ccontr, drule linorder_not_le [THEN iffD2])
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1049
apply (drule_tac x = t in SReal_add, assumption)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1050
apply (drule_tac y = "t + r" in isLubD1 [THEN setleD], auto)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1051
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1052
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  1053
lemma hypreal_setle_less_trans:
20541
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
  1054
     "[| S *<= (x::hypreal); x < y |] ==> S *<= y"
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1055
apply (simp add: setle_def)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1056
apply (auto dest!: bspec order_le_less_trans intro: order_less_imp_le)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1057
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1058
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1059
lemma hypreal_gt_isUb:
20541
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
  1060
     "[| isUb R S (x::hypreal); x < y; y \<in> R |] ==> isUb R S y"
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1061
apply (simp add: isUb_def)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1062
apply (blast intro: hypreal_setle_less_trans)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1063
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1064
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  1065
lemma lemma_st_part_gt_ub:
20541
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
  1066
     "[| (x::hypreal) \<in> HFinite; x < y; y \<in> Reals |]
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  1067
      ==> isUb Reals {s. s \<in> Reals & s < x} y"
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  1068
by (auto dest: order_less_trans intro: order_less_imp_le intro!: isUbI setleI)
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1069
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  1070
lemma lemma_minus_le_zero: "t \<le> t + -r ==> r \<le> (0::hypreal)"
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1071
apply (drule_tac c = "-t" in add_left_mono)
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17298
diff changeset
  1072
apply (auto simp add: add_assoc [symmetric])
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1073
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1074
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  1075
lemma lemma_st_part_le2:
20541
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
  1076
     "[| (x::hypreal) \<in> HFinite;
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1077
         isLub Reals {s. s \<in> Reals & s < x} t;
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1078
         r \<in> Reals; 0 < r |]
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  1079
      ==> t + -r \<le> x"
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1080
apply (frule isLubD1a)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1081
apply (rule ccontr, drule linorder_not_le [THEN iffD1])
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1082
apply (drule SReal_minus, drule_tac x = t in SReal_add, assumption)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1083
apply (drule lemma_st_part_gt_ub, assumption+)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1084
apply (drule isLub_le_isUb, assumption)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1085
apply (drule lemma_minus_le_zero)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1086
apply (auto dest: order_less_le_trans)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1087
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1088
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  1089
lemma lemma_st_part1a:
20541
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
  1090
     "[| (x::hypreal) \<in> HFinite;
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1091
         isLub Reals {s. s \<in> Reals & s < x} t;
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1092
         r \<in> Reals; 0 < r |]
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  1093
      ==> x + -t \<le> r"
15229
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15140
diff changeset
  1094
apply (subgoal_tac "x \<le> t+r") 
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15140
diff changeset
  1095
apply (auto intro: lemma_st_part_le1)
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15140
diff changeset
  1096
done
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1097
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  1098
lemma lemma_st_part2a:
20541
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
  1099
     "[| (x::hypreal) \<in> HFinite;
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1100
         isLub Reals {s. s \<in> Reals & s < x} t;
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1101
         r \<in> Reals;  0 < r |]
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  1102
      ==> -(x + -t) \<le> r"
15229
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15140
diff changeset
  1103
apply (subgoal_tac "(t + -r \<le> x)") 
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15140
diff changeset
  1104
apply (auto intro: lemma_st_part_le2)
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1105
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1106
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  1107
lemma lemma_SReal_ub:
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  1108
     "(x::hypreal) \<in> Reals ==> isUb Reals {s. s \<in> Reals & s < x} x"
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1109
by (auto intro: isUbI setleI order_less_imp_le)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1110
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  1111
lemma lemma_SReal_lub:
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  1112
     "(x::hypreal) \<in> Reals ==> isLub Reals {s. s \<in> Reals & s < x} x"
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1113
apply (auto intro!: isLubI2 lemma_SReal_ub setgeI)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1114
apply (frule isUbD2a)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1115
apply (rule_tac x = x and y = y in linorder_cases)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1116
apply (auto intro!: order_less_imp_le)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1117
apply (drule SReal_dense, assumption, assumption, safe)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1118
apply (drule_tac y = r in isUbD)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1119
apply (auto dest: order_less_le_trans)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1120
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1121
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  1122
lemma lemma_st_part_not_eq1:
20541
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
  1123
     "[| (x::hypreal) \<in> HFinite;
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1124
         isLub Reals {s. s \<in> Reals & s < x} t;
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1125
         r \<in> Reals; 0 < r |]
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1126
      ==> x + -t \<noteq> r"
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1127
apply auto
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1128
apply (frule isLubD1a [THEN SReal_minus])
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1129
apply (drule SReal_add_cancel, assumption)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1130
apply (drule_tac x = x in lemma_SReal_lub)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1131
apply (drule hypreal_isLub_unique, assumption, auto)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1132
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1133
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  1134
lemma lemma_st_part_not_eq2:
20541
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
  1135
     "[| (x::hypreal) \<in> HFinite;
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1136
         isLub Reals {s. s \<in> Reals & s < x} t;
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1137
         r \<in> Reals; 0 < r |]
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1138
      ==> -(x + -t) \<noteq> r"
15539
333a88244569 comprehensive cleanup, replacing sumr by setsum
nipkow
parents: 15531
diff changeset
  1139
apply (auto)
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1140
apply (frule isLubD1a)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1141
apply (drule SReal_add_cancel, assumption)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1142
apply (drule_tac x = "-x" in SReal_minus, simp)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1143
apply (drule_tac x = x in lemma_SReal_lub)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1144
apply (drule hypreal_isLub_unique, assumption, auto)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1145
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1146
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  1147
lemma lemma_st_part_major:
20541
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
  1148
     "[| (x::hypreal) \<in> HFinite;
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1149
         isLub Reals {s. s \<in> Reals & s < x} t;
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1150
         r \<in> Reals; 0 < r |]
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1151
      ==> abs (x + -t) < r"
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1152
apply (frule lemma_st_part1a)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1153
apply (frule_tac [4] lemma_st_part2a, auto)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1154
apply (drule order_le_imp_less_or_eq)+
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1155
apply (auto dest: lemma_st_part_not_eq1 lemma_st_part_not_eq2 simp add: abs_less_iff)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1156
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1157
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  1158
lemma lemma_st_part_major2:
20541
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
  1159
     "[| (x::hypreal) \<in> HFinite; isLub Reals {s. s \<in> Reals & s < x} t |]
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1160
      ==> \<forall>r \<in> Reals. 0 < r --> abs (x + -t) < r"
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  1161
by (blast dest!: lemma_st_part_major)
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1162
15229
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15140
diff changeset
  1163
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15140
diff changeset
  1164
text{*Existence of real and Standard Part Theorem*}
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  1165
lemma lemma_st_part_Ex:
20541
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
  1166
     "(x::hypreal) \<in> HFinite
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
  1167
       ==> \<exists>t \<in> Reals. \<forall>r \<in> Reals. 0 < r --> abs (x + -t) < r"
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1168
apply (frule lemma_st_part_lub, safe)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1169
apply (frule isLubD1a)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1170
apply (blast dest: lemma_st_part_major2)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1171
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1172
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1173
lemma st_part_Ex:
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1174
     "x \<in> HFinite ==> \<exists>t \<in> Reals. x @= t"
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1175
apply (simp add: approx_def Infinitesimal_def)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1176
apply (drule lemma_st_part_Ex, auto)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1177
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1178
15229
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15140
diff changeset
  1179
text{*There is a unique real infinitely close*}
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1180
lemma st_part_Ex1: "x \<in> HFinite ==> EX! t. t \<in> Reals & x @= t"
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1181
apply (drule st_part_Ex, safe)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1182
apply (drule_tac [2] approx_sym, drule_tac [2] approx_sym, drule_tac [2] approx_sym)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1183
apply (auto intro!: approx_unique_real)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1184
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1185
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  1186
subsection{* Finite, Infinite and Infinitesimal*}
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1187
15229
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15140
diff changeset
  1188
lemma HFinite_Int_HInfinite_empty [simp]: "HFinite Int HInfinite = {}"
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1189
apply (simp add: HFinite_def HInfinite_def)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1190
apply (auto dest: order_less_trans)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1191
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1192
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1193
lemma HFinite_not_HInfinite: 
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1194
  assumes x: "x \<in> HFinite" shows "x \<notin> HInfinite"
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1195
proof
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1196
  assume x': "x \<in> HInfinite"
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1197
  with x have "x \<in> HFinite \<inter> HInfinite" by blast
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1198
  thus False by auto
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1199
qed
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1200
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1201
lemma not_HFinite_HInfinite: "x\<notin> HFinite ==> x \<in> HInfinite"
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1202
apply (simp add: HInfinite_def HFinite_def, auto)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1203
apply (drule_tac x = "r + 1" in bspec)
15539
333a88244569 comprehensive cleanup, replacing sumr by setsum
nipkow
parents: 15531
diff changeset
  1204
apply (auto)
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1205
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1206
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1207
lemma HInfinite_HFinite_disj: "x \<in> HInfinite | x \<in> HFinite"
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1208
by (blast intro: not_HFinite_HInfinite)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1209
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1210
lemma HInfinite_HFinite_iff: "(x \<in> HInfinite) = (x \<notin> HFinite)"
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1211
by (blast dest: HFinite_not_HInfinite not_HFinite_HInfinite)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1212
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1213
lemma HFinite_HInfinite_iff: "(x \<in> HFinite) = (x \<notin> HInfinite)"
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  1214
by (simp add: HInfinite_HFinite_iff)
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  1215
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1216
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  1217
lemma HInfinite_diff_HFinite_Infinitesimal_disj:
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  1218
     "x \<notin> Infinitesimal ==> x \<in> HInfinite | x \<in> HFinite - Infinitesimal"
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1219
by (fast intro: not_HFinite_HInfinite)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1220
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  1221
lemma HFinite_inverse:
20541
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
  1222
  fixes x :: "'a::{real_normed_div_algebra,division_by_zero} star"
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
  1223
  shows "[| x \<in> HFinite; x \<notin> Infinitesimal |] ==> inverse x \<in> HFinite"
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1224
apply (cut_tac x = "inverse x" in HInfinite_HFinite_disj)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1225
apply (auto dest!: HInfinite_inverse_Infinitesimal)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1226
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1227
20541
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
  1228
lemma HFinite_inverse2:
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
  1229
  fixes x :: "'a::{real_normed_div_algebra,division_by_zero} star"
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
  1230
  shows "x \<in> HFinite - Infinitesimal ==> inverse x \<in> HFinite"
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1231
by (blast intro: HFinite_inverse)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1232
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1233
(* stronger statement possible in fact *)
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  1234
lemma Infinitesimal_inverse_HFinite:
20541
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
  1235
  fixes x :: "'a::{real_normed_div_algebra,division_by_zero} star"
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
  1236
  shows "x \<notin> Infinitesimal ==> inverse(x) \<in> HFinite"
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1237
apply (drule HInfinite_diff_HFinite_Infinitesimal_disj)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1238
apply (blast intro: HFinite_inverse HInfinite_inverse_Infinitesimal Infinitesimal_subset_HFinite [THEN subsetD])
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1239
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1240
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  1241
lemma HFinite_not_Infinitesimal_inverse:
20541
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
  1242
  fixes x :: "'a::{real_normed_div_algebra,division_by_zero} star"
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
  1243
  shows "x \<in> HFinite - Infinitesimal ==> inverse x \<in> HFinite - Infinitesimal"
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1244
apply (auto intro: Infinitesimal_inverse_HFinite)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1245
apply (drule Infinitesimal_HFinite_mult2, assumption)
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17298
diff changeset
  1246
apply (simp add: not_Infinitesimal_not_zero right_inverse)
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1247
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1248
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  1249
lemma approx_inverse:
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  1250
     "[| x @= y; y \<in>  HFinite - Infinitesimal |]
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1251
      ==> inverse x @= inverse y"
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1252
apply (frule HFinite_diff_Infinitesimal_approx, assumption)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1253
apply (frule not_Infinitesimal_not_zero2)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1254
apply (frule_tac x = x in not_Infinitesimal_not_zero2)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1255
apply (drule HFinite_inverse2)+
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1256
apply (drule approx_mult2, assumption, auto)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1257
apply (drule_tac c = "inverse x" in approx_mult1, assumption)
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17298
diff changeset
  1258
apply (auto intro: approx_sym simp add: mult_assoc)
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1259
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1260
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1261
(*Used for NSLIM_inverse, NSLIMSEQ_inverse*)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1262
lemmas hypreal_of_real_approx_inverse =  hypreal_of_real_HFinite_diff_Infinitesimal [THEN [2] approx_inverse]
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1263
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  1264
lemma inverse_add_Infinitesimal_approx:
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  1265
     "[| x \<in> HFinite - Infinitesimal;
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1266
         h \<in> Infinitesimal |] ==> inverse(x + h) @= inverse x"
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1267
apply (auto intro: approx_inverse approx_sym Infinitesimal_add_approx_self)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1268
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1269
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  1270
lemma inverse_add_Infinitesimal_approx2:
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  1271
     "[| x \<in> HFinite - Infinitesimal;
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1272
         h \<in> Infinitesimal |] ==> inverse(h + x) @= inverse x"
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17298
diff changeset
  1273
apply (rule add_commute [THEN subst])
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1274
apply (blast intro: inverse_add_Infinitesimal_approx)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1275
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1276
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  1277
lemma inverse_add_Infinitesimal_approx_Infinitesimal:
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  1278
     "[| x \<in> HFinite - Infinitesimal;
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1279
         h \<in> Infinitesimal |] ==> inverse(x + h) + -inverse x @= h"
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1280
apply (rule approx_trans2)
15229
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15140
diff changeset
  1281
apply (auto intro: inverse_add_Infinitesimal_approx 
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15140
diff changeset
  1282
            simp add: mem_infmal_iff approx_minus_iff [symmetric])
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1283
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1284
20541
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
  1285
lemma Infinitesimal_square_iff:
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
  1286
  fixes x :: "'a::{real_normed_div_algebra,division_by_zero} star"
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
  1287
  shows "(x \<in> Infinitesimal) = (x*x \<in> Infinitesimal)"
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1288
apply (auto intro: Infinitesimal_mult)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1289
apply (rule ccontr, frule Infinitesimal_inverse_HFinite)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1290
apply (frule not_Infinitesimal_not_zero)
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17298
diff changeset
  1291
apply (auto dest: Infinitesimal_HFinite_mult simp add: mult_assoc)
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1292
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1293
declare Infinitesimal_square_iff [symmetric, simp]
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1294
20541
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
  1295
lemma HFinite_square_iff [simp]:
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
  1296
  fixes x :: "'a::real_normed_div_algebra star"
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
  1297
  shows "(x*x \<in> HFinite) = (x \<in> HFinite)"
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1298
apply (auto intro: HFinite_mult)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1299
apply (auto dest: HInfinite_mult simp add: HFinite_HInfinite_iff)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1300
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1301
20541
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
  1302
lemma HInfinite_square_iff [simp]:
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
  1303
  fixes x :: "'a::real_normed_div_algebra star"
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
  1304
  shows "(x*x \<in> HInfinite) = (x \<in> HInfinite)"
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1305
by (auto simp add: HInfinite_HFinite_iff)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1306
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  1307
lemma approx_HFinite_mult_cancel:
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  1308
     "[| a: HFinite-Infinitesimal; a* w @= a*z |] ==> w @= z"
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1309
apply safe
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1310
apply (frule HFinite_inverse, assumption)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1311
apply (drule not_Infinitesimal_not_zero)
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17298
diff changeset
  1312
apply (auto dest: approx_mult2 simp add: mult_assoc [symmetric])
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1313
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1314
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  1315
lemma approx_HFinite_mult_cancel_iff1:
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  1316
     "a: HFinite-Infinitesimal ==> (a * w @= a * z) = (w @= z)"
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1317
by (auto intro: approx_mult2 approx_HFinite_mult_cancel)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1318
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  1319
lemma HInfinite_HFinite_add_cancel:
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  1320
     "[| x + y \<in> HInfinite; y \<in> HFinite |] ==> x \<in> HInfinite"
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1321
apply (rule ccontr)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1322
apply (drule HFinite_HInfinite_iff [THEN iffD2])
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1323
apply (auto dest: HFinite_add simp add: HInfinite_HFinite_iff)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1324
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1325
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  1326
lemma HInfinite_HFinite_add:
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  1327
     "[| x \<in> HInfinite; y \<in> HFinite |] ==> x + y \<in> HInfinite"
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1328
apply (rule_tac y = "-y" in HInfinite_HFinite_add_cancel)
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17298
diff changeset
  1329
apply (auto simp add: add_assoc HFinite_minus_iff)
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1330
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1331
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  1332
lemma HInfinite_ge_HInfinite:
20541
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
  1333
     "[| (x::hypreal) \<in> HInfinite; x \<le> y; 0 \<le> x |] ==> y \<in> HInfinite"
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1334
by (auto intro: HFinite_bounded simp add: HInfinite_HFinite_iff)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1335
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  1336
lemma Infinitesimal_inverse_HInfinite:
20541
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
  1337
  fixes x :: "'a::real_normed_div_algebra star"
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
  1338
  shows "[| x \<in> Infinitesimal; x \<noteq> 0 |] ==> inverse x \<in> HInfinite"
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1339
apply (rule ccontr, drule HFinite_HInfinite_iff [THEN iffD2])
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1340
apply (auto dest: Infinitesimal_HFinite_mult2)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1341
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1342
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  1343
lemma HInfinite_HFinite_not_Infinitesimal_mult:
20541
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
  1344
  fixes x y :: "'a::{real_normed_div_algebra,division_by_zero} star"
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
  1345
  shows "[| x \<in> HInfinite; y \<in> HFinite - Infinitesimal |]
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1346
      ==> x * y \<in> HInfinite"
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1347
apply (rule ccontr, drule HFinite_HInfinite_iff [THEN iffD2])
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1348
apply (frule HFinite_Infinitesimal_not_zero)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1349
apply (drule HFinite_not_Infinitesimal_inverse)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1350
apply (safe, drule HFinite_mult)
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17298
diff changeset
  1351
apply (auto simp add: mult_assoc HFinite_HInfinite_iff)
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1352
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1353
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  1354
lemma HInfinite_HFinite_not_Infinitesimal_mult2:
20541
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
  1355
  fixes x y :: "'a::{real_normed_div_algebra,division_by_zero} star"
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
  1356
  shows "[| x \<in> HInfinite; y \<in> HFinite - Infinitesimal |]
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1357
      ==> y * x \<in> HInfinite"
20541
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
  1358
apply (rule ccontr, drule HFinite_HInfinite_iff [THEN iffD2])
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
  1359
apply (frule HFinite_Infinitesimal_not_zero)
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
  1360
apply (drule HFinite_not_Infinitesimal_inverse)
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
  1361
apply (safe, drule_tac x="inverse y" in HFinite_mult)
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
  1362
apply assumption
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
  1363
apply (auto simp add: mult_assoc [symmetric] HFinite_HInfinite_iff)
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
  1364
done
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1365
20541
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
  1366
lemma HInfinite_gt_SReal:
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
  1367
  "[| (x::hypreal) \<in> HInfinite; 0 < x; y \<in> Reals |] ==> y < x"
15003
6145dd7538d7 replaced monomorphic abs definitions by abs_if
paulson
parents: 14653
diff changeset
  1368
by (auto dest!: bspec simp add: HInfinite_def abs_if order_less_imp_le)
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1369
20541
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
  1370
lemma HInfinite_gt_zero_gt_one:
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
  1371
  "[| (x::hypreal) \<in> HInfinite; 0 < x |] ==> 1 < x"
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1372
by (auto intro: HInfinite_gt_SReal)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1373
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1374
15229
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15140
diff changeset
  1375
lemma not_HInfinite_one [simp]: "1 \<notin> HInfinite"
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1376
apply (simp (no_asm) add: HInfinite_HFinite_iff)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1377
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1378
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1379
lemma approx_hrabs_disj: "abs x @= x | abs x @= -x"
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1380
by (cut_tac x = x in hrabs_disj, auto)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1381
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1382
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  1383
subsection{*Theorems about Monads*}
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  1384
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  1385
lemma monad_hrabs_Un_subset: "monad (abs x) \<le> monad(x) Un monad(-x)"
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1386
by (rule_tac x1 = x in hrabs_disj [THEN disjE], auto)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1387
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1388
lemma Infinitesimal_monad_eq: "e \<in> Infinitesimal ==> monad (x+e) = monad x"
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1389
by (fast intro!: Infinitesimal_add_approx_self [THEN approx_sym] approx_monad_iff [THEN iffD1])
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1390
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1391
lemma mem_monad_iff: "(u \<in> monad x) = (-u \<in> monad (-x))"
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1392
by (simp add: monad_def)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1393
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1394
lemma Infinitesimal_monad_zero_iff: "(x \<in> Infinitesimal) = (x \<in> monad 0)"
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1395
by (auto intro: approx_sym simp add: monad_def mem_infmal_iff)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1396
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1397
lemma monad_zero_minus_iff: "(x \<in> monad 0) = (-x \<in> monad 0)"
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1398
apply (simp (no_asm) add: Infinitesimal_monad_zero_iff [symmetric])
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1399
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1400
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1401
lemma monad_zero_hrabs_iff: "(x \<in> monad 0) = (abs x \<in> monad 0)"
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1402
apply (rule_tac x1 = x in hrabs_disj [THEN disjE])
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1403
apply (auto simp add: monad_zero_minus_iff [symmetric])
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1404
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1405
15229
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15140
diff changeset
  1406
lemma mem_monad_self [simp]: "x \<in> monad x"
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1407
by (simp add: monad_def)
15229
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15140
diff changeset
  1408
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1409
15229
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15140
diff changeset
  1410
subsection{*Proof that @{term "x @= y"} implies @{term"\<bar>x\<bar> @= \<bar>y\<bar>"}*}
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15140
diff changeset
  1411
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15140
diff changeset
  1412
lemma approx_subset_monad: "x @= y ==> {x,y} \<le> monad x"
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1413
apply (simp (no_asm))
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1414
apply (simp add: approx_monad_iff)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1415
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1416
15229
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15140
diff changeset
  1417
lemma approx_subset_monad2: "x @= y ==> {x,y} \<le> monad y"
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1418
apply (drule approx_sym)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1419
apply (fast dest: approx_subset_monad)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1420
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1421
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1422
lemma mem_monad_approx: "u \<in> monad x ==> x @= u"
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1423
by (simp add: monad_def)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1424
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1425
lemma approx_mem_monad: "x @= u ==> u \<in> monad x"
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1426
by (simp add: monad_def)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1427
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1428
lemma approx_mem_monad2: "x @= u ==> x \<in> monad u"
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1429
apply (simp add: monad_def)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1430
apply (blast intro!: approx_sym)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1431
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1432
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1433
lemma approx_mem_monad_zero: "[| x @= y;x \<in> monad 0 |] ==> y \<in> monad 0"
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1434
apply (drule mem_monad_approx)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1435
apply (fast intro: approx_mem_monad approx_trans)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1436
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1437
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  1438
lemma Infinitesimal_approx_hrabs:
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  1439
     "[| x @= y; x \<in> Infinitesimal |] ==> abs x @= abs y"
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1440
apply (drule Infinitesimal_monad_zero_iff [THEN iffD1])
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1441
apply (blast intro: approx_mem_monad_zero monad_zero_hrabs_iff [THEN iffD1] mem_monad_approx approx_trans3)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1442
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1443
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  1444
lemma less_Infinitesimal_less:
20541
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
  1445
     "[| 0 < x;  (x::hypreal) \<notin>Infinitesimal;  e :Infinitesimal |] ==> e < x"
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1446
apply (rule ccontr)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1447
apply (auto intro: Infinitesimal_zero [THEN [2] Infinitesimal_interval] 
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1448
            dest!: order_le_imp_less_or_eq simp add: linorder_not_less)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1449
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1450
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  1451
lemma Ball_mem_monad_gt_zero:
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  1452
     "[| 0 < x;  x \<notin> Infinitesimal; u \<in> monad x |] ==> 0 < u"
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1453
apply (drule mem_monad_approx [THEN approx_sym])
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1454
apply (erule bex_Infinitesimal_iff2 [THEN iffD2, THEN bexE])
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1455
apply (drule_tac e = "-xa" in less_Infinitesimal_less, auto)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1456
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1457
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  1458
lemma Ball_mem_monad_less_zero:
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  1459
     "[| x < 0; x \<notin> Infinitesimal; u \<in> monad x |] ==> u < 0"
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1460
apply (drule mem_monad_approx [THEN approx_sym])
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1461
apply (erule bex_Infinitesimal_iff [THEN iffD2, THEN bexE])
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1462
apply (cut_tac x = "-x" and e = xa in less_Infinitesimal_less, auto)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1463
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1464
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  1465
lemma lemma_approx_gt_zero:
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  1466
     "[|0 < x; x \<notin> Infinitesimal; x @= y|] ==> 0 < y"
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1467
by (blast dest: Ball_mem_monad_gt_zero approx_subset_monad)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1468
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  1469
lemma lemma_approx_less_zero:
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  1470
     "[|x < 0; x \<notin> Infinitesimal; x @= y|] ==> y < 0"
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1471
by (blast dest: Ball_mem_monad_less_zero approx_subset_monad)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1472
15229
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15140
diff changeset
  1473
theorem approx_hrabs: "x @= y ==> abs x @= abs y"
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15140
diff changeset
  1474
apply (case_tac "x \<in> Infinitesimal") 
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15140
diff changeset
  1475
apply (simp add: Infinitesimal_approx_hrabs)
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15140
diff changeset
  1476
apply (rule linorder_cases [of 0 x])  
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15140
diff changeset
  1477
apply (frule lemma_approx_gt_zero [of x y]) 
16775
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 15539
diff changeset
  1478
apply (auto simp add: lemma_approx_less_zero [of x y] abs_of_neg)
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1479
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1480
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1481
lemma approx_hrabs_zero_cancel: "abs(x) @= 0 ==> x @= 0"
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1482
apply (cut_tac x = x in hrabs_disj)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1483
apply (auto dest: approx_minus)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1484
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1485
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1486
lemma approx_hrabs_add_Infinitesimal: "e \<in> Infinitesimal ==> abs x @= abs(x+e)"
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1487
by (fast intro: approx_hrabs Infinitesimal_add_approx_self)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1488
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  1489
lemma approx_hrabs_add_minus_Infinitesimal:
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  1490
     "e \<in> Infinitesimal ==> abs x @= abs(x + -e)"
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1491
by (fast intro: approx_hrabs Infinitesimal_add_minus_approx_self)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1492
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  1493
lemma hrabs_add_Infinitesimal_cancel:
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  1494
     "[| e \<in> Infinitesimal; e' \<in> Infinitesimal;
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1495
         abs(x+e) = abs(y+e')|] ==> abs x @= abs y"
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1496
apply (drule_tac x = x in approx_hrabs_add_Infinitesimal)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1497
apply (drule_tac x = y in approx_hrabs_add_Infinitesimal)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1498
apply (auto intro: approx_trans2)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1499
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1500
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  1501
lemma hrabs_add_minus_Infinitesimal_cancel:
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  1502
     "[| e \<in> Infinitesimal; e' \<in> Infinitesimal;
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1503
         abs(x + -e) = abs(y + -e')|] ==> abs x @= abs y"
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1504
apply (drule_tac x = x in approx_hrabs_add_minus_Infinitesimal)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1505
apply (drule_tac x = y in approx_hrabs_add_minus_Infinitesimal)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1506
apply (auto intro: approx_trans2)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1507
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1508
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1509
(* interesting slightly counterintuitive theorem: necessary
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1510
   for proving that an open interval is an NS open set
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1511
*)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1512
lemma Infinitesimal_add_hypreal_of_real_less:
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1513
     "[| x < y;  u \<in> Infinitesimal |]
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1514
      ==> hypreal_of_real x + u < hypreal_of_real y"
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1515
apply (simp add: Infinitesimal_def)
17431
70311ad8bf11 fix names in hypreal_arith.ML
huffman
parents: 17429
diff changeset
  1516
apply (drule_tac x = "hypreal_of_real y + -hypreal_of_real x" in bspec, simp)
20254
58b71535ed00 lin_arith_prover splits certain operators (e.g. min, max, abs)
webertj
parents: 20217
diff changeset
  1517
apply (simp add: abs_less_iff)
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1518
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1519
14387
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
  1520
lemma Infinitesimal_add_hrabs_hypreal_of_real_less:
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
  1521
     "[| x \<in> Infinitesimal; abs(hypreal_of_real r) < hypreal_of_real y |]
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1522
      ==> abs (hypreal_of_real r + x) < hypreal_of_real y"
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1523
apply (drule_tac x = "hypreal_of_real r" in approx_hrabs_add_Infinitesimal)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1524
apply (drule approx_sym [THEN bex_Infinitesimal_iff2 [THEN iffD2]])
15229
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15140
diff changeset
  1525
apply (auto intro!: Infinitesimal_add_hypreal_of_real_less
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17298
diff changeset
  1526
            simp del: star_of_abs
15229
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15140
diff changeset
  1527
            simp add: hypreal_of_real_hrabs)
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1528
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1529
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  1530
lemma Infinitesimal_add_hrabs_hypreal_of_real_less2:
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  1531
     "[| x \<in> Infinitesimal;  abs(hypreal_of_real r) < hypreal_of_real y |]
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1532
      ==> abs (x + hypreal_of_real r) < hypreal_of_real y"
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17298
diff changeset
  1533
apply (rule add_commute [THEN subst])
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1534
apply (erule Infinitesimal_add_hrabs_hypreal_of_real_less, assumption)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1535
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1536
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  1537
lemma hypreal_of_real_le_add_Infininitesimal_cancel:
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  1538
     "[| u \<in> Infinitesimal; v \<in> Infinitesimal;
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  1539
         hypreal_of_real x + u \<le> hypreal_of_real y + v |]
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  1540
      ==> hypreal_of_real x \<le> hypreal_of_real y"
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1541
apply (simp add: linorder_not_less [symmetric], auto)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1542
apply (drule_tac u = "v-u" in Infinitesimal_add_hypreal_of_real_less)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1543
apply (auto simp add: Infinitesimal_diff)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1544
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1545
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  1546
lemma hypreal_of_real_le_add_Infininitesimal_cancel2:
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  1547
     "[| u \<in> Infinitesimal; v \<in> Infinitesimal;
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  1548
         hypreal_of_real x + u \<le> hypreal_of_real y + v |]
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  1549
      ==> x \<le> y"
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17298
diff changeset
  1550
by (blast intro: star_of_le [THEN iffD1] 
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17298
diff changeset
  1551
          intro!: hypreal_of_real_le_add_Infininitesimal_cancel)
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1552
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  1553
lemma hypreal_of_real_less_Infinitesimal_le_zero:
15229
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15140
diff changeset
  1554
    "[| hypreal_of_real x < e; e \<in> Infinitesimal |] ==> hypreal_of_real x \<le> 0"
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1555
apply (rule linorder_not_less [THEN iffD1], safe)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1556
apply (drule Infinitesimal_interval)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1557
apply (drule_tac [4] SReal_hypreal_of_real [THEN SReal_Infinitesimal_zero], auto)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1558
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1559
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1560
(*used once, in Lim/NSDERIV_inverse*)
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  1561
lemma Infinitesimal_add_not_zero:
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  1562
     "[| h \<in> Infinitesimal; x \<noteq> 0 |] ==> hypreal_of_real x + h \<noteq> 0"
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1563
apply auto
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1564
apply (subgoal_tac "h = - hypreal_of_real x", auto)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1565
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1566
15229
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15140
diff changeset
  1567
lemma Infinitesimal_square_cancel [simp]:
20541
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
  1568
     "(x::hypreal)*x + y*y \<in> Infinitesimal ==> x*x \<in> Infinitesimal"
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1569
apply (rule Infinitesimal_interval2)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1570
apply (rule_tac [3] zero_le_square, assumption)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1571
apply (auto simp add: zero_le_square)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1572
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1573
20541
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
  1574
lemma HFinite_square_cancel [simp]:
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
  1575
  "(x::hypreal)*x + y*y \<in> HFinite ==> x*x \<in> HFinite"
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1576
apply (rule HFinite_bounded, assumption)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1577
apply (auto simp add: zero_le_square)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1578
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1579
15229
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15140
diff changeset
  1580
lemma Infinitesimal_square_cancel2 [simp]:
20541
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
  1581
     "(x::hypreal)*x + y*y \<in> Infinitesimal ==> y*y \<in> Infinitesimal"
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1582
apply (rule Infinitesimal_square_cancel)
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17298
diff changeset
  1583
apply (rule add_commute [THEN subst])
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1584
apply (simp (no_asm))
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1585
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1586
20541
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
  1587
lemma HFinite_square_cancel2 [simp]:
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
  1588
  "(x::hypreal)*x + y*y \<in> HFinite ==> y*y \<in> HFinite"
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1589
apply (rule HFinite_square_cancel)
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17298
diff changeset
  1590
apply (rule add_commute [THEN subst])
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1591
apply (simp (no_asm))
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1592
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1593
15229
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15140
diff changeset
  1594
lemma Infinitesimal_sum_square_cancel [simp]:
20541
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
  1595
     "(x::hypreal)*x + y*y + z*z \<in> Infinitesimal ==> x*x \<in> Infinitesimal"
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1596
apply (rule Infinitesimal_interval2, assumption)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1597
apply (rule_tac [2] zero_le_square, simp)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1598
apply (insert zero_le_square [of y]) 
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1599
apply (insert zero_le_square [of z], simp)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1600
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1601
15229
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15140
diff changeset
  1602
lemma HFinite_sum_square_cancel [simp]:
20541
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
  1603
     "(x::hypreal)*x + y*y + z*z \<in> HFinite ==> x*x \<in> HFinite"
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1604
apply (rule HFinite_bounded, assumption)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1605
apply (rule_tac [2] zero_le_square)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1606
apply (insert zero_le_square [of y]) 
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1607
apply (insert zero_le_square [of z], simp)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1608
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1609
15229
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15140
diff changeset
  1610
lemma Infinitesimal_sum_square_cancel2 [simp]:
20541
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
  1611
     "(y::hypreal)*y + x*x + z*z \<in> Infinitesimal ==> x*x \<in> Infinitesimal"
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1612
apply (rule Infinitesimal_sum_square_cancel)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1613
apply (simp add: add_ac)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1614
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1615
15229
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15140
diff changeset
  1616
lemma HFinite_sum_square_cancel2 [simp]:
20541
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
  1617
     "(y::hypreal)*y + x*x + z*z \<in> HFinite ==> x*x \<in> HFinite"
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1618
apply (rule HFinite_sum_square_cancel)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1619
apply (simp add: add_ac)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1620
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1621
15229
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15140
diff changeset
  1622
lemma Infinitesimal_sum_square_cancel3 [simp]:
20541
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
  1623
     "(z::hypreal)*z + y*y + x*x \<in> Infinitesimal ==> x*x \<in> Infinitesimal"
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1624
apply (rule Infinitesimal_sum_square_cancel)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1625
apply (simp add: add_ac)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1626
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1627
15229
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15140
diff changeset
  1628
lemma HFinite_sum_square_cancel3 [simp]:
20541
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
  1629
     "(z::hypreal)*z + y*y + x*x \<in> HFinite ==> x*x \<in> HFinite"
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1630
apply (rule HFinite_sum_square_cancel)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1631
apply (simp add: add_ac)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1632
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1633
15229
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15140
diff changeset
  1634
lemma monad_hrabs_less:
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15140
diff changeset
  1635
     "[| y \<in> monad x; 0 < hypreal_of_real e |]
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1636
      ==> abs (y + -x) < hypreal_of_real e"
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1637
apply (drule mem_monad_approx [THEN approx_sym])
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1638
apply (drule bex_Infinitesimal_iff [THEN iffD2])
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1639
apply (auto dest!: InfinitesimalD)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1640
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1641
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  1642
lemma mem_monad_SReal_HFinite:
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  1643
     "x \<in> monad (hypreal_of_real  a) ==> x \<in> HFinite"
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1644
apply (drule mem_monad_approx [THEN approx_sym])
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1645
apply (drule bex_Infinitesimal_iff2 [THEN iffD2])
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1646
apply (safe dest!: Infinitesimal_subset_HFinite [THEN subsetD])
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1647
apply (erule SReal_hypreal_of_real [THEN SReal_subset_HFinite [THEN subsetD], THEN HFinite_add])
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1648
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1649
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  1650
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  1651
subsection{* Theorems about Standard Part*}
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1652
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1653
lemma st_approx_self: "x \<in> HFinite ==> st x @= x"
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1654
apply (simp add: st_def)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1655
apply (frule st_part_Ex, safe)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1656
apply (rule someI2)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1657
apply (auto intro: approx_sym)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1658
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1659
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1660
lemma st_SReal: "x \<in> HFinite ==> st x \<in> Reals"
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1661
apply (simp add: st_def)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1662
apply (frule st_part_Ex, safe)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1663
apply (rule someI2)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1664
apply (auto intro: approx_sym)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1665
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1666
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1667
lemma st_HFinite: "x \<in> HFinite ==> st x \<in> HFinite"
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1668
by (erule st_SReal [THEN SReal_subset_HFinite [THEN subsetD]])
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1669
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1670
lemma st_SReal_eq: "x \<in> Reals ==> st x = x"
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1671
apply (simp add: st_def)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1672
apply (rule some_equality)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1673
apply (fast intro: SReal_subset_HFinite [THEN subsetD])
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1674
apply (blast dest: SReal_approx_iff [THEN iffD1])
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1675
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1676
15229
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15140
diff changeset
  1677
lemma st_hypreal_of_real [simp]: "st (hypreal_of_real x) = hypreal_of_real x"
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1678
by (rule SReal_hypreal_of_real [THEN st_SReal_eq])
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1679
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1680
lemma st_eq_approx: "[| x \<in> HFinite; y \<in> HFinite; st x = st y |] ==> x @= y"
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1681
by (auto dest!: st_approx_self elim!: approx_trans3)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1682
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1683
lemma approx_st_eq: 
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1684
  assumes "x \<in> HFinite" and "y \<in> HFinite" and "x @= y" 
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1685
  shows "st x = st y"
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1686
proof -
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1687
  have "st x @= x" "st y @= y" "st x \<in> Reals" "st y \<in> Reals"
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1688
    by (simp_all add: st_approx_self st_SReal prems) 
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1689
  with prems show ?thesis 
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1690
    by (fast elim: approx_trans approx_trans2 SReal_approx_iff [THEN iffD1])
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1691
qed
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1692
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  1693
lemma st_eq_approx_iff:
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  1694
     "[| x \<in> HFinite; y \<in> HFinite|]
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1695
                   ==> (x @= y) = (st x = st y)"
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1696
by (blast intro: approx_st_eq st_eq_approx)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1697
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  1698
lemma st_Infinitesimal_add_SReal:
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  1699
     "[| x \<in> Reals; e \<in> Infinitesimal |] ==> st(x + e) = x"
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1700
apply (frule st_SReal_eq [THEN subst])
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1701
prefer 2 apply assumption
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1702
apply (frule SReal_subset_HFinite [THEN subsetD])
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1703
apply (frule Infinitesimal_subset_HFinite [THEN subsetD])
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1704
apply (drule st_SReal_eq)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1705
apply (rule approx_st_eq)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1706
apply (auto intro: HFinite_add simp add: Infinitesimal_add_approx_self [THEN approx_sym])
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1707
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1708
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  1709
lemma st_Infinitesimal_add_SReal2:
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  1710
     "[| x \<in> Reals; e \<in> Infinitesimal |] ==> st(e + x) = x"
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17298
diff changeset
  1711
apply (rule add_commute [THEN subst])
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1712
apply (blast intro!: st_Infinitesimal_add_SReal)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1713
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1714
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  1715
lemma HFinite_st_Infinitesimal_add:
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  1716
     "x \<in> HFinite ==> \<exists>e \<in> Infinitesimal. x = st(x) + e"
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  1717
by (blast dest!: st_approx_self [THEN approx_sym] bex_Infinitesimal_iff2 [THEN iffD2])
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1718
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1719
lemma st_add: 
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1720
  assumes x: "x \<in> HFinite" and y: "y \<in> HFinite"
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1721
  shows "st (x + y) = st(x) + st(y)"
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1722
proof -
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1723
  from HFinite_st_Infinitesimal_add [OF x]
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1724
  obtain ex where ex: "ex \<in> Infinitesimal" "st x + ex = x" 
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1725
    by (blast intro: sym)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1726
  from HFinite_st_Infinitesimal_add [OF y]
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1727
  obtain ey where ey: "ey \<in> Infinitesimal" "st y + ey = y" 
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1728
    by (blast intro: sym)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1729
  have "st (x + y) = st ((st x + ex) + (st y + ey))"
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1730
    by (simp add: ex ey) 
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1731
  also have "... = st ((ex + ey) + (st x + st y))" by (simp add: add_ac)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1732
  also have "... = st x + st y" 
15539
333a88244569 comprehensive cleanup, replacing sumr by setsum
nipkow
parents: 15531
diff changeset
  1733
    by (simp add: prems st_SReal Infinitesimal_add 
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1734
                  st_Infinitesimal_add_SReal2) 
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1735
  finally show ?thesis .
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1736
qed
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1737
15229
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15140
diff changeset
  1738
lemma st_number_of [simp]: "st (number_of w) = number_of w"
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1739
by (rule SReal_number_of [THEN st_SReal_eq])
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1740
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1741
(*the theorem above for the special cases of zero and one*)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1742
lemma [simp]: "st 0 = 0" "st 1 = 1"
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1743
by (simp_all add: st_SReal_eq)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1744
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1745
lemma st_minus: assumes "y \<in> HFinite" shows "st(-y) = -st(y)"
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1746
proof -
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1747
  have "st (- y) + st y = 0"
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1748
   by (simp add: prems st_add [symmetric] HFinite_minus_iff) 
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1749
  thus ?thesis by arith
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1750
qed
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1751
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  1752
lemma st_diff: "[| x \<in> HFinite; y \<in> HFinite |] ==> st (x-y) = st(x) - st(y)"
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17298
diff changeset
  1753
apply (simp add: diff_def)
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1754
apply (frule_tac y1 = y in st_minus [symmetric])
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1755
apply (drule_tac x1 = y in HFinite_minus_iff [THEN iffD2])
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1756
apply (simp (no_asm_simp) add: st_add)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1757
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1758
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  1759
lemma lemma_st_mult:
20541
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
  1760
  fixes x y e ea :: "'a::real_normed_algebra star"
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
  1761
  shows "[| x \<in> HFinite; y \<in> HFinite; e \<in> Infinitesimal; ea \<in> Infinitesimal |]
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  1762
      ==> e*y + x*ea + e*ea \<in> Infinitesimal"
20541
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
  1763
apply (intro Infinitesimal_add)
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
  1764
apply (erule (1) Infinitesimal_HFinite_mult)
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
  1765
apply (erule (1) Infinitesimal_HFinite_mult2)
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
  1766
apply (erule (1) Infinitesimal_mult)
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1767
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1768
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  1769
lemma st_mult: "[| x \<in> HFinite; y \<in> HFinite |] ==> st (x * y) = st(x) * st(y)"
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1770
apply (frule HFinite_st_Infinitesimal_add)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1771
apply (frule_tac x = y in HFinite_st_Infinitesimal_add, safe)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1772
apply (subgoal_tac "st (x * y) = st ((st x + e) * (st y + ea))")
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1773
apply (drule_tac [2] sym, drule_tac [2] sym)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1774
 prefer 2 apply simp 
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1775
apply (erule_tac V = "x = st x + e" in thin_rl)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1776
apply (erule_tac V = "y = st y + ea" in thin_rl)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1777
apply (simp add: left_distrib right_distrib)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1778
apply (drule st_SReal)+
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17298
diff changeset
  1779
apply (simp (no_asm_use) add: add_assoc)
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1780
apply (rule st_Infinitesimal_add_SReal)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1781
apply (blast intro!: SReal_mult)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1782
apply (drule SReal_subset_HFinite [THEN subsetD])+
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17298
diff changeset
  1783
apply (rule add_assoc [THEN subst])
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1784
apply (blast intro!: lemma_st_mult)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1785
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1786
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1787
lemma st_Infinitesimal: "x \<in> Infinitesimal ==> st x = 0"
14387
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
  1788
apply (subst numeral_0_eq_0 [symmetric])
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1789
apply (rule st_number_of [THEN subst])
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1790
apply (rule approx_st_eq)
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  1791
apply (auto intro: Infinitesimal_subset_HFinite [THEN subsetD]
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  1792
            simp add: mem_infmal_iff [symmetric])
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1793
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1794
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1795
lemma st_not_Infinitesimal: "st(x) \<noteq> 0 ==> x \<notin> Infinitesimal"
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1796
by (fast intro: st_Infinitesimal)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1797
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  1798
lemma st_inverse:
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  1799
     "[| x \<in> HFinite; st x \<noteq> 0 |]
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1800
      ==> st(inverse x) = inverse (st x)"
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1801
apply (rule_tac c1 = "st x" in hypreal_mult_left_cancel [THEN iffD1])
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1802
apply (auto simp add: st_mult [symmetric] st_not_Infinitesimal HFinite_inverse)
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17298
diff changeset
  1803
apply (subst right_inverse, auto)
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1804
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1805
15229
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15140
diff changeset
  1806
lemma st_divide [simp]:
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  1807
     "[| x \<in> HFinite; y \<in> HFinite; st y \<noteq> 0 |]
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1808
      ==> st(x/y) = (st x) / (st y)"
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17298
diff changeset
  1809
by (simp add: divide_inverse st_mult st_not_Infinitesimal HFinite_inverse st_inverse)
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1810
15229
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15140
diff changeset
  1811
lemma st_idempotent [simp]: "x \<in> HFinite ==> st(st(x)) = st(x)"
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1812
by (blast intro: st_HFinite st_approx_self approx_st_eq)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1813
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  1814
lemma Infinitesimal_add_st_less:
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  1815
     "[| x \<in> HFinite; y \<in> HFinite; u \<in> Infinitesimal; st x < st y |] 
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  1816
      ==> st x + u < st y"
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1817
apply (drule st_SReal)+
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1818
apply (auto intro!: Infinitesimal_add_hypreal_of_real_less simp add: SReal_iff)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1819
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1820
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  1821
lemma Infinitesimal_add_st_le_cancel:
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  1822
     "[| x \<in> HFinite; y \<in> HFinite;
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  1823
         u \<in> Infinitesimal; st x \<le> st y + u
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  1824
      |] ==> st x \<le> st y"
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1825
apply (simp add: linorder_not_less [symmetric])
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1826
apply (auto dest: Infinitesimal_add_st_less)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1827
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1828
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  1829
lemma st_le: "[| x \<in> HFinite; y \<in> HFinite; x \<le> y |] ==> st(x) \<le> st(y)"
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1830
apply (frule HFinite_st_Infinitesimal_add)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1831
apply (rotate_tac 1)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1832
apply (frule HFinite_st_Infinitesimal_add, safe)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1833
apply (rule Infinitesimal_add_st_le_cancel)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1834
apply (rule_tac [3] x = ea and y = e in Infinitesimal_diff)
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17298
diff changeset
  1835
apply (auto simp add: add_assoc [symmetric])
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1836
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1837
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  1838
lemma st_zero_le: "[| 0 \<le> x;  x \<in> HFinite |] ==> 0 \<le> st x"
14387
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
  1839
apply (subst numeral_0_eq_0 [symmetric])
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1840
apply (rule st_number_of [THEN subst])
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1841
apply (rule st_le, auto)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1842
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1843
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  1844
lemma st_zero_ge: "[| x \<le> 0;  x \<in> HFinite |] ==> st x \<le> 0"
14387
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
  1845
apply (subst numeral_0_eq_0 [symmetric])
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1846
apply (rule st_number_of [THEN subst])
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1847
apply (rule st_le, auto)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1848
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1849
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1850
lemma st_hrabs: "x \<in> HFinite ==> abs(st x) = st(abs x)"
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1851
apply (simp add: linorder_not_le st_zero_le abs_if st_minus
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1852
   linorder_not_less)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1853
apply (auto dest!: st_zero_ge [OF order_less_imp_le]) 
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1854
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1855
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1856
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1857
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  1858
subsection{*Alternative Definitions for @{term HFinite} using Free Ultrafilter*}
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1859
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  1860
lemma FreeUltrafilterNat_Rep_hypreal:
17298
ad73fb6144cf replace type hypreal with real star
huffman
parents: 16924
diff changeset
  1861
     "[| X \<in> Rep_star x; Y \<in> Rep_star x |]
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1862
      ==> {n. X n = Y n} \<in> FreeUltrafilterNat"
17429
e8d6ed3aacfe merged Transfer.thy and StarType.thy into StarDef.thy; renamed Ifun2_of to starfun2; cleaned up
huffman
parents: 17318
diff changeset
  1863
by (cases x, unfold star_n_def, auto, ultra)
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1864
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1865
lemma HFinite_FreeUltrafilterNat:
20541
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
  1866
    "(x::hypreal) \<in> HFinite 
17298
ad73fb6144cf replace type hypreal with real star
huffman
parents: 16924
diff changeset
  1867
     ==> \<exists>X \<in> Rep_star x. \<exists>u. {n. abs (X n) < u} \<in> FreeUltrafilterNat"
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17298
diff changeset
  1868
apply (cases x)
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1869
apply (auto simp add: HFinite_def abs_less_iff minus_less_iff [of x] 
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17298
diff changeset
  1870
              star_of_def
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17298
diff changeset
  1871
              star_n_less SReal_iff star_n_minus)
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17298
diff changeset
  1872
apply (rule_tac x=X in bexI)
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17298
diff changeset
  1873
apply (rule_tac x=y in exI, ultra)
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17298
diff changeset
  1874
apply simp
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1875
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1876
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1877
lemma FreeUltrafilterNat_HFinite:
17298
ad73fb6144cf replace type hypreal with real star
huffman
parents: 16924
diff changeset
  1878
     "\<exists>X \<in> Rep_star x.
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1879
       \<exists>u. {n. abs (X n) < u} \<in> FreeUltrafilterNat
20541
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
  1880
       ==>  (x::hypreal) \<in> HFinite"
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17298
diff changeset
  1881
apply (cases x)
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1882
apply (auto simp add: HFinite_def abs_less_iff minus_less_iff [of x])
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1883
apply (rule_tac x = "hypreal_of_real u" in bexI)
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17298
diff changeset
  1884
apply (auto simp add: star_n_less SReal_iff star_n_minus star_of_def)
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17298
diff changeset
  1885
apply ultra+
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1886
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1887
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  1888
lemma HFinite_FreeUltrafilterNat_iff:
20541
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
  1889
     "((x::hypreal) \<in> HFinite) = (\<exists>X \<in> Rep_star x.
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1890
           \<exists>u. {n. abs (X n) < u} \<in> FreeUltrafilterNat)"
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17298
diff changeset
  1891
by (blast intro!: HFinite_FreeUltrafilterNat FreeUltrafilterNat_HFinite)
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1892
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  1893
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  1894
subsection{*Alternative Definitions for @{term HInfinite} using Free Ultrafilter*}
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  1895
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  1896
lemma lemma_Compl_eq: "- {n. (u::real) < abs (xa n)} = {n. abs (xa n) \<le> u}"
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1897
by auto
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1898
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  1899
lemma lemma_Compl_eq2: "- {n. abs (xa n) < (u::real)} = {n. u \<le> abs (xa n)}"
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1900
by auto
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1901
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  1902
lemma lemma_Int_eq1:
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  1903
     "{n. abs (xa n) \<le> (u::real)} Int {n. u \<le> abs (xa n)}
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1904
          = {n. abs(xa n) = u}"
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17298
diff changeset
  1905
by auto
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1906
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  1907
lemma lemma_FreeUltrafilterNat_one:
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  1908
     "{n. abs (xa n) = u} \<le> {n. abs (xa n) < u + (1::real)}"
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1909
by auto
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1910
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1911
(*-------------------------------------
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1912
  Exclude this type of sets from free
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1913
  ultrafilter for Infinite numbers!
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1914
 -------------------------------------*)
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  1915
lemma FreeUltrafilterNat_const_Finite:
17298
ad73fb6144cf replace type hypreal with real star
huffman
parents: 16924
diff changeset
  1916
     "[| xa: Rep_star x;
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1917
                  {n. abs (xa n) = u} \<in> FreeUltrafilterNat
20541
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
  1918
               |] ==> (x::hypreal) \<in> HFinite"
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1919
apply (rule FreeUltrafilterNat_HFinite)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1920
apply (rule_tac x = xa in bexI)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1921
apply (rule_tac x = "u + 1" in exI)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1922
apply (ultra, assumption)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1923
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1924
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1925
lemma HInfinite_FreeUltrafilterNat:
20541
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
  1926
     "(x::hypreal) \<in> HInfinite ==> \<exists>X \<in> Rep_star x.
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1927
           \<forall>u. {n. u < abs (X n)} \<in> FreeUltrafilterNat"
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1928
apply (frule HInfinite_HFinite_iff [THEN iffD1])
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1929
apply (cut_tac x = x in Rep_hypreal_nonempty)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1930
apply (auto simp del: Rep_hypreal_nonempty simp add: HFinite_FreeUltrafilterNat_iff Bex_def)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1931
apply (drule spec)+
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1932
apply auto
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1933
apply (drule_tac x = u in spec)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1934
apply (drule FreeUltrafilterNat_Compl_mem)+
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1935
apply (drule FreeUltrafilterNat_Int, assumption)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1936
apply (simp add: lemma_Compl_eq lemma_Compl_eq2 lemma_Int_eq1)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1937
apply (auto dest: FreeUltrafilterNat_const_Finite simp
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1938
      add: HInfinite_HFinite_iff [THEN iffD1])
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1939
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1940
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  1941
lemma lemma_Int_HI:
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  1942
     "{n. abs (Xa n) < u} Int {n. X n = Xa n} \<subseteq> {n. abs (X n) < (u::real)}"
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  1943
by auto
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1944
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1945
lemma lemma_Int_HIa: "{n. u < abs (X n)} Int {n. abs (X n) < (u::real)} = {}"
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1946
by (auto intro: order_less_asym)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1947
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  1948
lemma FreeUltrafilterNat_HInfinite:
17298
ad73fb6144cf replace type hypreal with real star
huffman
parents: 16924
diff changeset
  1949
     "\<exists>X \<in> Rep_star x. \<forall>u.
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1950
               {n. u < abs (X n)} \<in> FreeUltrafilterNat
20541
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
  1951
               ==>  (x::hypreal) \<in> HInfinite"
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1952
apply (rule HInfinite_HFinite_iff [THEN iffD2])
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1953
apply (safe, drule HFinite_FreeUltrafilterNat, auto)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1954
apply (drule_tac x = u in spec)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1955
apply (drule FreeUltrafilterNat_Rep_hypreal, assumption)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1956
apply (drule_tac Y = "{n. X n = Xa n}" in FreeUltrafilterNat_Int, simp) 
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1957
apply (drule lemma_Int_HI [THEN [2] FreeUltrafilterNat_subset])
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1958
apply (drule_tac Y = "{n. abs (X n) < u}" in FreeUltrafilterNat_Int)
15539
333a88244569 comprehensive cleanup, replacing sumr by setsum
nipkow
parents: 15531
diff changeset
  1959
apply (auto simp add: lemma_Int_HIa)
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1960
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1961
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  1962
lemma HInfinite_FreeUltrafilterNat_iff:
20541
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
  1963
     "((x::hypreal) \<in> HInfinite) = (\<exists>X \<in> Rep_star x.
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1964
           \<forall>u. {n. u < abs (X n)} \<in> FreeUltrafilterNat)"
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  1965
by (blast intro!: HInfinite_FreeUltrafilterNat FreeUltrafilterNat_HInfinite)
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1966
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1967
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  1968
subsection{*Alternative Definitions for @{term Infinitesimal} using Free Ultrafilter*}
10751
a81ea5d3dd41 separation of HOL-Hyperreal from HOL-Real
paulson
parents:
diff changeset
  1969
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1970
lemma Infinitesimal_FreeUltrafilterNat:
20541
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
  1971
          "(x::hypreal) \<in> Infinitesimal ==> \<exists>X \<in> Rep_star x.
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1972
           \<forall>u. 0 < u --> {n. abs (X n) < u} \<in> FreeUltrafilterNat"
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1973
apply (simp add: Infinitesimal_def)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1974
apply (auto simp add: abs_less_iff minus_less_iff [of x])
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17298
diff changeset
  1975
apply (cases x)
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17298
diff changeset
  1976
apply (auto, rule bexI [OF _ Rep_star_star_n], safe)
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17298
diff changeset
  1977
apply (drule star_of_less [THEN iffD2])
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1978
apply (drule_tac x = "hypreal_of_real u" in bspec, auto)
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17298
diff changeset
  1979
apply (auto simp add: star_n_less star_n_minus star_of_def, ultra)
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1980
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1981
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1982
lemma FreeUltrafilterNat_Infinitesimal:
17298
ad73fb6144cf replace type hypreal with real star
huffman
parents: 16924
diff changeset
  1983
     "\<exists>X \<in> Rep_star x.
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1984
            \<forall>u. 0 < u --> {n. abs (X n) < u} \<in> FreeUltrafilterNat
20541
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
  1985
      ==> (x::hypreal) \<in> Infinitesimal"
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1986
apply (simp add: Infinitesimal_def)
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17298
diff changeset
  1987
apply (cases x)
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1988
apply (auto simp add: abs_less_iff abs_interval_iff minus_less_iff [of x])
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1989
apply (auto simp add: SReal_iff)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1990
apply (drule_tac [!] x=y in spec) 
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17298
diff changeset
  1991
apply (auto simp add: star_n_less star_n_minus star_of_def, ultra+)
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1992
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1993
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  1994
lemma Infinitesimal_FreeUltrafilterNat_iff:
20541
f614c619b1e1 generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents: 20485
diff changeset
  1995
     "((x::hypreal) \<in> Infinitesimal) = (\<exists>X \<in> Rep_star x.
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1996
           \<forall>u. 0 < u --> {n. abs (X n) < u} \<in> FreeUltrafilterNat)"
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17298
diff changeset
  1997
by (blast intro!: Infinitesimal_FreeUltrafilterNat FreeUltrafilterNat_Infinitesimal)
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1998
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  1999
(*------------------------------------------------------------------------
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  2000
         Infinitesimals as smaller than 1/n for all n::nat (> 0)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  2001
 ------------------------------------------------------------------------*)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  2002
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  2003
lemma lemma_Infinitesimal:
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  2004
     "(\<forall>r. 0 < r --> x < r) = (\<forall>n. x < inverse(real (Suc n)))"
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  2005
apply (auto simp add: real_of_nat_Suc_gt_zero)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  2006
apply (blast dest!: reals_Archimedean intro: order_less_trans)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  2007
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  2008
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14371
diff changeset
  2009
lemma of_nat_in_Reals [simp]: "(of_nat n::hypreal) \<in> \<real>"
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14371
diff changeset
  2010
apply (induct n)
15539
333a88244569 comprehensive cleanup, replacing sumr by setsum
nipkow
parents: 15531
diff changeset
  2011
apply (simp_all)
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14371
diff changeset
  2012
done 
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14371
diff changeset
  2013
 
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  2014
lemma lemma_Infinitesimal2:
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  2015
     "(\<forall>r \<in> Reals. 0 < r --> x < r) =
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  2016
      (\<forall>n. x < inverse(hypreal_of_nat (Suc n)))"
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  2017
apply safe
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  2018
apply (drule_tac x = "inverse (hypreal_of_real (real (Suc n))) " in bspec)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  2019
apply (simp (no_asm_use) add: SReal_inverse)
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17298
diff changeset
  2020
apply (rule real_of_nat_Suc_gt_zero [THEN positive_imp_inverse_positive, THEN star_of_less [THEN iffD2], THEN [2] impE])
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  2021
prefer 2 apply assumption
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14371
diff changeset
  2022
apply (simp add: real_of_nat_Suc_gt_zero hypreal_of_nat_eq)
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  2023
apply (auto dest!: reals_Archimedean simp add: SReal_iff)
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17298
diff changeset
  2024
apply (drule star_of_less [THEN iffD2])
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14371
diff changeset
  2025
apply (simp add: real_of_nat_Suc_gt_zero hypreal_of_nat_eq)
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  2026
apply (blast intro: order_less_trans)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  2027
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  2028
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14371
diff changeset
  2029
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  2030
lemma Infinitesimal_hypreal_of_nat_iff:
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  2031
     "Infinitesimal = {x. \<forall>n. abs x < inverse (hypreal_of_nat (Suc n))}"
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  2032
apply (simp add: Infinitesimal_def)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  2033
apply (auto simp add: lemma_Infinitesimal2)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  2034
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  2035
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  2036
15229
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15140
diff changeset
  2037
subsection{*Proof that @{term omega} is an infinite number*}
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15140
diff changeset
  2038
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15140
diff changeset
  2039
text{*It will follow that epsilon is an infinitesimal number.*}
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15140
diff changeset
  2040
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  2041
lemma Suc_Un_eq: "{n. n < Suc m} = {n. n < m} Un {n. n = m}"
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  2042
by (auto simp add: less_Suc_eq)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  2043
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  2044
(*-------------------------------------------
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  2045
  Prove that any segment is finite and
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  2046
  hence cannot belong to FreeUltrafilterNat
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  2047
 -------------------------------------------*)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  2048
lemma finite_nat_segment: "finite {n::nat. n < m}"
15251
bb6f072c8d10 converted some induct_tac to induct
paulson
parents: 15234
diff changeset
  2049
apply (induct "m")
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  2050
apply (auto simp add: Suc_Un_eq)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  2051
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  2052
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  2053
lemma finite_real_of_nat_segment: "finite {n::nat. real n < real (m::nat)}"
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  2054
by (auto intro: finite_nat_segment)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  2055
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  2056
lemma finite_real_of_nat_less_real: "finite {n::nat. real n < u}"
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  2057
apply (cut_tac x = u in reals_Archimedean2, safe)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  2058
apply (rule finite_real_of_nat_segment [THEN [2] finite_subset])
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  2059
apply (auto dest: order_less_trans)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  2060
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  2061
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  2062
lemma lemma_real_le_Un_eq:
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  2063
     "{n. f n \<le> u} = {n. f n < u} Un {n. u = (f n :: real)}"
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  2064
by (auto dest: order_le_imp_less_or_eq simp add: order_less_imp_le)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  2065
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  2066
lemma finite_real_of_nat_le_real: "finite {n::nat. real n \<le> u}"
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  2067
by (auto simp add: lemma_real_le_Un_eq lemma_finite_omega_set finite_real_of_nat_less_real)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  2068
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  2069
lemma finite_rabs_real_of_nat_le_real: "finite {n::nat. abs(real n) \<le> u}"
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  2070
apply (simp (no_asm) add: real_of_nat_Suc_gt_zero finite_real_of_nat_le_real)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  2071
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  2072
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  2073
lemma rabs_real_of_nat_le_real_FreeUltrafilterNat:
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  2074
     "{n. abs(real n) \<le> u} \<notin> FreeUltrafilterNat"
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  2075
by (blast intro!: FreeUltrafilterNat_finite finite_rabs_real_of_nat_le_real)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  2076
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  2077
lemma FreeUltrafilterNat_nat_gt_real: "{n. u < real n} \<in> FreeUltrafilterNat"
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  2078
apply (rule ccontr, drule FreeUltrafilterNat_Compl_mem)
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  2079
apply (subgoal_tac "- {n::nat. u < real n} = {n. real n \<le> u}")
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  2080
prefer 2 apply force
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  2081
apply (simp add: finite_real_of_nat_le_real [THEN FreeUltrafilterNat_finite])
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  2082
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  2083
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  2084
(*--------------------------------------------------------------
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  2085
 The complement of {n. abs(real n) \<le> u} =
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  2086
 {n. u < abs (real n)} is in FreeUltrafilterNat
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  2087
 by property of (free) ultrafilters
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  2088
 --------------------------------------------------------------*)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  2089
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  2090
lemma Compl_real_le_eq: "- {n::nat. real n \<le> u} = {n. u < real n}"
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  2091
by (auto dest!: order_le_less_trans simp add: linorder_not_le)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  2092
15229
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15140
diff changeset
  2093
text{*@{term omega} is a member of @{term HInfinite}*}
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  2094
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  2095
lemma FreeUltrafilterNat_omega: "{n. u < real n} \<in> FreeUltrafilterNat"
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  2096
apply (cut_tac u = u in rabs_real_of_nat_le_real_FreeUltrafilterNat)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  2097
apply (auto dest: FreeUltrafilterNat_Compl_mem simp add: Compl_real_le_eq)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  2098
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  2099
15229
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15140
diff changeset
  2100
theorem HInfinite_omega [simp]: "omega \<in> HInfinite"
17298
ad73fb6144cf replace type hypreal with real star
huffman
parents: 16924
diff changeset
  2101
apply (simp add: omega_def star_n_def)
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  2102
apply (auto intro!: FreeUltrafilterNat_HInfinite)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  2103
apply (rule bexI)
17298
ad73fb6144cf replace type hypreal with real star
huffman
parents: 16924
diff changeset
  2104
apply (rule_tac [2] lemma_starrel_refl, auto)
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  2105
apply (simp (no_asm) add: real_of_nat_Suc diff_less_eq [symmetric] FreeUltrafilterNat_omega)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  2106
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  2107
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  2108
(*-----------------------------------------------
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  2109
       Epsilon is a member of Infinitesimal
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  2110
 -----------------------------------------------*)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  2111
15229
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15140
diff changeset
  2112
lemma Infinitesimal_epsilon [simp]: "epsilon \<in> Infinitesimal"
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  2113
by (auto intro!: HInfinite_inverse_Infinitesimal HInfinite_omega simp add: hypreal_epsilon_inverse_omega)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  2114
15229
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15140
diff changeset
  2115
lemma HFinite_epsilon [simp]: "epsilon \<in> HFinite"
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  2116
by (auto intro: Infinitesimal_subset_HFinite [THEN subsetD])
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  2117
15229
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15140
diff changeset
  2118
lemma epsilon_approx_zero [simp]: "epsilon @= 0"
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  2119
apply (simp (no_asm) add: mem_infmal_iff [symmetric])
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  2120
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  2121
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  2122
(*------------------------------------------------------------------------
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  2123
  Needed for proof that we define a hyperreal [<X(n)] @= hypreal_of_real a given
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  2124
  that \<forall>n. |X n - a| < 1/n. Used in proof of NSLIM => LIM.
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  2125
 -----------------------------------------------------------------------*)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  2126
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  2127
lemma real_of_nat_less_inverse_iff:
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  2128
     "0 < u  ==> (u < inverse (real(Suc n))) = (real(Suc n) < inverse u)"
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  2129
apply (simp add: inverse_eq_divide)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  2130
apply (subst pos_less_divide_eq, assumption)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  2131
apply (subst pos_less_divide_eq)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  2132
 apply (simp add: real_of_nat_Suc_gt_zero)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  2133
apply (simp add: real_mult_commute)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  2134
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  2135
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  2136
lemma finite_inverse_real_of_posnat_gt_real:
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  2137
     "0 < u ==> finite {n. u < inverse(real(Suc n))}"
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  2138
apply (simp (no_asm_simp) add: real_of_nat_less_inverse_iff)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  2139
apply (simp (no_asm_simp) add: real_of_nat_Suc less_diff_eq [symmetric])
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  2140
apply (rule finite_real_of_nat_less_real)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  2141
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  2142
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  2143
lemma lemma_real_le_Un_eq2:
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  2144
     "{n. u \<le> inverse(real(Suc n))} =
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  2145
     {n. u < inverse(real(Suc n))} Un {n. u = inverse(real(Suc n))}"
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  2146
apply (auto dest: order_le_imp_less_or_eq simp add: order_less_imp_le)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  2147
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  2148
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  2149
lemma real_of_nat_inverse_le_iff:
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  2150
     "(inverse (real(Suc n)) \<le> r) = (1 \<le> r * real(Suc n))"
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  2151
apply (simp (no_asm) add: linorder_not_less [symmetric])
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  2152
apply (simp (no_asm) add: inverse_eq_divide)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  2153
apply (subst pos_less_divide_eq)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  2154
apply (simp (no_asm) add: real_of_nat_Suc_gt_zero)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  2155
apply (simp (no_asm) add: real_mult_commute)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  2156
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  2157
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  2158
lemma real_of_nat_inverse_eq_iff:
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  2159
     "(u = inverse (real(Suc n))) = (real(Suc n) = inverse u)"
15539
333a88244569 comprehensive cleanup, replacing sumr by setsum
nipkow
parents: 15531
diff changeset
  2160
by (auto simp add: real_of_nat_Suc_gt_zero real_not_refl2 [THEN not_sym])
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  2161
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  2162
lemma lemma_finite_omega_set2: "finite {n::nat. u = inverse(real(Suc n))}"
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  2163
apply (simp (no_asm_simp) add: real_of_nat_inverse_eq_iff)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  2164
apply (cut_tac x = "inverse u - 1" in lemma_finite_omega_set)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  2165
apply (simp add: real_of_nat_Suc diff_eq_eq [symmetric] eq_commute)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  2166
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  2167
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  2168
lemma finite_inverse_real_of_posnat_ge_real:
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  2169
     "0 < u ==> finite {n. u \<le> inverse(real(Suc n))}"
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  2170
by (auto simp add: lemma_real_le_Un_eq2 lemma_finite_omega_set2 finite_inverse_real_of_posnat_gt_real)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  2171
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  2172
lemma inverse_real_of_posnat_ge_real_FreeUltrafilterNat:
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  2173
     "0 < u ==> {n. u \<le> inverse(real(Suc n))} \<notin> FreeUltrafilterNat"
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  2174
by (blast intro!: FreeUltrafilterNat_finite finite_inverse_real_of_posnat_ge_real)
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  2175
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  2176
(*--------------------------------------------------------------
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  2177
    The complement of  {n. u \<le> inverse(real(Suc n))} =
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  2178
    {n. inverse(real(Suc n)) < u} is in FreeUltrafilterNat
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  2179
    by property of (free) ultrafilters
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  2180
 --------------------------------------------------------------*)
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  2181
lemma Compl_le_inverse_eq:
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  2182
     "- {n. u \<le> inverse(real(Suc n))} =
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  2183
      {n. inverse(real(Suc n)) < u}"
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  2184
apply (auto dest!: order_le_less_trans simp add: linorder_not_le)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  2185
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  2186
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  2187
lemma FreeUltrafilterNat_inverse_real_of_posnat:
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  2188
     "0 < u ==>
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  2189
      {n. inverse(real(Suc n)) < u} \<in> FreeUltrafilterNat"
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  2190
apply (cut_tac u = u in inverse_real_of_posnat_ge_real_FreeUltrafilterNat)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  2191
apply (auto dest: FreeUltrafilterNat_Compl_mem simp add: Compl_le_inverse_eq)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  2192
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  2193
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  2194
text{* Example where we get a hyperreal from a real sequence
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  2195
      for which a particular property holds. The theorem is
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  2196
      used in proofs about equivalence of nonstandard and
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  2197
      standard neighbourhoods. Also used for equivalence of
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  2198
      nonstandard ans standard definitions of pointwise
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  2199
      limit.*}
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  2200
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  2201
(*-----------------------------------------------------
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  2202
    |X(n) - x| < 1/n ==> [<X n>] - hypreal_of_real x| \<in> Infinitesimal
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  2203
 -----------------------------------------------------*)
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  2204
lemma real_seq_to_hypreal_Infinitesimal:
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  2205
     "\<forall>n. abs(X n + -x) < inverse(real(Suc n))
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17298
diff changeset
  2206
     ==> star_n X + -hypreal_of_real x \<in> Infinitesimal"
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17298
diff changeset
  2207
apply (auto intro!: bexI dest: FreeUltrafilterNat_inverse_real_of_posnat FreeUltrafilterNat_all FreeUltrafilterNat_Int intro: order_less_trans FreeUltrafilterNat_subset simp add: star_n_minus star_of_def star_n_add Infinitesimal_FreeUltrafilterNat_iff star_n_inverse)
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  2208
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  2209
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  2210
lemma real_seq_to_hypreal_approx:
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  2211
     "\<forall>n. abs(X n + -x) < inverse(real(Suc n))
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17298
diff changeset
  2212
      ==> star_n X @= hypreal_of_real x"
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  2213
apply (subst approx_minus_iff)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  2214
apply (rule mem_infmal_iff [THEN subst])
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  2215
apply (erule real_seq_to_hypreal_Infinitesimal)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  2216
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  2217
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  2218
lemma real_seq_to_hypreal_approx2:
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  2219
     "\<forall>n. abs(x + -X n) < inverse(real(Suc n))
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17298
diff changeset
  2220
               ==> star_n X @= hypreal_of_real x"
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  2221
apply (simp add: abs_minus_add_cancel real_seq_to_hypreal_approx)
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  2222
done
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  2223
14420
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  2224
lemma real_seq_to_hypreal_Infinitesimal2:
4e72cd222e0b converted Hyperreal/HTranscendental to Isar script
paulson
parents: 14387
diff changeset
  2225
     "\<forall>n. abs(X n + -Y n) < inverse(real(Suc n))
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17298
diff changeset
  2226
      ==> star_n X + -star_n Y \<in> Infinitesimal"
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  2227
by (auto intro!: bexI
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  2228
	 dest: FreeUltrafilterNat_inverse_real_of_posnat 
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  2229
	       FreeUltrafilterNat_all FreeUltrafilterNat_Int
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  2230
	 intro: order_less_trans FreeUltrafilterNat_subset 
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17298
diff changeset
  2231
	 simp add: Infinitesimal_FreeUltrafilterNat_iff star_n_minus 
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17298
diff changeset
  2232
                   star_n_add star_n_inverse)
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 12114
diff changeset
  2233
10751
a81ea5d3dd41 separation of HOL-Hyperreal from HOL-Real
paulson
parents:
diff changeset
  2234
end