src/HOL/Hyperreal/Star.thy
author urbanc
Tue, 13 Dec 2005 18:11:21 +0100
changeset 18396 b3e7da94b51f
parent 17429 e8d6ed3aacfe
child 19765 dfe940911617
permissions -rw-r--r--
added a fresh_left lemma that contains all instantiation for the various atom-types.
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
14468
6be497cacab5 heavy tidying
paulson
parents: 14378
diff changeset
     4
    Conversion to Isar and new proofs by Lawrence C Paulson, 2003/4
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
     5
*)
10751
a81ea5d3dd41 separation of HOL-Hyperreal from HOL-Real
paulson
parents:
diff changeset
     6
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
     7
header{*Star-Transforms in Non-Standard Analysis*}
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
     8
15131
c69542757a4d New theory header syntax.
nipkow
parents: 15003
diff changeset
     9
theory Star
15140
322485b816ac import -> imports
nipkow
parents: 15131
diff changeset
    10
imports NSA
15131
c69542757a4d New theory header syntax.
nipkow
parents: 15003
diff changeset
    11
begin
10751
a81ea5d3dd41 separation of HOL-Hyperreal from HOL-Real
paulson
parents:
diff changeset
    12
a81ea5d3dd41 separation of HOL-Hyperreal from HOL-Real
paulson
parents:
diff changeset
    13
constdefs
a81ea5d3dd41 separation of HOL-Hyperreal from HOL-Real
paulson
parents:
diff changeset
    14
    (* internal sets *)
17302
25aab757672b generalized types
huffman
parents: 17298
diff changeset
    15
    starset_n :: "(nat => 'a set) => 'a star set"        ("*sn* _" [80] 80)
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17303
diff changeset
    16
    "*sn* As == Iset (star_n As)"
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
    17
17302
25aab757672b generalized types
huffman
parents: 17298
diff changeset
    18
    InternalSets :: "'a star set set"
14468
6be497cacab5 heavy tidying
paulson
parents: 14378
diff changeset
    19
    "InternalSets == {X. \<exists>As. X = *sn* As}"
10751
a81ea5d3dd41 separation of HOL-Hyperreal from HOL-Real
paulson
parents:
diff changeset
    20
a81ea5d3dd41 separation of HOL-Hyperreal from HOL-Real
paulson
parents:
diff changeset
    21
    (* nonstandard extension of function *)
17302
25aab757672b generalized types
huffman
parents: 17298
diff changeset
    22
    is_starext  :: "['a star => 'a star, 'a => 'a] => bool"
17298
ad73fb6144cf replace type hypreal with real star
huffman
parents: 15170
diff changeset
    23
    "is_starext F f == (\<forall>x y. \<exists>X \<in> Rep_star(x). \<exists>Y \<in> Rep_star(y).
10751
a81ea5d3dd41 separation of HOL-Hyperreal from HOL-Real
paulson
parents:
diff changeset
    24
                        ((y = (F x)) = ({n. Y n = f(X n)} : FreeUltrafilterNat)))"
a81ea5d3dd41 separation of HOL-Hyperreal from HOL-Real
paulson
parents:
diff changeset
    25
    (* internal functions *)
17303
560cf01f4772 generalized types more
huffman
parents: 17302
diff changeset
    26
    starfun_n :: "(nat => ('a => 'b)) => 'a star => 'b star"
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
    27
                 ("*fn* _" [80] 80)
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17303
diff changeset
    28
    "*fn* F == Ifun (star_n F)"
10751
a81ea5d3dd41 separation of HOL-Hyperreal from HOL-Real
paulson
parents:
diff changeset
    29
17303
560cf01f4772 generalized types more
huffman
parents: 17302
diff changeset
    30
    InternalFuns :: "('a star => 'b star) set"
14468
6be497cacab5 heavy tidying
paulson
parents: 14378
diff changeset
    31
    "InternalFuns == {X. \<exists>F. X = *fn* F}"
10751
a81ea5d3dd41 separation of HOL-Hyperreal from HOL-Real
paulson
parents:
diff changeset
    32
a81ea5d3dd41 separation of HOL-Hyperreal from HOL-Real
paulson
parents:
diff changeset
    33
a81ea5d3dd41 separation of HOL-Hyperreal from HOL-Real
paulson
parents:
diff changeset
    34
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
    35
(*--------------------------------------------------------
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
    36
   Preamble - Pulling "EX" over "ALL"
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
    37
 ---------------------------------------------------------*)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
    38
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
    39
(* This proof does not need AC and was suggested by the
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
    40
   referee for the JCM Paper: let f(x) be least y such
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
    41
   that  Q(x,y)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
    42
*)
14468
6be497cacab5 heavy tidying
paulson
parents: 14378
diff changeset
    43
lemma no_choice: "\<forall>x. \<exists>y. Q x y ==> \<exists>(f :: nat => nat). \<forall>x. Q x (f x)"
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
    44
apply (rule_tac x = "%x. LEAST y. Q x y" in exI)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
    45
apply (blast intro: LeastI)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
    46
done
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
    47
15169
2b5da07a0b89 new "respects" syntax for quotienting
paulson
parents: 15140
diff changeset
    48
subsection{*Properties of the Star-transform Applied to Sets of Reals*}
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
    49
17429
e8d6ed3aacfe merged Transfer.thy and StarType.thy into StarDef.thy; renamed Ifun2_of to starfun2; cleaned up
huffman
parents: 17318
diff changeset
    50
lemma STAR_UNIV_set: "*s*(UNIV::'a set) = (UNIV::'a star set)"
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17303
diff changeset
    51
by (transfer UNIV_def, rule refl)
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
    52
17429
e8d6ed3aacfe merged Transfer.thy and StarType.thy into StarDef.thy; renamed Ifun2_of to starfun2; cleaned up
huffman
parents: 17318
diff changeset
    53
lemma STAR_empty_set: "*s* {} = {}"
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17303
diff changeset
    54
by (transfer empty_def, rule refl)
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
    55
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
    56
lemma STAR_Un: "*s* (A Un B) = *s* A Un *s* B"
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17303
diff changeset
    57
by (transfer Un_def, rule refl)
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
    58
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
    59
lemma STAR_Int: "*s* (A Int B) = *s* A Int *s* B"
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17303
diff changeset
    60
by (transfer Int_def, rule refl)
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
    61
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
    62
lemma STAR_Compl: "*s* -A = -( *s* A)"
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17303
diff changeset
    63
by (transfer Compl_def, rule refl)
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
    64
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17303
diff changeset
    65
lemma STAR_mem_Compl: "!!x. x \<notin> *s* F ==> x : *s* (- F)"
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17303
diff changeset
    66
by (transfer Compl_def, simp)
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
    67
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
    68
lemma STAR_diff: "*s* (A - B) = *s* A - *s* B"
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17303
diff changeset
    69
by (transfer set_diff_def, rule refl)
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
    70
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
    71
lemma STAR_subset: "A <= B ==> *s* A <= *s* B"
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17303
diff changeset
    72
by (transfer subset_def, simp)
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17303
diff changeset
    73
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17303
diff changeset
    74
lemma STAR_mem: "a \<in> A ==> star_of a : *s* A"
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17303
diff changeset
    75
by transfer
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
    76
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17303
diff changeset
    77
lemma STAR_mem_iff: "(star_of x \<in> *s* A) = (x \<in> A)"
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17303
diff changeset
    78
by (transfer, rule refl)
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
    79
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17303
diff changeset
    80
lemma STAR_star_of_image_subset: "star_of ` A <= *s* A"
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17303
diff changeset
    81
by (auto simp add: STAR_mem)
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
    82
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
    83
lemma STAR_hypreal_of_real_Int: "*s* X Int Reals = hypreal_of_real ` X"
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17303
diff changeset
    84
by (auto simp add: SReal_def STAR_mem_iff)
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
    85
14468
6be497cacab5 heavy tidying
paulson
parents: 14378
diff changeset
    86
lemma lemma_not_hyprealA: "x \<notin> hypreal_of_real ` A ==> \<forall>y \<in> A. x \<noteq> hypreal_of_real y"
14371
c78c7da09519 Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents: 14370
diff changeset
    87
by auto
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
    88
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
    89
lemma lemma_Compl_eq: "- {n. X n = xa} = {n. X n \<noteq> xa}"
14371
c78c7da09519 Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents: 14370
diff changeset
    90
by auto
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
    91
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
    92
lemma STAR_real_seq_to_hypreal:
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17303
diff changeset
    93
    "\<forall>n. (X n) \<notin> M ==> star_n X \<notin> *s* M"
17429
e8d6ed3aacfe merged Transfer.thy and StarType.thy into StarDef.thy; renamed Ifun2_of to starfun2; cleaned up
huffman
parents: 17318
diff changeset
    94
apply (unfold starset_def star_of_def)
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17303
diff changeset
    95
apply (simp add: Iset_star_n)
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
    96
done
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
    97
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17303
diff changeset
    98
lemma STAR_singleton: "*s* {x} = {star_of x}"
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17303
diff changeset
    99
by simp
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17303
diff changeset
   100
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17303
diff changeset
   101
lemma STAR_not_mem: "x \<notin> F ==> star_of x \<notin> *s* F"
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17303
diff changeset
   102
by transfer
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   103
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   104
lemma STAR_subset_closed: "[| x : *s* A; A <= B |] ==> x : *s* B"
14371
c78c7da09519 Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents: 14370
diff changeset
   105
by (blast dest: STAR_subset)
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   106
15169
2b5da07a0b89 new "respects" syntax for quotienting
paulson
parents: 15140
diff changeset
   107
text{*Nonstandard extension of a set (defined using a constant
2b5da07a0b89 new "respects" syntax for quotienting
paulson
parents: 15140
diff changeset
   108
   sequence) as a special case of an internal set*}
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   109
14468
6be497cacab5 heavy tidying
paulson
parents: 14378
diff changeset
   110
lemma starset_n_starset: "\<forall>n. (As n = A) ==> *sn* As = *s* A"
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17303
diff changeset
   111
apply (drule expand_fun_eq [THEN iffD2])
17429
e8d6ed3aacfe merged Transfer.thy and StarType.thy into StarDef.thy; renamed Ifun2_of to starfun2; cleaned up
huffman
parents: 17318
diff changeset
   112
apply (simp add: starset_n_def starset_def star_of_def)
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17303
diff changeset
   113
done
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   114
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   115
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   116
(*----------------------------------------------------------------*)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   117
(* Theorems about nonstandard extensions of functions             *)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   118
(*----------------------------------------------------------------*)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   119
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   120
(*----------------------------------------------------------------*)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   121
(* Nonstandard extension of a function (defined using a           *)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   122
(* constant sequence) as a special case of an internal function   *)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   123
(*----------------------------------------------------------------*)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   124
14468
6be497cacab5 heavy tidying
paulson
parents: 14378
diff changeset
   125
lemma starfun_n_starfun: "\<forall>n. (F n = f) ==> *fn* F = *f* f"
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17303
diff changeset
   126
apply (drule expand_fun_eq [THEN iffD2])
17429
e8d6ed3aacfe merged Transfer.thy and StarType.thy into StarDef.thy; renamed Ifun2_of to starfun2; cleaned up
huffman
parents: 17318
diff changeset
   127
apply (simp add: starfun_n_def starfun_def star_of_def)
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17303
diff changeset
   128
done
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   129
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   130
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   131
(*
15003
6145dd7538d7 replaced monomorphic abs definitions by abs_if
paulson
parents: 14477
diff changeset
   132
   Prove that abs for hypreal is a nonstandard extension of abs for real w/o
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   133
   use of congruence property (proved after this for general
15003
6145dd7538d7 replaced monomorphic abs definitions by abs_if
paulson
parents: 14477
diff changeset
   134
   nonstandard extensions of real valued functions). 
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   135
15003
6145dd7538d7 replaced monomorphic abs definitions by abs_if
paulson
parents: 14477
diff changeset
   136
   Proof now Uses the ultrafilter tactic!
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   137
*)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   138
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   139
lemma hrabs_is_starext_rabs: "is_starext abs abs"
14468
6be497cacab5 heavy tidying
paulson
parents: 14378
diff changeset
   140
apply (simp add: is_starext_def, safe)
17429
e8d6ed3aacfe merged Transfer.thy and StarType.thy into StarDef.thy; renamed Ifun2_of to starfun2; cleaned up
huffman
parents: 17318
diff changeset
   141
apply (rule_tac x=x in star_cases)
e8d6ed3aacfe merged Transfer.thy and StarType.thy into StarDef.thy; renamed Ifun2_of to starfun2; cleaned up
huffman
parents: 17318
diff changeset
   142
apply (rule_tac x=y in star_cases)
e8d6ed3aacfe merged Transfer.thy and StarType.thy into StarDef.thy; renamed Ifun2_of to starfun2; cleaned up
huffman
parents: 17318
diff changeset
   143
apply (unfold star_n_def, auto)
17298
ad73fb6144cf replace type hypreal with real star
huffman
parents: 15170
diff changeset
   144
apply (rule bexI, rule_tac [2] lemma_starrel_refl)
ad73fb6144cf replace type hypreal with real star
huffman
parents: 15170
diff changeset
   145
apply (rule bexI, rule_tac [2] lemma_starrel_refl)
17302
25aab757672b generalized types
huffman
parents: 17298
diff changeset
   146
apply (fold star_n_def)
17429
e8d6ed3aacfe merged Transfer.thy and StarType.thy into StarDef.thy; renamed Ifun2_of to starfun2; cleaned up
huffman
parents: 17318
diff changeset
   147
apply (unfold star_abs_def starfun_def star_of_def)
17302
25aab757672b generalized types
huffman
parents: 17298
diff changeset
   148
apply (simp add: Ifun_star_n star_n_eq_iff)
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   149
done
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   150
17298
ad73fb6144cf replace type hypreal with real star
huffman
parents: 15170
diff changeset
   151
lemma Rep_star_FreeUltrafilterNat:
ad73fb6144cf replace type hypreal with real star
huffman
parents: 15170
diff changeset
   152
     "[| X \<in> Rep_star z; Y \<in> Rep_star z |]
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   153
      ==> {n. X n = Y n} : FreeUltrafilterNat"
17429
e8d6ed3aacfe merged Transfer.thy and StarType.thy into StarDef.thy; renamed Ifun2_of to starfun2; cleaned up
huffman
parents: 17318
diff changeset
   154
by (rule FreeUltrafilterNat_Rep_hypreal)
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   155
15169
2b5da07a0b89 new "respects" syntax for quotienting
paulson
parents: 15140
diff changeset
   156
text{*Nonstandard extension of functions*}
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   157
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   158
lemma starfun:
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17303
diff changeset
   159
      "( *f* f) (star_n X) = star_n (%n. f (X n))"
17429
e8d6ed3aacfe merged Transfer.thy and StarType.thy into StarDef.thy; renamed Ifun2_of to starfun2; cleaned up
huffman
parents: 17318
diff changeset
   160
by (simp add: starfun_def star_of_def Ifun_star_n)
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   161
14477
cc61fd03e589 conversion of Hyperreal/Lim to new-style
paulson
parents: 14468
diff changeset
   162
lemma starfun_if_eq:
17429
e8d6ed3aacfe merged Transfer.thy and StarType.thy into StarDef.thy; renamed Ifun2_of to starfun2; cleaned up
huffman
parents: 17318
diff changeset
   163
     "!!w. w \<noteq> star_of x
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17303
diff changeset
   164
       ==> ( *f* (\<lambda>z. if z = x then a else g z)) w = ( *f* g) w"
17429
e8d6ed3aacfe merged Transfer.thy and StarType.thy into StarDef.thy; renamed Ifun2_of to starfun2; cleaned up
huffman
parents: 17318
diff changeset
   165
by (transfer, simp)
14477
cc61fd03e589 conversion of Hyperreal/Lim to new-style
paulson
parents: 14468
diff changeset
   166
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   167
(*-------------------------------------------
14371
c78c7da09519 Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents: 14370
diff changeset
   168
  multiplication: ( *f) x ( *g) = *(f x g)
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   169
 ------------------------------------------*)
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17303
diff changeset
   170
lemma starfun_mult: "!!x. ( *f* f) x * ( *f* g) x = ( *f* (%x. f x * g x)) x"
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17303
diff changeset
   171
by (transfer, rule refl)
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   172
declare starfun_mult [symmetric, simp]
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   173
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   174
(*---------------------------------------
14371
c78c7da09519 Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents: 14370
diff changeset
   175
  addition: ( *f) + ( *g) = *(f + g)
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   176
 ---------------------------------------*)
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17303
diff changeset
   177
lemma starfun_add: "!!x. ( *f* f) x + ( *f* g) x = ( *f* (%x. f x + g x)) x"
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17303
diff changeset
   178
by (transfer, rule refl)
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   179
declare starfun_add [symmetric, simp]
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   180
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   181
(*--------------------------------------------
14371
c78c7da09519 Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents: 14370
diff changeset
   182
  subtraction: ( *f) + -( *g) = *(f + -g)
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   183
 -------------------------------------------*)
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17303
diff changeset
   184
lemma starfun_minus: "!!x. - ( *f* f) x = ( *f* (%x. - f x)) x"
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17303
diff changeset
   185
by (transfer, rule refl)
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   186
declare starfun_minus [symmetric, simp]
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   187
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   188
(*FIXME: delete*)
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17303
diff changeset
   189
lemma starfun_add_minus: "!!x. ( *f* f) x + -( *f* g) x = ( *f* (%x. f x + -g x)) x"
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17303
diff changeset
   190
by (transfer, rule refl)
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   191
declare starfun_add_minus [symmetric, simp]
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   192
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17303
diff changeset
   193
lemma starfun_diff: "!!x. ( *f* f) x  - ( *f* g) x = ( *f* (%x. f x - g x)) x"
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17303
diff changeset
   194
by (transfer, rule refl)
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   195
declare starfun_diff [symmetric, simp]
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   196
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   197
(*--------------------------------------
14371
c78c7da09519 Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents: 14370
diff changeset
   198
  composition: ( *f) o ( *g) = *(f o g)
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   199
 ---------------------------------------*)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   200
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   201
lemma starfun_o2: "(%x. ( *f* f) (( *f* g) x)) = *f* (%x. f (g x))"
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17303
diff changeset
   202
by (transfer, rule refl)
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   203
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   204
lemma starfun_o: "( *f* f) o ( *f* g) = ( *f* (f o g))"
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17303
diff changeset
   205
by (transfer o_def, rule refl)
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   206
15169
2b5da07a0b89 new "respects" syntax for quotienting
paulson
parents: 15140
diff changeset
   207
text{*NS extension of constant function*}
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17303
diff changeset
   208
lemma starfun_const_fun [simp]: "!!x. ( *f* (%x. k)) x = star_of k"
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17303
diff changeset
   209
by (transfer, rule refl)
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   210
15169
2b5da07a0b89 new "respects" syntax for quotienting
paulson
parents: 15140
diff changeset
   211
text{*the NS extension of the identity function*}
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   212
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17303
diff changeset
   213
lemma starfun_Id [simp]: "!!x. ( *f* (%x. x)) x = x"
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17303
diff changeset
   214
by (transfer, rule refl)
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   215
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17303
diff changeset
   216
(* this is trivial, given starfun_Id *)
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17303
diff changeset
   217
lemma starfun_Idfun_approx:
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17303
diff changeset
   218
  "x @= hypreal_of_real a ==> ( *f* (%x. x)) x @= hypreal_of_real a"
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17303
diff changeset
   219
by (simp only: starfun_Id)
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   220
15169
2b5da07a0b89 new "respects" syntax for quotienting
paulson
parents: 15140
diff changeset
   221
text{*The Star-function is a (nonstandard) extension of the function*}
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   222
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   223
lemma is_starext_starfun: "is_starext ( *f* f) f"
14468
6be497cacab5 heavy tidying
paulson
parents: 14378
diff changeset
   224
apply (simp add: is_starext_def, auto)
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17303
diff changeset
   225
apply (rule_tac x = x in star_cases)
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17303
diff changeset
   226
apply (rule_tac x = y in star_cases)
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17303
diff changeset
   227
apply (auto intro!: bexI [OF _ Rep_star_star_n]
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17303
diff changeset
   228
            simp add: starfun star_n_eq_iff)
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   229
done
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   230
15169
2b5da07a0b89 new "respects" syntax for quotienting
paulson
parents: 15140
diff changeset
   231
text{*Any nonstandard extension is in fact the Star-function*}
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   232
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   233
lemma is_starfun_starext: "is_starext F f ==> F = *f* f"
14468
6be497cacab5 heavy tidying
paulson
parents: 14378
diff changeset
   234
apply (simp add: is_starext_def)
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   235
apply (rule ext)
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17303
diff changeset
   236
apply (rule_tac x = x in star_cases)
14371
c78c7da09519 Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents: 14370
diff changeset
   237
apply (drule_tac x = x in spec)
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   238
apply (drule_tac x = "( *f* f) x" in spec)
14371
c78c7da09519 Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents: 14370
diff changeset
   239
apply (auto dest!: FreeUltrafilterNat_Compl_mem simp add: starfun, ultra)
14370
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 is_starext_starfun_iff: "(is_starext F f) = (F = *f* f)"
14371
c78c7da09519 Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents: 14370
diff changeset
   243
by (blast intro: is_starfun_starext is_starext_starfun)
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   244
15169
2b5da07a0b89 new "respects" syntax for quotienting
paulson
parents: 15140
diff changeset
   245
text{*extented function has same solution as its standard
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   246
   version for real arguments. i.e they are the same
15169
2b5da07a0b89 new "respects" syntax for quotienting
paulson
parents: 15140
diff changeset
   247
   for all real arguments*}
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17303
diff changeset
   248
lemma starfun_eq [simp]: "( *f* f) (star_of a) = star_of (f a)"
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17303
diff changeset
   249
by (transfer, rule refl)
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   250
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17303
diff changeset
   251
lemma starfun_approx: "( *f* f) (star_of a) @= hypreal_of_real (f a)"
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17303
diff changeset
   252
by simp
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   253
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   254
(* useful for NS definition of derivatives *)
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17303
diff changeset
   255
lemma starfun_lambda_cancel:
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17303
diff changeset
   256
  "!!x'. ( *f* (%h. f (x + h))) x'  = ( *f* f) (star_of x + x')"
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17303
diff changeset
   257
by (transfer, rule refl)
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   258
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17303
diff changeset
   259
lemma starfun_lambda_cancel2:
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17303
diff changeset
   260
  "( *f* (%h. f(g(x + h)))) x' = ( *f* (f o g)) (star_of x + x')"
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17303
diff changeset
   261
by (unfold o_def, rule starfun_lambda_cancel)
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   262
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17303
diff changeset
   263
lemma starfun_mult_HFinite_approx: "[| ( *f* f) x @= l; ( *f* g) x @= m;
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   264
                  l: HFinite; m: HFinite
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17303
diff changeset
   265
               |] ==>  ( *f* (%x. f x * g x)) x @= l * m"
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17303
diff changeset
   266
apply (drule (3) approx_mult_HFinite)
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   267
apply (auto intro: approx_HFinite [OF _ approx_sym])
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   268
done
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   269
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17303
diff changeset
   270
lemma starfun_add_approx: "[| ( *f* f) x @= l; ( *f* g) x @= m
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17303
diff changeset
   271
               |] ==>  ( *f* (%x. f x + g x)) x @= l + m"
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17303
diff changeset
   272
by (auto intro: approx_add)
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   273
15169
2b5da07a0b89 new "respects" syntax for quotienting
paulson
parents: 15140
diff changeset
   274
text{*Examples: hrabs is nonstandard extension of rabs
2b5da07a0b89 new "respects" syntax for quotienting
paulson
parents: 15140
diff changeset
   275
              inverse is nonstandard extension of inverse*}
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   276
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   277
(* can be proved easily using theorem "starfun" and *)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   278
(* properties of ultrafilter as for inverse below we  *)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   279
(* use the theorem we proved above instead          *)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   280
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   281
lemma starfun_rabs_hrabs: "*f* abs = abs"
14371
c78c7da09519 Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents: 14370
diff changeset
   282
by (rule hrabs_is_starext_rabs [THEN is_starext_starfun_iff [THEN iffD1], symmetric])
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   283
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17303
diff changeset
   284
lemma starfun_inverse_inverse [simp]: "( *f* inverse) x = inverse(x)"
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17303
diff changeset
   285
by (unfold star_inverse_def, rule refl)
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   286
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17303
diff changeset
   287
lemma starfun_inverse: "!!x. inverse (( *f* f) x) = ( *f* (%x. inverse (f x))) x"
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17303
diff changeset
   288
by (transfer, rule refl)
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   289
declare starfun_inverse [symmetric, simp]
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   290
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17303
diff changeset
   291
lemma starfun_divide: "!!x. ( *f* f) x / ( *f* g) x = ( *f* (%x. f x / g x)) x"
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17303
diff changeset
   292
by (transfer, rule refl)
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   293
declare starfun_divide [symmetric, simp]
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   294
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17303
diff changeset
   295
lemma starfun_inverse2: "!!x. inverse (( *f* f) x) = ( *f* (%x. inverse (f x))) x"
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17303
diff changeset
   296
by (transfer, rule refl)
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   297
15169
2b5da07a0b89 new "respects" syntax for quotienting
paulson
parents: 15140
diff changeset
   298
text{*General lemma/theorem needed for proofs in elementary
2b5da07a0b89 new "respects" syntax for quotienting
paulson
parents: 15140
diff changeset
   299
    topology of the reals*}
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   300
lemma starfun_mem_starset:
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17303
diff changeset
   301
      "!!x. ( *f* f) x : *s* A ==> x : *s* {x. f x  \<in> A}"
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17303
diff changeset
   302
by (transfer, simp)
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   303
15169
2b5da07a0b89 new "respects" syntax for quotienting
paulson
parents: 15140
diff changeset
   304
text{*Alternative definition for hrabs with rabs function
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   305
   applied entrywise to equivalence class representative.
15169
2b5da07a0b89 new "respects" syntax for quotienting
paulson
parents: 15140
diff changeset
   306
   This is easily proved using starfun and ns extension thm*}
15170
e7d4d3314f4c fixed presentation
paulson
parents: 15169
diff changeset
   307
lemma hypreal_hrabs:
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17303
diff changeset
   308
     "abs (star_n X) = star_n (%n. abs (X n))"
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17303
diff changeset
   309
by (simp only: starfun_rabs_hrabs [symmetric] starfun)
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   310
15169
2b5da07a0b89 new "respects" syntax for quotienting
paulson
parents: 15140
diff changeset
   311
text{*nonstandard extension of set through nonstandard extension
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   312
   of rabs function i.e hrabs. A more general result should be
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   313
   where we replace rabs by some arbitrary function f and hrabs
15169
2b5da07a0b89 new "respects" syntax for quotienting
paulson
parents: 15140
diff changeset
   314
   by its NS extenson. See second NS set extension below.*}
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   315
lemma STAR_rabs_add_minus:
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   316
   "*s* {x. abs (x + - y) < r} =
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   317
     {x. abs(x + -hypreal_of_real y) < hypreal_of_real r}"
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17303
diff changeset
   318
by (transfer, rule refl)
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   319
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   320
lemma STAR_starfun_rabs_add_minus:
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   321
  "*s* {x. abs (f x + - y) < r} =
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   322
       {x. abs(( *f* f) x + -hypreal_of_real y) < hypreal_of_real r}"
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17303
diff changeset
   323
by (transfer, rule refl)
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   324
15169
2b5da07a0b89 new "respects" syntax for quotienting
paulson
parents: 15140
diff changeset
   325
text{*Another characterization of Infinitesimal and one of @= relation.
15170
e7d4d3314f4c fixed presentation
paulson
parents: 15169
diff changeset
   326
   In this theory since @{text hypreal_hrabs} proved here. Maybe
e7d4d3314f4c fixed presentation
paulson
parents: 15169
diff changeset
   327
   move both theorems??*}
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14371
diff changeset
   328
lemma Infinitesimal_FreeUltrafilterNat_iff2:
14468
6be497cacab5 heavy tidying
paulson
parents: 14378
diff changeset
   329
     "(x \<in> Infinitesimal) =
17298
ad73fb6144cf replace type hypreal with real star
huffman
parents: 15170
diff changeset
   330
      (\<exists>X \<in> Rep_star(x).
14468
6be497cacab5 heavy tidying
paulson
parents: 14378
diff changeset
   331
        \<forall>m. {n. abs(X n) < inverse(real(Suc m))}
6be497cacab5 heavy tidying
paulson
parents: 14378
diff changeset
   332
                \<in>  FreeUltrafilterNat)"
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17303
diff changeset
   333
apply (cases x)
17298
ad73fb6144cf replace type hypreal with real star
huffman
parents: 15170
diff changeset
   334
apply (auto intro!: bexI lemma_starrel_refl 
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17303
diff changeset
   335
            simp add: Infinitesimal_hypreal_of_nat_iff star_of_def
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17303
diff changeset
   336
     star_n_inverse star_n_abs star_n_less hypreal_of_nat_eq)
14371
c78c7da09519 Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents: 14370
diff changeset
   337
apply (drule_tac x = n in spec, ultra)
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   338
done
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   339
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17303
diff changeset
   340
lemma approx_FreeUltrafilterNat_iff: "star_n X @= star_n Y =
14468
6be497cacab5 heavy tidying
paulson
parents: 14378
diff changeset
   341
      (\<forall>m. {n. abs (X n + - Y n) <
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   342
                  inverse(real(Suc m))} : FreeUltrafilterNat)"
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   343
apply (subst approx_minus_iff)
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   344
apply (rule mem_infmal_iff [THEN subst])
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17303
diff changeset
   345
apply (auto simp add: star_n_minus star_n_add Infinitesimal_FreeUltrafilterNat_iff2)
14371
c78c7da09519 Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents: 14370
diff changeset
   346
apply (drule_tac x = m in spec, ultra)
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   347
done
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   348
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   349
lemma inj_starfun: "inj starfun"
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   350
apply (rule inj_onI)
14371
c78c7da09519 Conversion of HyperNat to Isar format and its declaration as a semiring
paulson
parents: 14370
diff changeset
   351
apply (rule ext, rule ccontr)
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17303
diff changeset
   352
apply (drule_tac x = "star_n (%n. xa)" in fun_cong)
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17303
diff changeset
   353
apply (auto simp add: starfun star_n_eq_iff)
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   354
done
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   355
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   356
ML
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   357
{*
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   358
val starset_n_def = thm"starset_n_def";
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   359
val InternalSets_def = thm"InternalSets_def";
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   360
val is_starext_def = thm"is_starext_def";
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   361
val starfun_n_def = thm"starfun_n_def";
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   362
val InternalFuns_def = thm"InternalFuns_def";
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   363
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   364
val no_choice = thm "no_choice";
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17303
diff changeset
   365
val STAR_UNIV_set = thm "STAR_UNIV_set";
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   366
val STAR_empty_set = thm "STAR_empty_set";
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   367
val STAR_Un = thm "STAR_Un";
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   368
val STAR_Int = thm "STAR_Int";
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   369
val STAR_Compl = thm "STAR_Compl";
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   370
val STAR_mem_Compl = thm "STAR_mem_Compl";
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   371
val STAR_diff = thm "STAR_diff";
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   372
val STAR_subset = thm "STAR_subset";
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   373
val STAR_mem = thm "STAR_mem";
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17303
diff changeset
   374
val STAR_star_of_image_subset = thm "STAR_star_of_image_subset";
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   375
val STAR_hypreal_of_real_Int = thm "STAR_hypreal_of_real_Int";
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   376
val STAR_real_seq_to_hypreal = thm "STAR_real_seq_to_hypreal";
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   377
val STAR_singleton = thm "STAR_singleton";
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   378
val STAR_not_mem = thm "STAR_not_mem";
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   379
val STAR_subset_closed = thm "STAR_subset_closed";
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   380
val starset_n_starset = thm "starset_n_starset";
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   381
val starfun_n_starfun = thm "starfun_n_starfun";
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   382
val hrabs_is_starext_rabs = thm "hrabs_is_starext_rabs";
17298
ad73fb6144cf replace type hypreal with real star
huffman
parents: 15170
diff changeset
   383
val Rep_star_FreeUltrafilterNat = thm "Rep_star_FreeUltrafilterNat";
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   384
val starfun = thm "starfun";
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   385
val starfun_mult = thm "starfun_mult";
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   386
val starfun_add = thm "starfun_add";
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   387
val starfun_minus = thm "starfun_minus";
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   388
val starfun_add_minus = thm "starfun_add_minus";
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   389
val starfun_diff = thm "starfun_diff";
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   390
val starfun_o2 = thm "starfun_o2";
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   391
val starfun_o = thm "starfun_o";
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   392
val starfun_const_fun = thm "starfun_const_fun";
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   393
val starfun_Idfun_approx = thm "starfun_Idfun_approx";
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   394
val starfun_Id = thm "starfun_Id";
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   395
val is_starext_starfun = thm "is_starext_starfun";
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   396
val is_starfun_starext = thm "is_starfun_starext";
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   397
val is_starext_starfun_iff = thm "is_starext_starfun_iff";
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   398
val starfun_eq = thm "starfun_eq";
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   399
val starfun_approx = thm "starfun_approx";
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   400
val starfun_lambda_cancel = thm "starfun_lambda_cancel";
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   401
val starfun_lambda_cancel2 = thm "starfun_lambda_cancel2";
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   402
val starfun_mult_HFinite_approx = thm "starfun_mult_HFinite_approx";
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   403
val starfun_add_approx = thm "starfun_add_approx";
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   404
val starfun_rabs_hrabs = thm "starfun_rabs_hrabs";
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   405
val starfun_inverse_inverse = thm "starfun_inverse_inverse";
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   406
val starfun_inverse = thm "starfun_inverse";
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   407
val starfun_divide = thm "starfun_divide";
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   408
val starfun_inverse2 = thm "starfun_inverse2";
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   409
val starfun_mem_starset = thm "starfun_mem_starset";
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   410
val hypreal_hrabs = thm "hypreal_hrabs";
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   411
val STAR_rabs_add_minus = thm "STAR_rabs_add_minus";
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   412
val STAR_starfun_rabs_add_minus = thm "STAR_starfun_rabs_add_minus";
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   413
val Infinitesimal_FreeUltrafilterNat_iff2 = thm "Infinitesimal_FreeUltrafilterNat_iff2";
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   414
val approx_FreeUltrafilterNat_iff = thm "approx_FreeUltrafilterNat_iff";
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   415
val inj_starfun = thm "inj_starfun";
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   416
*}
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   417
b0064703967b simplifications in the hyperreals
paulson
parents: 10834
diff changeset
   418
end