src/HOL/Hyperreal/Star.thy
author paulson
Thu, 29 Jan 2004 16:51:17 +0100
changeset 14370 b0064703967b
parent 10834 a7897aebbffc
child 14371 c78c7da09519
permissions -rw-r--r--
simplifications in the hyperreals
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       : Star.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
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
     4
    Description : defining *-transforms in NSA which extends sets of reals,
10751
a81ea5d3dd41 separation of HOL-Hyperreal from HOL-Real
paulson
parents:
diff changeset
     5
                  and real=>real functions
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
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: 10834
diff changeset
     8
header{*Star-Transforms in Non-Standard Analysis*}
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
     9
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
    10
theory Star = NSA:
10751
a81ea5d3dd41 separation of HOL-Hyperreal from HOL-Real
paulson
parents:
diff changeset
    11
a81ea5d3dd41 separation of HOL-Hyperreal from HOL-Real
paulson
parents:
diff changeset
    12
constdefs
a81ea5d3dd41 separation of HOL-Hyperreal from HOL-Real
paulson
parents:
diff changeset
    13
    (* nonstandard extension of sets *)
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
    14
    starset :: "real set => hypreal set"          ("*s* _" [80] 80)
10751
a81ea5d3dd41 separation of HOL-Hyperreal from HOL-Real
paulson
parents:
diff changeset
    15
    "*s* A  == {x. ALL X: Rep_hypreal(x). {n::nat. X n : A}: FreeUltrafilterNat}"
a81ea5d3dd41 separation of HOL-Hyperreal from HOL-Real
paulson
parents:
diff changeset
    16
a81ea5d3dd41 separation of HOL-Hyperreal from HOL-Real
paulson
parents:
diff changeset
    17
    (* internal sets *)
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
    18
    starset_n :: "(nat => real set) => hypreal set"        ("*sn* _" [80] 80)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
    19
    "*sn* As  == {x. ALL X: Rep_hypreal(x). {n::nat. X n : (As n)}: FreeUltrafilterNat}"
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
    20
10751
a81ea5d3dd41 separation of HOL-Hyperreal from HOL-Real
paulson
parents:
diff changeset
    21
    InternalSets :: "hypreal set set"
a81ea5d3dd41 separation of HOL-Hyperreal from HOL-Real
paulson
parents:
diff changeset
    22
    "InternalSets == {X. EX As. X = *sn* As}"
a81ea5d3dd41 separation of HOL-Hyperreal from HOL-Real
paulson
parents:
diff changeset
    23
a81ea5d3dd41 separation of HOL-Hyperreal from HOL-Real
paulson
parents:
diff changeset
    24
    (* nonstandard extension of function *)
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
    25
    is_starext  :: "[hypreal => hypreal, real => real] => bool"
10751
a81ea5d3dd41 separation of HOL-Hyperreal from HOL-Real
paulson
parents:
diff changeset
    26
    "is_starext F f == (ALL x y. EX X: Rep_hypreal(x). EX Y: Rep_hypreal(y).
a81ea5d3dd41 separation of HOL-Hyperreal from HOL-Real
paulson
parents:
diff changeset
    27
                        ((y = (F x)) = ({n. Y n = f(X n)} : FreeUltrafilterNat)))"
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
    28
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
    29
    starfun :: "(real => real) => hypreal => hypreal"       ("*f* _" [80] 80)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
    30
    "*f* f  == (%x. Abs_hypreal(UN X: Rep_hypreal(x). hyprel``{%n. f(X n)}))"
10751
a81ea5d3dd41 separation of HOL-Hyperreal from HOL-Real
paulson
parents:
diff changeset
    31
a81ea5d3dd41 separation of HOL-Hyperreal from HOL-Real
paulson
parents:
diff changeset
    32
    (* internal functions *)
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
    33
    starfun_n :: "(nat => (real => real)) => hypreal => hypreal"
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
    34
                 ("*fn* _" [80] 80)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
    35
    "*fn* F  == (%x. Abs_hypreal(UN X: Rep_hypreal(x). hyprel``{%n. (F n)(X n)}))"
10751
a81ea5d3dd41 separation of HOL-Hyperreal from HOL-Real
paulson
parents:
diff changeset
    36
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
    37
    InternalFuns :: "(hypreal => hypreal) set"
10751
a81ea5d3dd41 separation of HOL-Hyperreal from HOL-Real
paulson
parents:
diff changeset
    38
    "InternalFuns == {X. EX F. X = *fn* F}"
a81ea5d3dd41 separation of HOL-Hyperreal from HOL-Real
paulson
parents:
diff changeset
    39
a81ea5d3dd41 separation of HOL-Hyperreal from HOL-Real
paulson
parents:
diff changeset
    40
a81ea5d3dd41 separation of HOL-Hyperreal from HOL-Real
paulson
parents:
diff changeset
    41
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
    42
(*--------------------------------------------------------
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
    43
   Preamble - Pulling "EX" over "ALL"
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
    44
 ---------------------------------------------------------*)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
    45
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
    46
(* This proof does not need AC and was suggested by the
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
    47
   referee for the JCM Paper: let f(x) be least y such
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
    48
   that  Q(x,y)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
    49
*)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
    50
lemma no_choice: "ALL x. EX y. Q x y ==> EX (f :: nat => nat). ALL x. Q x (f x)"
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
    51
apply (rule_tac x = "%x. LEAST y. Q x y" in exI)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
    52
apply (blast intro: LeastI)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
    53
done
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
    54
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
    55
(*------------------------------------------------------------
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
    56
    Properties of the *-transform applied to sets of reals
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
    57
 ------------------------------------------------------------*)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
    58
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
    59
lemma STAR_real_set: "*s*(UNIV::real set) = (UNIV::hypreal set)"
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
    60
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
    61
apply (unfold starset_def)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
    62
apply auto
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
    63
done
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
    64
declare STAR_real_set [simp]
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
    65
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
    66
lemma STAR_empty_set: "*s* {} = {}"
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
    67
apply (unfold starset_def)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
    68
apply safe
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
    69
apply (rule_tac z = "x" in eq_Abs_hypreal)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
    70
apply (drule_tac x = "%n. xa n" in bspec)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
    71
apply auto
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
    72
done
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
    73
declare STAR_empty_set [simp]
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
    74
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
    75
lemma STAR_Un: "*s* (A Un B) = *s* A Un *s* B"
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
    76
apply (unfold starset_def)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
    77
apply auto
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
    78
  prefer 3 apply (blast intro: FreeUltrafilterNat_subset)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
    79
 prefer 2 apply (blast intro: FreeUltrafilterNat_subset)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
    80
apply (drule FreeUltrafilterNat_Compl_mem)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
    81
apply (drule bspec , assumption)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
    82
apply (rule_tac z = "x" in eq_Abs_hypreal)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
    83
apply auto
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
    84
apply ultra
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
    85
done
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
    86
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
    87
lemma STAR_Int: "*s* (A Int B) = *s* A Int *s* B"
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
    88
apply (unfold starset_def)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
    89
apply auto
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
    90
prefer 3 apply (blast intro: FreeUltrafilterNat_Int FreeUltrafilterNat_subset)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
    91
apply (blast intro: FreeUltrafilterNat_subset)+
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
    92
done
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
    93
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
    94
lemma STAR_Compl: "*s* -A = -( *s* A)"
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
    95
apply (auto simp add: starset_def)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
    96
apply (rule_tac [!] z = "x" in eq_Abs_hypreal)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
    97
apply (auto dest!: bspec);
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
    98
apply ultra
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
    99
apply (drule FreeUltrafilterNat_Compl_mem)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   100
apply ultra
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   101
done
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   102
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   103
lemma STAR_mem_Compl: "x \<notin> *s* F ==> x : *s* (- F)"
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   104
apply (auto simp add: STAR_Compl)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   105
done
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   106
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   107
lemma STAR_diff: "*s* (A - B) = *s* A - *s* B"
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   108
apply (auto simp add: Diff_eq STAR_Int STAR_Compl)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   109
done
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   110
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   111
lemma STAR_subset: "A <= B ==> *s* A <= *s* B"
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   112
apply (unfold starset_def)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   113
apply (blast intro: FreeUltrafilterNat_subset)+
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   114
done
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   115
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   116
lemma STAR_mem: "a : A ==> hypreal_of_real a : *s* A"
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   117
apply (unfold starset_def hypreal_of_real_def)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   118
apply (auto intro: FreeUltrafilterNat_subset)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   119
done
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   120
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   121
lemma STAR_hypreal_of_real_image_subset: "hypreal_of_real ` A <= *s* A"
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   122
apply (unfold starset_def)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   123
apply (auto simp add: hypreal_of_real_def)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   124
apply (blast intro: FreeUltrafilterNat_subset)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   125
done
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   126
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   127
lemma STAR_hypreal_of_real_Int: "*s* X Int Reals = hypreal_of_real ` X"
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   128
apply (unfold starset_def)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   129
apply (auto simp add: hypreal_of_real_def SReal_def)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   130
apply (simp add: hypreal_of_real_def [symmetric])
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   131
apply (rule imageI , rule ccontr)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   132
apply (drule bspec)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   133
apply (rule lemma_hyprel_refl)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   134
prefer 2 apply (blast intro: FreeUltrafilterNat_subset)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   135
apply auto
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   136
done
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   137
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   138
lemma lemma_not_hyprealA: "x \<notin> hypreal_of_real ` A ==> ALL y: A. x \<noteq> hypreal_of_real y"
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   139
apply auto
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   140
done
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   141
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   142
lemma lemma_Compl_eq: "- {n. X n = xa} = {n. X n \<noteq> xa}"
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   143
apply auto
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   144
done
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   145
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   146
lemma STAR_real_seq_to_hypreal:
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   147
    "ALL n. (X n) \<notin> M
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   148
          ==> Abs_hypreal(hyprel``{X}) \<notin> *s* M"
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   149
apply (unfold starset_def)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   150
apply (auto , rule bexI , rule_tac [2] lemma_hyprel_refl)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   151
apply auto
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   152
done
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   153
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   154
lemma STAR_singleton: "*s* {x} = {hypreal_of_real x}"
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   155
apply (unfold starset_def)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   156
apply (auto simp add: hypreal_of_real_def)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   157
apply (rule_tac z = "xa" in eq_Abs_hypreal)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   158
apply (auto intro: FreeUltrafilterNat_subset)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   159
done
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   160
declare STAR_singleton [simp]
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   161
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   162
lemma STAR_not_mem: "x \<notin> F ==> hypreal_of_real x \<notin> *s* F"
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   163
apply (auto simp add: starset_def hypreal_of_real_def)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   164
apply (rule bexI , rule_tac [2] lemma_hyprel_refl)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   165
apply auto
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   166
done
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   167
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   168
lemma STAR_subset_closed: "[| x : *s* A; A <= B |] ==> x : *s* B"
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   169
apply (blast dest: STAR_subset)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   170
done
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   171
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   172
(*------------------------------------------------------------------
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   173
   Nonstandard extension of a set (defined using a constant
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   174
   sequence) as a special case of an internal set
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   175
 -----------------------------------------------------------------*)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   176
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   177
lemma starset_n_starset:
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   178
     "ALL n. (As n = A) ==> *sn* As = *s* A"
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   179
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   180
apply (unfold starset_n_def starset_def)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   181
apply auto
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   182
done
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   183
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   184
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   185
(*----------------------------------------------------------------*)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   186
(* Theorems about nonstandard extensions of functions             *)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   187
(*----------------------------------------------------------------*)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   188
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   189
(*----------------------------------------------------------------*)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   190
(* Nonstandard extension of a function (defined using a           *)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   191
(* constant sequence) as a special case of an internal function   *)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   192
(*----------------------------------------------------------------*)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   193
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   194
lemma starfun_n_starfun:
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   195
     "ALL n. (F n = f) ==> *fn* F = *f* f"
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   196
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   197
apply (unfold starfun_n_def starfun_def)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   198
apply auto
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   199
done
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   200
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   201
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   202
(*
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   203
   Prove that hrabs is a nonstandard extension of rabs without
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   204
   use of congruence property (proved after this for general
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   205
   nonstandard extensions of real valued functions). This makes
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   206
   proof much longer- see comments at end of HREALABS.thy where
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   207
   we proved a congruence theorem for hrabs.
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   208
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   209
   NEW!!! No need to prove all the lemmas anymore. Use the ultrafilter
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   210
   tactic!
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   211
*)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   212
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   213
lemma hrabs_is_starext_rabs: "is_starext abs abs"
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   214
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   215
apply (unfold is_starext_def)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   216
apply safe
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   217
apply (rule_tac z = "x" in eq_Abs_hypreal)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   218
apply (rule_tac z = "y" in eq_Abs_hypreal)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   219
apply auto
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   220
apply (rule bexI , rule_tac [2] lemma_hyprel_refl)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   221
apply (rule bexI , rule_tac [2] lemma_hyprel_refl)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   222
apply (auto dest!: spec simp add: hypreal_minus hrabs_def hypreal_zero_def hypreal_le_def hypreal_less_def)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   223
apply (arith | ultra)+
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   224
done
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   225
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   226
lemma Rep_hypreal_FreeUltrafilterNat: "[| X: Rep_hypreal z; Y: Rep_hypreal z |]
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   227
      ==> {n. X n = Y n} : FreeUltrafilterNat"
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   228
apply (rule_tac z = "z" in eq_Abs_hypreal)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   229
apply (auto , ultra)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   230
done
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   231
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   232
(*-----------------------------------------------------------------------
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   233
    Nonstandard extension of functions- congruence
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   234
 -----------------------------------------------------------------------*)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   235
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   236
lemma starfun_congruent: "congruent hyprel (%X. hyprel``{%n. f (X n)})"
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   237
apply (unfold congruent_def)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   238
apply auto
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   239
apply ultra
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   240
done
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   241
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   242
lemma starfun:
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   243
      "( *f* f) (Abs_hypreal(hyprel``{%n. X n})) =
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   244
       Abs_hypreal(hyprel `` {%n. f (X n)})"
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   245
apply (unfold starfun_def)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   246
apply (rule_tac f = "Abs_hypreal" in arg_cong)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   247
apply (simp add: hyprel_in_hypreal [THEN Abs_hypreal_inverse] 
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   248
                 UN_equiv_class [OF equiv_hyprel starfun_congruent])
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   249
done
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   250
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   251
(*-------------------------------------------
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   252
  multiplication: ( *f ) x ( *g ) = *(f x g)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   253
 ------------------------------------------*)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   254
lemma starfun_mult: "( *f* f) xa * ( *f* g) xa = ( *f* (%x. f x * g x)) xa"
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   255
apply (rule_tac z = "xa" in eq_Abs_hypreal)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   256
apply (auto simp add: starfun hypreal_mult)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   257
done
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   258
declare starfun_mult [symmetric, simp]
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   259
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   260
(*---------------------------------------
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   261
  addition: ( *f ) + ( *g ) = *(f + g)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   262
 ---------------------------------------*)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   263
lemma starfun_add: "( *f* f) xa + ( *f* g) xa = ( *f* (%x. f x + g x)) xa"
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   264
apply (rule_tac z = "xa" in eq_Abs_hypreal)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   265
apply (auto simp add: starfun hypreal_add)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   266
done
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   267
declare starfun_add [symmetric, simp]
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   268
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   269
(*--------------------------------------------
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   270
  subtraction: ( *f ) + -( *g ) = *(f + -g)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   271
 -------------------------------------------*)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   272
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   273
lemma starfun_minus: "- ( *f* f) x = ( *f* (%x. - f x)) x"
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   274
apply (rule_tac z = "x" in eq_Abs_hypreal)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   275
apply (auto simp add: starfun hypreal_minus)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   276
done
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   277
declare starfun_minus [symmetric, simp]
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   278
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   279
(*FIXME: delete*)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   280
lemma starfun_add_minus: "( *f* f) xa + -( *f* g) xa = ( *f* (%x. f x + -g x)) xa"
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   281
apply (simp (no_asm))
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   282
done
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   283
declare starfun_add_minus [symmetric, simp]
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   284
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   285
lemma starfun_diff:
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   286
  "( *f* f) xa  - ( *f* g) xa = ( *f* (%x. f x - g x)) xa"
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   287
apply (unfold hypreal_diff_def real_diff_def)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   288
apply (rule starfun_add_minus)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   289
done
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   290
declare starfun_diff [symmetric, simp]
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   291
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   292
(*--------------------------------------
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   293
  composition: ( *f ) o ( *g ) = *(f o g)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   294
 ---------------------------------------*)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   295
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   296
lemma starfun_o2: "(%x. ( *f* f) (( *f* g) x)) = *f* (%x. f (g x))"
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   297
apply (rule ext)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   298
apply (rule_tac z = "x" in eq_Abs_hypreal)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   299
apply (auto simp add: starfun)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   300
done
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   301
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   302
lemma starfun_o: "( *f* f) o ( *f* g) = ( *f* (f o g))"
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   303
apply (unfold o_def)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   304
apply (simp (no_asm) add: starfun_o2)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   305
done
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   306
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   307
(*--------------------------------------
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   308
  NS extension of constant function
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   309
 --------------------------------------*)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   310
lemma starfun_const_fun: "( *f* (%x. k)) xa = hypreal_of_real  k"
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   311
apply (rule_tac z = "xa" in eq_Abs_hypreal)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   312
apply (auto simp add: starfun hypreal_of_real_def)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   313
done
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   314
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   315
declare starfun_const_fun [simp]
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   316
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   317
(*----------------------------------------------------
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   318
   the NS extension of the identity function
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   319
 ----------------------------------------------------*)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   320
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   321
lemma starfun_Idfun_approx: "x @= hypreal_of_real a ==> ( *f* (%x. x)) x @= hypreal_of_real  a"
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   322
apply (rule_tac z = "x" in eq_Abs_hypreal)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   323
apply (auto simp add: starfun)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   324
done
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   325
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   326
lemma starfun_Id: "( *f* (%x. x)) x = x"
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   327
apply (rule_tac z = "x" in eq_Abs_hypreal)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   328
apply (auto simp add: starfun)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   329
done
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   330
declare starfun_Id [simp]
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   331
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   332
(*----------------------------------------------------------------------
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   333
      the *-function is a (nonstandard) extension of the function
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   334
 ----------------------------------------------------------------------*)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   335
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   336
lemma is_starext_starfun: "is_starext ( *f* f) f"
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   337
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   338
apply (unfold is_starext_def)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   339
apply auto
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   340
apply (rule_tac z = "x" in eq_Abs_hypreal)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   341
apply (rule_tac z = "y" in eq_Abs_hypreal)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   342
apply (auto intro!: bexI simp add: starfun)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   343
done
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   344
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   345
(*----------------------------------------------------------------------
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   346
     Any nonstandard extension is in fact the *-function
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   347
 ----------------------------------------------------------------------*)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   348
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   349
lemma is_starfun_starext: "is_starext F f ==> F = *f* f"
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   350
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   351
apply (unfold is_starext_def)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   352
apply (rule ext)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   353
apply (rule_tac z = "x" in eq_Abs_hypreal)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   354
apply (drule_tac x = "x" in spec)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   355
apply (drule_tac x = "( *f* f) x" in spec)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   356
apply (auto dest!: FreeUltrafilterNat_Compl_mem simp add: starfun)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   357
apply ultra
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   358
done
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   359
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   360
lemma is_starext_starfun_iff: "(is_starext F f) = (F = *f* f)"
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   361
apply (blast intro: is_starfun_starext is_starext_starfun)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   362
done
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   363
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   364
(*--------------------------------------------------------
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   365
   extented function has same solution as its standard
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   366
   version for real arguments. i.e they are the same
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   367
   for all real arguments
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   368
 -------------------------------------------------------*)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   369
lemma starfun_eq: "( *f* f) (hypreal_of_real a) = hypreal_of_real (f a)"
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   370
apply (auto simp add: starfun hypreal_of_real_def)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   371
done
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   372
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   373
declare starfun_eq [simp]
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   374
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   375
lemma starfun_approx: "( *f* f) (hypreal_of_real a) @= hypreal_of_real (f a)"
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   376
apply auto
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   377
done
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   378
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   379
(* useful for NS definition of derivatives *)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   380
lemma starfun_lambda_cancel: "( *f* (%h. f (x + h))) xa  = ( *f* f) (hypreal_of_real  x + xa)"
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   381
apply (rule_tac z = "xa" in eq_Abs_hypreal)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   382
apply (auto simp add: starfun hypreal_of_real_def hypreal_add)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   383
done
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   384
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   385
lemma starfun_lambda_cancel2: "( *f* (%h. f(g(x + h)))) xa = ( *f* (f o g)) (hypreal_of_real x + xa)"
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   386
apply (rule_tac z = "xa" in eq_Abs_hypreal)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   387
apply (auto simp add: starfun hypreal_of_real_def hypreal_add)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   388
done
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   389
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   390
lemma starfun_mult_HFinite_approx: "[| ( *f* f) xa @= l; ( *f* g) xa @= m;
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   391
                  l: HFinite; m: HFinite
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   392
               |] ==>  ( *f* (%x. f x * g x)) xa @= l * m"
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   393
apply (drule approx_mult_HFinite)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   394
apply (assumption)+
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   395
apply (auto intro: approx_HFinite [OF _ approx_sym])
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   396
done
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   397
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   398
lemma starfun_add_approx: "[| ( *f* f) xa @= l; ( *f* g) xa @= m
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   399
               |] ==>  ( *f* (%x. f x + g x)) xa @= l + m"
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   400
apply (auto intro: approx_add)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   401
done
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   402
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   403
(*----------------------------------------------------
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   404
    Examples: hrabs is nonstandard extension of rabs
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   405
              inverse is nonstandard extension of inverse
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   406
 ---------------------------------------------------*)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   407
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   408
(* can be proved easily using theorem "starfun" and *)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   409
(* properties of ultrafilter as for inverse below we  *)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   410
(* use the theorem we proved above instead          *)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   411
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   412
lemma starfun_rabs_hrabs: "*f* abs = abs"
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   413
apply (rule hrabs_is_starext_rabs [THEN is_starext_starfun_iff [THEN iffD1], symmetric])
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   414
done
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   415
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   416
lemma starfun_inverse_inverse: "( *f* inverse) x = inverse(x)"
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   417
apply (rule_tac z = "x" in eq_Abs_hypreal)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   418
apply (auto simp add: starfun hypreal_inverse hypreal_zero_def)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   419
done
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   420
declare starfun_inverse_inverse [simp]
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   421
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   422
lemma starfun_inverse: "inverse (( *f* f) x) = ( *f* (%x. inverse (f x))) x"
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   423
apply (rule_tac z = "x" in eq_Abs_hypreal)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   424
apply (auto simp add: starfun hypreal_inverse)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   425
done
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   426
declare starfun_inverse [symmetric, simp]
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   427
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   428
lemma starfun_divide:
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   429
  "( *f* f) xa  / ( *f* g) xa = ( *f* (%x. f x / g x)) xa"
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   430
apply (unfold hypreal_divide_def real_divide_def)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   431
apply auto
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   432
done
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   433
declare starfun_divide [symmetric, simp]
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   434
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   435
lemma starfun_inverse2: "inverse (( *f* f) x) = ( *f* (%x. inverse (f x))) x"
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   436
apply (rule_tac z = "x" in eq_Abs_hypreal)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   437
apply (auto intro: FreeUltrafilterNat_subset dest!: FreeUltrafilterNat_Compl_mem simp add: starfun hypreal_inverse hypreal_zero_def)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   438
done
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   439
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   440
(*-------------------------------------------------------------
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   441
    General lemma/theorem needed for proofs in elementary
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   442
    topology of the reals
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   443
 ------------------------------------------------------------*)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   444
lemma starfun_mem_starset:
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   445
      "( *f* f) x : *s* A ==> x : *s* {x. f x : A}"
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   446
apply (unfold starset_def)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   447
apply (rule_tac z = "x" in eq_Abs_hypreal)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   448
apply (auto simp add: starfun)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   449
apply (rename_tac "X")
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   450
apply (drule_tac x = "%n. f (X n) " in bspec)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   451
apply (auto , ultra)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   452
done
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   453
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   454
(*------------------------------------------------------------
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   455
   Alternative definition for hrabs with rabs function
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   456
   applied entrywise to equivalence class representative.
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   457
   This is easily proved using starfun and ns extension thm
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   458
 ------------------------------------------------------------*)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   459
lemma hypreal_hrabs: "abs (Abs_hypreal (hyprel `` {X})) =
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   460
                  Abs_hypreal(hyprel `` {%n. abs (X n)})"
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   461
apply (simp (no_asm) add: starfun_rabs_hrabs [symmetric] starfun)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   462
done
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   463
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   464
(*----------------------------------------------------------------
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   465
   nonstandard extension of set through nonstandard extension
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   466
   of rabs function i.e hrabs. A more general result should be
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   467
   where we replace rabs by some arbitrary function f and hrabs
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   468
   by its NS extenson ( *f* f). See second NS set extension below.
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   469
 ----------------------------------------------------------------*)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   470
lemma STAR_rabs_add_minus:
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   471
   "*s* {x. abs (x + - y) < r} =
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   472
     {x. abs(x + -hypreal_of_real y) < hypreal_of_real r}"
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   473
apply (unfold starset_def)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   474
apply safe
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   475
apply (rule_tac [!] z = "x" in eq_Abs_hypreal)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   476
apply (auto intro!: exI dest!: bspec simp add: hypreal_minus hypreal_of_real_def hypreal_add hypreal_hrabs hypreal_less)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   477
apply ultra
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   478
done
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   479
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   480
lemma STAR_starfun_rabs_add_minus:
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   481
  "*s* {x. abs (f x + - y) < r} =
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   482
       {x. abs(( *f* f) x + -hypreal_of_real y) < hypreal_of_real r}"
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   483
apply (unfold starset_def)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   484
apply safe
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   485
apply (rule_tac [!] z = "x" in eq_Abs_hypreal)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   486
apply (auto intro!: exI dest!: bspec simp add: hypreal_minus hypreal_of_real_def hypreal_add hypreal_hrabs hypreal_less starfun)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   487
apply ultra
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   488
done
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   489
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   490
(*-------------------------------------------------------------------
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   491
   Another characterization of Infinitesimal and one of @= relation.
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   492
   In this theory since hypreal_hrabs proved here. (To Check:) Maybe
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   493
   move both if possible?
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   494
 -------------------------------------------------------------------*)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   495
lemma Infinitesimal_FreeUltrafilterNat_iff2: "(x:Infinitesimal) =
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   496
      (EX X:Rep_hypreal(x).
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   497
        ALL m. {n. abs(X n) < inverse(real(Suc m))}
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   498
               : FreeUltrafilterNat)"
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   499
apply (rule_tac z = "x" in eq_Abs_hypreal)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   500
apply (auto intro!: bexI lemma_hyprel_refl 
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   501
            simp add: Infinitesimal_hypreal_of_nat_iff hypreal_of_real_def 
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   502
     hypreal_inverse hypreal_hrabs hypreal_less hypreal_of_nat_def)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   503
apply (drule_tac x = "n" in spec)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   504
apply ultra
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   505
done
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   506
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   507
lemma approx_FreeUltrafilterNat_iff: "(Abs_hypreal(hyprel``{X}) @= Abs_hypreal(hyprel``{Y})) =
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   508
      (ALL m. {n. abs (X n + - Y n) <
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   509
                  inverse(real(Suc m))} : FreeUltrafilterNat)"
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   510
apply (subst approx_minus_iff)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   511
apply (rule mem_infmal_iff [THEN subst])
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   512
apply (auto simp add: hypreal_minus hypreal_add Infinitesimal_FreeUltrafilterNat_iff2)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   513
apply (drule_tac x = "m" in spec)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   514
apply ultra
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   515
done
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   516
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   517
lemma inj_starfun: "inj starfun"
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   518
apply (rule inj_onI)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   519
apply (rule ext , rule ccontr)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   520
apply (drule_tac x = "Abs_hypreal (hyprel ``{%n. xa}) " in fun_cong)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   521
apply (auto simp add: starfun)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   522
done
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   523
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   524
ML
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   525
{*
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   526
val starset_def = thm"starset_def";
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   527
val starset_n_def = thm"starset_n_def";
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   528
val InternalSets_def = thm"InternalSets_def";
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   529
val is_starext_def = thm"is_starext_def";
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   530
val starfun_def = thm"starfun_def";
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   531
val starfun_n_def = thm"starfun_n_def";
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   532
val InternalFuns_def = thm"InternalFuns_def";
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   533
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   534
val no_choice = thm "no_choice";
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   535
val STAR_real_set = thm "STAR_real_set";
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   536
val STAR_empty_set = thm "STAR_empty_set";
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   537
val STAR_Un = thm "STAR_Un";
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   538
val STAR_Int = thm "STAR_Int";
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   539
val STAR_Compl = thm "STAR_Compl";
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   540
val STAR_mem_Compl = thm "STAR_mem_Compl";
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   541
val STAR_diff = thm "STAR_diff";
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   542
val STAR_subset = thm "STAR_subset";
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   543
val STAR_mem = thm "STAR_mem";
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   544
val STAR_hypreal_of_real_image_subset = thm "STAR_hypreal_of_real_image_subset";
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   545
val STAR_hypreal_of_real_Int = thm "STAR_hypreal_of_real_Int";
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   546
val STAR_real_seq_to_hypreal = thm "STAR_real_seq_to_hypreal";
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   547
val STAR_singleton = thm "STAR_singleton";
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   548
val STAR_not_mem = thm "STAR_not_mem";
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   549
val STAR_subset_closed = thm "STAR_subset_closed";
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   550
val starset_n_starset = thm "starset_n_starset";
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   551
val starfun_n_starfun = thm "starfun_n_starfun";
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   552
val hrabs_is_starext_rabs = thm "hrabs_is_starext_rabs";
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   553
val Rep_hypreal_FreeUltrafilterNat = thm "Rep_hypreal_FreeUltrafilterNat";
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   554
val starfun_congruent = thm "starfun_congruent";
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   555
val starfun = thm "starfun";
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   556
val starfun_mult = thm "starfun_mult";
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   557
val starfun_add = thm "starfun_add";
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   558
val starfun_minus = thm "starfun_minus";
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   559
val starfun_add_minus = thm "starfun_add_minus";
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   560
val starfun_diff = thm "starfun_diff";
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   561
val starfun_o2 = thm "starfun_o2";
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   562
val starfun_o = thm "starfun_o";
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   563
val starfun_const_fun = thm "starfun_const_fun";
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   564
val starfun_Idfun_approx = thm "starfun_Idfun_approx";
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   565
val starfun_Id = thm "starfun_Id";
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   566
val is_starext_starfun = thm "is_starext_starfun";
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   567
val is_starfun_starext = thm "is_starfun_starext";
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   568
val is_starext_starfun_iff = thm "is_starext_starfun_iff";
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   569
val starfun_eq = thm "starfun_eq";
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   570
val starfun_approx = thm "starfun_approx";
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   571
val starfun_lambda_cancel = thm "starfun_lambda_cancel";
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   572
val starfun_lambda_cancel2 = thm "starfun_lambda_cancel2";
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   573
val starfun_mult_HFinite_approx = thm "starfun_mult_HFinite_approx";
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   574
val starfun_add_approx = thm "starfun_add_approx";
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   575
val starfun_rabs_hrabs = thm "starfun_rabs_hrabs";
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   576
val starfun_inverse_inverse = thm "starfun_inverse_inverse";
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   577
val starfun_inverse = thm "starfun_inverse";
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   578
val starfun_divide = thm "starfun_divide";
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   579
val starfun_inverse2 = thm "starfun_inverse2";
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   580
val starfun_mem_starset = thm "starfun_mem_starset";
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   581
val hypreal_hrabs = thm "hypreal_hrabs";
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   582
val STAR_rabs_add_minus = thm "STAR_rabs_add_minus";
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   583
val STAR_starfun_rabs_add_minus = thm "STAR_starfun_rabs_add_minus";
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   584
val Infinitesimal_FreeUltrafilterNat_iff2 = thm "Infinitesimal_FreeUltrafilterNat_iff2";
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   585
val approx_FreeUltrafilterNat_iff = thm "approx_FreeUltrafilterNat_iff";
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   586
val inj_starfun = thm "inj_starfun";
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   587
*}
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   588
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   589
end