src/HOL/Real/Hyperreal/Lim.ML
author wenzelm
Fri, 06 Oct 2000 17:35:58 +0200
changeset 10168 50be659d4222
parent 10045 c76b73e16711
child 10607 352f6f209775
permissions -rw-r--r--
final tuning;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
10045
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
     1
(*  Title       : Lim.ML
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
     2
    Author      : Jacques D. Fleuriot
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
     3
    Copyright   : 1998  University of Cambridge
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
     4
    Description : Theory of limits, continuity and 
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
     5
                  differentiation of real=>real functions
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
     6
*)
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
     7
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
     8
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
     9
fun ARITH_PROVE str = prove_goal thy str 
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    10
                      (fn prems => [cut_facts_tac prems 1,arith_tac 1]);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    11
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    12
(*---------------------------------------------------------------
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    13
   Theory of limits, continuity and differentiation of 
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    14
   real=>real functions 
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    15
 ----------------------------------------------------------------*)
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    16
Goalw [LIM_def] "(f -- a --> L) = (ALL r. #0 < r --> \
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    17
\    (EX s. #0 < s & (ALL x. (#0 < abs(x + -a) & (abs(x + -a) < s) \
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    18
\          --> abs(f x + -L) < r))))";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    19
by (Blast_tac 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    20
qed "LIM_iff";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    21
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    22
Goalw [LIM_def] 
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    23
      "!!a. [| f -- a --> L; #0 < r |] \
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    24
\           ==> (EX s. #0 < s & (ALL x. (#0 < abs(x + -a) \
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    25
\           & (abs(x + -a) < s) --> abs(f x + -L) < r)))";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    26
by (Blast_tac 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    27
qed "LIMD";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    28
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    29
Goalw [LIM_def] "(%x. k) -- x --> k";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    30
by (Auto_tac);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    31
qed "LIM_const";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    32
Addsimps [LIM_const];
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    33
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    34
(***-----------------------------------------------------------***)
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    35
(***  Some Purely Standard Proofs - Can be used for comparison ***)
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    36
(***-----------------------------------------------------------***)
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    37
 
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    38
(*--------------- 
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    39
    LIM_add    
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    40
 ---------------*)
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    41
Goalw [LIM_def] 
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    42
     "[| f -- x --> l; g -- x --> m |] \
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    43
\     ==> (%x. f(x) + g(x)) -- x --> (l + m)";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    44
by (Step_tac 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    45
by (REPEAT(dres_inst_tac [("x","r*rinv(#1 + #1)")] spec 1));
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    46
by (dtac (rename_numerals (real_zero_less_two RS real_rinv_gt_zero 
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    47
    RSN (2,real_mult_less_mono2))) 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    48
by (Asm_full_simp_tac 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    49
by (Clarify_tac 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    50
by (res_inst_tac [("R1.0","s"),("R2.0","sa")] 
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    51
    real_linear_less2 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    52
by (res_inst_tac [("x","s")] exI 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    53
by (res_inst_tac [("x","sa")] exI 2);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    54
by (res_inst_tac [("x","sa")] exI 3);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    55
by (Step_tac 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    56
by (REPEAT(dres_inst_tac [("x","xa")] spec 1) 
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    57
    THEN step_tac (claset() addSEs [real_less_trans]) 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    58
by (REPEAT(dres_inst_tac [("x","xa")] spec 2) 
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    59
    THEN step_tac (claset() addSEs [real_less_trans]) 2);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    60
by (REPEAT(dres_inst_tac [("x","xa")] spec 3) 
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    61
    THEN step_tac (claset() addSEs [real_less_trans]) 3);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    62
by (ALLGOALS(rtac (abs_sum_triangle_ineq RS real_le_less_trans)));
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    63
by (ALLGOALS(rtac (real_sum_of_halves RS subst)));
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    64
by (auto_tac (claset() addIs [real_add_less_mono],simpset()));
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    65
qed "LIM_add";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    66
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    67
Goalw [LIM_def] "f -- a --> L ==> (%x. -f(x)) -- a --> -L";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    68
by (full_simp_tac (simpset() addsimps [real_minus_add_distrib RS sym,
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    69
    abs_minus_cancel] delsimps [real_minus_add_distrib,real_minus_minus]) 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    70
qed "LIM_minus";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    71
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    72
(*----------------------------------------------
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    73
     LIM_add_minus
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    74
 ----------------------------------------------*)
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    75
Goal "[| f -- x --> l; g -- x --> m |] \
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    76
\     ==> (%x. f(x) + -g(x)) -- x --> (l + -m)";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    77
by (blast_tac (claset() addDs [LIM_add,LIM_minus]) 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    78
qed "LIM_add_minus";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    79
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    80
(*----------------------------------------------
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    81
     LIM_zero
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    82
 ----------------------------------------------*)
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    83
Goal "f -- a --> l ==> (%x. f(x) + -l) -- a --> #0";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    84
by (res_inst_tac [("z1","l")] (rename_numerals 
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    85
    (real_add_minus RS subst)) 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    86
by (rtac LIM_add_minus 1 THEN Auto_tac);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    87
qed "LIM_zero";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    88
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    89
(*--------------------------
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    90
   Limit not zero
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    91
 --------------------------*)
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    92
Goalw [LIM_def] "k ~= #0 ==> ~ ((%x. k) -- x --> #0)";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    93
by (res_inst_tac [("R1.0","k"),("R2.0","#0")] real_linear_less2 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    94
by (auto_tac (claset(),simpset() addsimps [abs_minus_eqI2,abs_eqI2]));
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    95
by (dtac (rename_numerals (real_minus_zero_less_iff RS iffD2)) 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    96
by (res_inst_tac [("x","-k")] exI 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    97
by (res_inst_tac [("x","k")] exI 2);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    98
by Auto_tac;
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    99
by (ALLGOALS(dres_inst_tac [("y","s")] real_dense));
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   100
by (Step_tac 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   101
by (ALLGOALS(res_inst_tac [("x","r + x")] exI));
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   102
by (auto_tac (claset(),simpset() addsimps [real_add_assoc,abs_eqI2]));
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   103
qed "LIM_not_zero";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   104
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   105
(* [| k ~= #0; (%x. k) -- x --> #0 |] ==> R *)
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   106
bind_thm("LIM_not_zeroE", (LIM_not_zero RS notE));
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   107
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   108
Goal "(%x. k) -- x --> L ==> k = L";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   109
by (rtac ccontr 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   110
by (dtac LIM_zero 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   111
by (rtac LIM_not_zeroE 1 THEN assume_tac 2);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   112
by (arith_tac 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   113
qed "LIM_const_eq";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   114
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   115
(*------------------------
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   116
     Limit is Unique
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   117
 ------------------------*)
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   118
Goal "[| f -- x --> L; f -- x --> M |] ==> L = M";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   119
by (dtac LIM_minus 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   120
by (dtac LIM_add 1 THEN assume_tac 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   121
by (auto_tac (claset() addSDs [LIM_const_eq RS sym],
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   122
    simpset()));
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   123
qed "LIM_unique";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   124
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   125
(*-------------
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   126
    LIM_mult_zero
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   127
 -------------*)
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   128
Goalw [LIM_def] "!!f. [| f -- x --> #0; g -- x --> #0 |] \
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   129
\         ==> (%x. f(x)*g(x)) -- x --> #0";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   130
by (Step_tac 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   131
by (dres_inst_tac [("x","#1")] spec 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   132
by (dres_inst_tac [("x","r")] spec 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   133
by (cut_facts_tac [real_zero_less_one] 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   134
by (asm_full_simp_tac (simpset() addsimps 
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   135
    [abs_mult]) 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   136
by (Clarify_tac 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   137
by (res_inst_tac [("R1.0","s"),("R2.0","sa")] 
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   138
    real_linear_less2 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   139
by (res_inst_tac [("x","s")] exI 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   140
by (res_inst_tac [("x","sa")] exI 2);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   141
by (res_inst_tac [("x","sa")] exI 3);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   142
by (Step_tac 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   143
by (REPEAT(dres_inst_tac [("x","xa")] spec 1) 
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   144
    THEN step_tac (claset() addSEs [real_less_trans]) 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   145
by (REPEAT(dres_inst_tac [("x","xa")] spec 2) 
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   146
    THEN step_tac (claset() addSEs [real_less_trans]) 2);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   147
by (REPEAT(dres_inst_tac [("x","xa")] spec 3) 
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   148
    THEN step_tac (claset() addSEs [real_less_trans]) 3);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   149
by (ALLGOALS(res_inst_tac [("t","r")] (real_mult_1 RS subst)));
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   150
by (ALLGOALS(rtac abs_mult_less2));
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   151
by Auto_tac;
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   152
qed "LIM_mult_zero";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   153
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   154
Goalw [LIM_def] "(%x. x) -- a --> a";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   155
by (Auto_tac);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   156
qed "LIM_self";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   157
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   158
Goal "(#0::real) < abs (x + -a) ==> x ~= a";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   159
by (arith_tac 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   160
qed "lemma_rabs_not_eq";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   161
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   162
(*--------------------------------------------------------------
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   163
   Limits are equal for functions equal except at limit point
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   164
 --------------------------------------------------------------*)
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   165
Goalw [LIM_def] 
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   166
      "[| ALL x. x ~= a --> (f x = g x) |] \
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   167
\           ==> (f -- a --> l) = (g -- a --> l)";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   168
by (auto_tac (claset(),simpset() addsimps [lemma_rabs_not_eq]));
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   169
qed "LIM_equal";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   170
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   171
Goal "[| (%x. f(x) + -g(x)) -- a --> #0; \
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   172
\        g -- a --> l |] \
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   173
\      ==> f -- a --> l";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   174
by (dtac LIM_add 1 THEN assume_tac 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   175
by (auto_tac (claset(),simpset() addsimps 
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   176
    [real_add_assoc]));
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   177
qed "LIM_trans";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   178
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   179
(***-------------------------------------------------------------***)
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   180
(***           End of Purely Standard Proofs                     ***)
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   181
(***-------------------------------------------------------------***)
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   182
(*--------------------------------------------------------------
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   183
       Standard and NS definitions of Limit
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   184
 --------------------------------------------------------------*)
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   185
Goalw [LIM_def,NSLIM_def,inf_close_def] 
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   186
      "f -- x --> L ==> f -- x --NS> L";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   187
by (asm_full_simp_tac (simpset() addsimps 
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   188
    [Infinitesimal_FreeUltrafilterNat_iff]) 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   189
by (Step_tac 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   190
by (res_inst_tac [("z","xa")] eq_Abs_hypreal 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   191
by (auto_tac (claset(),simpset() addsimps [starfun,
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   192
    hypreal_minus,hypreal_of_real_minus RS sym,
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   193
    hypreal_of_real_def,hypreal_add]));
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   194
by (rtac bexI 1 THEN rtac lemma_hyprel_refl 2 THEN Step_tac 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   195
by (dres_inst_tac [("x","u")] spec 1 THEN Clarify_tac 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   196
by (dres_inst_tac [("x","s")] spec 1 THEN Clarify_tac 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   197
by (subgoal_tac "ALL n::nat. (#0::real) < abs ((xa n) + - x) & \
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   198
\  abs ((xa n) + - x) < s --> abs (f (xa n) + - L) < u" 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   199
by (Blast_tac 2);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   200
by (dtac FreeUltrafilterNat_all 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   201
by (Ultra_tac 1 THEN arith_tac 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   202
qed "LIM_NSLIM";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   203
 
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   204
(*---------------------------------------------------------------------
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   205
    Limit: NS definition ==> standard definition
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   206
 ---------------------------------------------------------------------*)
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   207
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   208
Goal "ALL s. #0 < s --> (EX xa.  #0 < abs (xa + - x) & \
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   209
\        abs (xa + - x) < s  & r <= abs (f xa + -L)) \
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   210
\     ==> ALL n::nat. EX xa.  #0 < abs (xa + - x) & \
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   211
\             abs(xa + -x) < rinv(real_of_posnat n) & r <= abs(f xa + -L)";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   212
by (Step_tac 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   213
by (cut_inst_tac [("n1","n")] (real_of_posnat_gt_zero RS real_rinv_gt_zero) 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   214
by Auto_tac;
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   215
val lemma_LIM = result();
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   216
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   217
Goal "ALL s. #0 < s --> (EX xa.  #0 < abs (xa + - x) & \
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   218
\        abs (xa + - x) < s  & r <= abs (f xa + -L)) \
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   219
\     ==> EX X. ALL n::nat. #0 < abs (X n + - x) & \
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   220
\               abs(X n + -x) < rinv(real_of_posnat n) & r <= abs(f (X n) + -L)";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   221
by (dtac lemma_LIM 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   222
by (dtac choice 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   223
by (Blast_tac 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   224
val lemma_skolemize_LIM2 = result();
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   225
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   226
Goal "{n. f (X n) + - L = Y n} Int \
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   227
\         {n. #0 < abs (X n + - x) & \
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   228
\             abs (X n + - x) < rinv (real_of_posnat  n) & \
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   229
\             r <= abs (f (X n) + - L)} <= \
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   230
\         {n. r <= abs (Y n)}";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   231
by (Auto_tac);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   232
val lemma_Int = result ();
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   233
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   234
Goal "!!X. ALL n. #0 < abs (X n + - x) & \
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   235
\         abs (X n + - x) < rinv (real_of_posnat  n) & \
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   236
\         r <= abs (f (X n) + - L) ==> \
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   237
\         ALL n. abs (X n + - x) < rinv (real_of_posnat  n)";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   238
by (Auto_tac );
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   239
val lemma_simp = result();
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   240
 
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   241
(*-------------------
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   242
    NSLIM => LIM
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   243
 -------------------*)
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   244
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   245
Goalw [LIM_def,NSLIM_def,inf_close_def] 
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   246
      "!!f. f -- x --NS> L ==> f -- x --> L";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   247
by (asm_full_simp_tac (simpset() addsimps 
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   248
    [Infinitesimal_FreeUltrafilterNat_iff]) 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   249
by (EVERY1[Step_tac, rtac ccontr, Asm_full_simp_tac]);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   250
by (fold_tac [real_le_def]);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   251
by (dtac lemma_skolemize_LIM2 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   252
by (Step_tac 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   253
by (dres_inst_tac [("x","Abs_hypreal(hyprel^^{X})")] spec 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   254
by (asm_full_simp_tac (simpset() addsimps [starfun,
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   255
    hypreal_minus,hypreal_of_real_minus RS sym,
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   256
    hypreal_of_real_def,hypreal_add]) 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   257
by (Step_tac 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   258
by (dtac FreeUltrafilterNat_Ex_P 1 THEN etac exE 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   259
by (dres_inst_tac [("x","n")] spec 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   260
by (asm_full_simp_tac (simpset() addsimps 
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   261
    [real_add_minus,real_less_not_refl]) 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   262
by (dtac (lemma_simp RS real_seq_to_hypreal_Infinitesimal) 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   263
by (asm_full_simp_tac (simpset() addsimps 
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   264
    [Infinitesimal_FreeUltrafilterNat_iff,hypreal_of_real_def,
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   265
     hypreal_minus,hypreal_of_real_minus,hypreal_add]) 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   266
by (rotate_tac 2 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   267
by (dres_inst_tac [("x","r")] spec 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   268
by (Clarify_tac 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   269
by (dtac FreeUltrafilterNat_all 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   270
by (Ultra_tac 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   271
qed "NSLIM_LIM";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   272
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   273
(**** Key result ****)
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   274
Goal "(f -- x --> L) = (f -- x --NS> L)";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   275
by (blast_tac (claset() addIs [LIM_NSLIM,NSLIM_LIM]) 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   276
qed "LIM_NSLIM_iff";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   277
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   278
(*-------------------------------------------------------------------*)
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   279
(*   Proving properties of limits using nonstandard definition and   *)
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   280
(*   hence, the properties hold for standard limits as well          *)
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   281
(*-------------------------------------------------------------------*)
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   282
(*------------------------------------------------
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   283
      NSLIM_mult and hence (trivially) LIM_mult
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   284
 ------------------------------------------------*)
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   285
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   286
Goalw [NSLIM_def]
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   287
     "[| f -- x --NS> l; g -- x --NS> m |] \
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   288
\     ==> (%x. f(x) * g(x)) -- x --NS> (l * m)";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   289
by (auto_tac (claset() addSIs [starfun_mult_HFinite_inf_close],
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   290
              simpset() addsimps [hypreal_of_real_mult]));
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   291
qed "NSLIM_mult";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   292
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   293
Goal "[| f -- x --> l; g -- x --> m |] \
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   294
\     ==> (%x. f(x) * g(x)) -- x --> (l * m)";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   295
by (asm_full_simp_tac (simpset() addsimps [LIM_NSLIM_iff,
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   296
    NSLIM_mult]) 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   297
qed "LIM_mult2";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   298
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   299
(*----------------------------------------------
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   300
      NSLIM_add and hence (trivially) LIM_add
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   301
      Note the much shorter proof
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   302
 ----------------------------------------------*)
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   303
Goalw [NSLIM_def]
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   304
     "[| f -- x --NS> l; g -- x --NS> m |] \
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   305
\     ==> (%x. f(x) + g(x)) -- x --NS> (l + m)";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   306
by (auto_tac (claset() addSIs [starfun_add_inf_close],
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   307
              simpset() addsimps [hypreal_of_real_add]));
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   308
qed "NSLIM_add";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   309
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   310
Goal "[| f -- x --> l; g -- x --> m |] \
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   311
\     ==> (%x. f(x) + g(x)) -- x --> (l + m)";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   312
by (asm_full_simp_tac (simpset() addsimps [LIM_NSLIM_iff,
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   313
    NSLIM_add]) 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   314
qed "LIM_add2";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   315
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   316
(*----------------------------------------------
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   317
     NSLIM_const
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   318
 ----------------------------------------------*)
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   319
Goalw [NSLIM_def] "(%x. k) -- x --NS> k";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   320
by (Auto_tac);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   321
qed "NSLIM_const";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   322
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   323
Addsimps [NSLIM_const];
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   324
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   325
Goal "(%x. k) -- x --> k";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   326
by (asm_full_simp_tac (simpset() addsimps [LIM_NSLIM_iff]) 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   327
qed "LIM_const2";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   328
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   329
(*----------------------------------------------
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   330
     NSLIM_minus
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   331
 ----------------------------------------------*)
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   332
Goalw [NSLIM_def] 
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   333
      "f -- a --NS> L ==> (%x. -f(x)) -- a --NS> -L";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   334
by (auto_tac (claset() addIs [inf_close_minus],
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   335
    simpset() addsimps [starfun_minus RS sym,
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   336
    hypreal_of_real_minus]));
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   337
qed "NSLIM_minus";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   338
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   339
Goal "f -- a --> L ==> (%x. -f(x)) -- a --> -L";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   340
by (asm_full_simp_tac (simpset() addsimps [LIM_NSLIM_iff,
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   341
    NSLIM_minus]) 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   342
qed "LIM_minus2";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   343
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   344
(*----------------------------------------------
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   345
     NSLIM_add_minus
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   346
 ----------------------------------------------*)
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   347
Goal "!!f. [| f -- x --NS> l; g -- x --NS> m |] \
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   348
\     ==> (%x. f(x) + -g(x)) -- x --NS> (l + -m)";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   349
by (blast_tac (claset() addDs [NSLIM_add,NSLIM_minus]) 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   350
qed "NSLIM_add_minus";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   351
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   352
Goal "!!f. [| f -- x --> l; g -- x --> m |] \
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   353
\     ==> (%x. f(x) + -g(x)) -- x --> (l + -m)";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   354
by (asm_full_simp_tac (simpset() addsimps [LIM_NSLIM_iff,
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   355
    NSLIM_add_minus]) 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   356
qed "LIM_add_minus2";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   357
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   358
(*-----------------------------
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   359
    NSLIM_rinv
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   360
 -----------------------------*)
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   361
Goalw [NSLIM_def] 
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   362
      "!!f. [| f -- a --NS> L; \
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   363
\              L ~= #0 \
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   364
\           |] ==> (%x. rinv(f(x))) -- a --NS> (rinv L)";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   365
by (Step_tac 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   366
by (dtac spec 1 THEN auto_tac (claset(),simpset() 
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   367
    addsimps [hypreal_of_real_not_zero_iff RS sym,
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   368
    hypreal_of_real_hrinv RS sym]));
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   369
by (forward_tac [inf_close_hypreal_of_real_not_zero] 1 
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   370
    THEN assume_tac 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   371
by (auto_tac (claset() addSEs [(starfun_hrinv2 RS subst),
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   372
    inf_close_hrinv,hypreal_of_real_HFinite_diff_Infinitesimal],
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   373
    simpset()));
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   374
qed "NSLIM_rinv";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   375
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   376
Goal "[| f -- a --> L; \
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   377
\        L ~= #0 |] ==> (%x. rinv(f(x))) -- a --> (rinv L)";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   378
by (asm_full_simp_tac (simpset() addsimps [LIM_NSLIM_iff,
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   379
    NSLIM_rinv]) 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   380
qed "LIM_rinv";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   381
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   382
(*------------------------------
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   383
    NSLIM_zero
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   384
 ------------------------------*)
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   385
Goal "f -- a --NS> l ==> (%x. f(x) + -l) -- a --NS> #0";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   386
by (res_inst_tac [("z1","l")] (rename_numerals 
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   387
    (real_add_minus RS subst)) 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   388
by (rtac NSLIM_add_minus 1 THEN Auto_tac);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   389
qed "NSLIM_zero";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   390
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   391
Goal "f -- a --> l ==> (%x. f(x) + -l) -- a --> #0";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   392
by (asm_full_simp_tac (simpset() addsimps [LIM_NSLIM_iff,
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   393
    NSLIM_zero]) 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   394
qed "LIM_zero2";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   395
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   396
Goal "(%x. f(x) - l) -- x --NS> #0 ==> f -- x --NS> l";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   397
by (dres_inst_tac [("g","%x. l"),("m","l")] NSLIM_add 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   398
by (auto_tac (claset(),simpset() addsimps [real_diff_def,
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   399
    real_add_assoc]));
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   400
qed "NSLIM_zero_cancel";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   401
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   402
Goal "(%x. f(x) - l) -- x --> #0 ==> f -- x --> l";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   403
by (dres_inst_tac [("g","%x. l"),("m","l")] LIM_add 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   404
by (auto_tac (claset(),simpset() addsimps [real_diff_def,
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   405
    real_add_assoc]));
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   406
qed "LIM_zero_cancel";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   407
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   408
(*--------------------------
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   409
   NSLIM_not_zero
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   410
 --------------------------*)
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   411
Goalw [NSLIM_def] "k ~= #0 ==> ~ ((%x. k) -- x --NS> #0)";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   412
by (Auto_tac);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   413
by (res_inst_tac [("x","hypreal_of_real x + ehr")] exI 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   414
by (auto_tac (claset() addIs [Infinitesimal_add_inf_close_self 
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   415
    RS inf_close_sym],simpset()));
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   416
by (dres_inst_tac [("x1","-hypreal_of_real x")] (hypreal_add_left_cancel RS iffD2) 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   417
by (asm_full_simp_tac (simpset() addsimps [hypreal_add_assoc RS sym,
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   418
    hypreal_epsilon_not_zero]) 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   419
qed "NSLIM_not_zero";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   420
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   421
(* [| k ~= #0; (%x. k) -- x --NS> #0 |] ==> R *)
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   422
bind_thm("NSLIM_not_zeroE", (NSLIM_not_zero RS notE));
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   423
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   424
Goal "k ~= #0 ==> ~ ((%x. k) -- x --> #0)";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   425
by (asm_full_simp_tac (simpset() addsimps [LIM_NSLIM_iff,
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   426
    NSLIM_not_zero]) 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   427
qed "LIM_not_zero2";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   428
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   429
(*-------------------------------------
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   430
   NSLIM of constant function
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   431
 -------------------------------------*)
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   432
Goal "(%x. k) -- x --NS> L ==> k = L";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   433
by (rtac ccontr 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   434
by (dtac NSLIM_zero 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   435
by (rtac NSLIM_not_zeroE 1 THEN assume_tac 2);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   436
by (arith_tac 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   437
qed "NSLIM_const_eq";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   438
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   439
Goal "(%x. k) -- x --> L ==> k = L";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   440
by (asm_full_simp_tac (simpset() addsimps [LIM_NSLIM_iff,
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   441
    NSLIM_const_eq]) 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   442
qed "LIM_const_eq2";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   443
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   444
(*------------------------
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   445
     NS Limit is Unique
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   446
 ------------------------*)
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   447
(* can actually be proved more easily by unfolding def! *)
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   448
Goal "[| f -- x --NS> L; f -- x --NS> M |] ==> L = M";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   449
by (dtac NSLIM_minus 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   450
by (dtac NSLIM_add 1 THEN assume_tac 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   451
by (auto_tac (claset() addSDs [NSLIM_const_eq RS sym],
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   452
    simpset() addsimps [real_add_minus]));
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   453
qed "NSLIM_unique";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   454
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   455
Goal "[| f -- x --> L; f -- x --> M |] ==> L = M";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   456
by (asm_full_simp_tac (simpset() addsimps [LIM_NSLIM_iff,
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   457
    NSLIM_unique]) 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   458
qed "LIM_unique2";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   459
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   460
(*--------------------
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   461
    NSLIM_mult_zero
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   462
 --------------------*)
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   463
Goal "[| f -- x --NS> #0; g -- x --NS> #0 |] \
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   464
\         ==> (%x. f(x)*g(x)) -- x --NS> #0";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   465
by (dtac NSLIM_mult 1 THEN Auto_tac);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   466
qed "NSLIM_mult_zero";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   467
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   468
(* we can use the corresponding thm LIM_mult2 *)
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   469
(* for standard definition of limit           *)
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   470
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   471
Goal "[| f -- x --> #0; g -- x --> #0 |] \
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   472
\     ==> (%x. f(x)*g(x)) -- x --> #0";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   473
by (dtac LIM_mult2 1 THEN Auto_tac);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   474
qed "LIM_mult_zero2";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   475
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   476
(*----------------------------
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   477
    NSLIM_self
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   478
 ----------------------------*)
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   479
Goalw [NSLIM_def] "(%x. x) -- a --NS> a";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   480
by (auto_tac (claset() addIs [starfun_Idfun_inf_close],simpset()));
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   481
qed "NSLIM_self";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   482
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   483
Goal "(%x. x) -- a --> a";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   484
by (simp_tac (simpset() addsimps [LIM_NSLIM_iff,NSLIM_self]) 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   485
qed "LIM_self2";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   486
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   487
(*-----------------------------------------------------------------------------
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   488
   Derivatives and Continuity - NS and Standard properties
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   489
 -----------------------------------------------------------------------------*)
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   490
(*---------------
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   491
    Continuity 
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   492
 ---------------*)
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   493
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   494
Goalw [isNSCont_def] 
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   495
      "!!f. [| isNSCont f a; y @= hypreal_of_real a |] \
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   496
\           ==> (*f* f) y @= hypreal_of_real (f a)";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   497
by (Blast_tac 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   498
qed "isNSContD";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   499
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   500
Goalw [isNSCont_def,NSLIM_def] 
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   501
      "!!f. isNSCont f a ==> f -- a --NS> (f a) ";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   502
by (Blast_tac 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   503
qed "isNSCont_NSLIM";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   504
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   505
Goalw [isNSCont_def,NSLIM_def] 
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   506
      "!!f. f -- a --NS> (f a) ==> isNSCont f a";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   507
by (Auto_tac);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   508
by (res_inst_tac [("Q","y = hypreal_of_real a")] 
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   509
    (excluded_middle RS disjE) 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   510
by (Auto_tac);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   511
qed "NSLIM_isNSCont";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   512
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   513
(*-----------------------------------------------------
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   514
    NS continuity can be defined using NS Limit in
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   515
    similar fashion to standard def of continuity
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   516
 -----------------------------------------------------*)
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   517
Goal "(isNSCont f a) = (f -- a --NS> (f a))";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   518
by (blast_tac (claset() addIs [isNSCont_NSLIM,NSLIM_isNSCont]) 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   519
qed "isNSCont_NSLIM_iff";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   520
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   521
(*----------------------------------------------
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   522
  Hence, NS continuity can be given
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   523
  in terms of standard limit
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   524
 ---------------------------------------------*)
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   525
Goal "(isNSCont f a) = (f -- a --> (f a))";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   526
by (asm_full_simp_tac (simpset() addsimps 
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   527
    [LIM_NSLIM_iff,isNSCont_NSLIM_iff]) 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   528
qed "isNSCont_LIM_iff";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   529
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   530
(*-----------------------------------------------
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   531
  Moreover, it's trivial now that NS continuity 
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   532
  is equivalent to standard continuity
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   533
 -----------------------------------------------*)
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   534
Goalw [isCont_def] "(isNSCont f a) = (isCont f a)";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   535
by (rtac isNSCont_LIM_iff 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   536
qed "isNSCont_isCont_iff";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   537
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   538
(*----------------------------------------
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   539
  Standard continuity ==> NS continuity 
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   540
 ----------------------------------------*)
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   541
Goal "!!f. isCont f a ==> isNSCont f a";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   542
by (etac (isNSCont_isCont_iff RS iffD2) 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   543
qed "isCont_isNSCont";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   544
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   545
(*----------------------------------------
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   546
  NS continuity ==> Standard continuity 
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   547
 ----------------------------------------*)
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   548
Goal "!!f. isNSCont f a ==> isCont f a";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   549
by (etac (isNSCont_isCont_iff RS iffD1) 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   550
qed "isNSCont_isCont";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   551
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   552
(*--------------------------------------------------------------------------
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   553
                 Alternative definition of continuity
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   554
 --------------------------------------------------------------------------*)
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   555
(* Prove equivalence between NS limits - *)
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   556
(* seems easier than using standard def  *)
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   557
Goalw [NSLIM_def] 
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   558
      "(f -- a --NS> L) = ((%h. f(a + h)) -- #0 --NS> L)";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   559
by (auto_tac (claset(),simpset() addsimps [hypreal_of_real_zero]));
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   560
by (dres_inst_tac [("x","hypreal_of_real a + x")] spec 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   561
by (dres_inst_tac [("x","-hypreal_of_real a + x")] spec 2);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   562
by (Step_tac 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   563
by (dtac (sym RS (hypreal_eq_minus_iff4 RS iffD1)) 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   564
by (rtac ((mem_infmal_iff RS iffD2) RS 
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   565
    (Infinitesimal_add_inf_close_self RS inf_close_sym)) 2);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   566
by (rtac (inf_close_minus_iff2 RS iffD1) 5);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   567
by (asm_full_simp_tac (simpset() addsimps [hypreal_add_commute]) 4);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   568
by (dtac (sym RS (hypreal_eq_minus_iff RS iffD2)) 4);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   569
by (res_inst_tac [("z","x")] eq_Abs_hypreal 3);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   570
by (res_inst_tac [("z","x")] eq_Abs_hypreal 6);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   571
by (auto_tac (claset(),simpset() addsimps [starfun,
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   572
    hypreal_of_real_def,hypreal_minus,hypreal_add,
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   573
    real_add_assoc,inf_close_refl,hypreal_zero_def]));
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   574
qed "NSLIM_h_iff";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   575
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   576
Goal "(f -- a --NS> f a) = ((%h. f(a + h)) -- #0 --NS> f a)";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   577
by (rtac NSLIM_h_iff 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   578
qed "NSLIM_isCont_iff";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   579
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   580
Goal "(f -- a --> f a) = ((%h. f(a + h)) -- #0 --> f(a))";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   581
by (asm_full_simp_tac (simpset() addsimps [LIM_NSLIM_iff,
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   582
    NSLIM_isCont_iff]) 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   583
qed "LIM_isCont_iff";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   584
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   585
Goalw [isCont_def] 
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   586
      "(isCont f x) = ((%h. f(x + h)) -- #0 --> f(x))";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   587
by (simp_tac (simpset() addsimps [LIM_isCont_iff]) 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   588
qed "isCont_iff";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   589
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   590
(*--------------------------------------------------------------------------
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   591
   Immediate application of nonstandard criterion for continuity can offer 
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   592
   very simple proofs of some standard property of continuous functions
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   593
 --------------------------------------------------------------------------*)
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   594
(*------------------------
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   595
     sum continuous
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   596
 ------------------------*)
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   597
Goal "!!f. [| isCont f a; isCont g a |] ==> isCont (%x. f(x) + g(x)) a";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   598
by (auto_tac (claset() addIs [starfun_add_inf_close],simpset() addsimps 
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   599
              [isNSCont_isCont_iff RS sym,isNSCont_def,hypreal_of_real_add]));
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   600
qed "isCont_add";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   601
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   602
(*------------------------
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   603
     mult continuous
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   604
 ------------------------*)
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   605
Goal "[| isCont f a; isCont g a |] ==> isCont (%x. f(x) * g(x)) a";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   606
by (auto_tac (claset() addSIs [starfun_mult_HFinite_inf_close],simpset() addsimps 
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   607
              [isNSCont_isCont_iff RS sym,isNSCont_def,hypreal_of_real_mult]));
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   608
qed "isCont_mult";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   609
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   610
(*-------------------------------------------
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   611
     composition of continuous functions
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   612
     Note very short straightforard proof!
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   613
 ------------------------------------------*)
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   614
Goal "[| isCont f a; isCont g (f a) |] \
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   615
\     ==> isCont (g o f) a";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   616
by (auto_tac (claset(),simpset() addsimps [isNSCont_isCont_iff RS sym,
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   617
              isNSCont_def,starfun_o RS sym]));
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   618
qed "isCont_o";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   619
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   620
Goal "[| isCont f a; isCont g (f a) |] \
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   621
\     ==> isCont (%x. g (f x)) a";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   622
by (auto_tac (claset() addDs [isCont_o],simpset() addsimps [o_def]));
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   623
qed "isCont_o2";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   624
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   625
Goalw [isNSCont_def] "isNSCont f a ==> isNSCont (%x. - f x) a";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   626
by (auto_tac (claset(),simpset() addsimps [starfun_minus RS sym, 
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   627
    hypreal_of_real_minus])); 
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   628
qed "isNSCont_minus";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   629
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   630
Goal "isCont f a ==> isCont (%x. - f x) a";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   631
by (auto_tac (claset(),simpset() addsimps [isNSCont_isCont_iff RS sym,
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   632
              isNSCont_minus]));
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   633
qed "isCont_minus";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   634
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   635
Goalw [isCont_def]  
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   636
      "[| isCont f x; f x ~= #0 |] ==> isCont (%x. rinv (f x)) x";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   637
by (blast_tac (claset() addIs [LIM_rinv]) 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   638
qed "isCont_rinv";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   639
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   640
Goal "[| isNSCont f x; f x ~= #0 |] ==> isNSCont (%x. rinv (f x)) x";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   641
by (auto_tac (claset() addIs [isCont_rinv],simpset() addsimps 
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   642
    [isNSCont_isCont_iff]));
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   643
qed "isNSCont_rinv";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   644
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   645
Goalw [real_diff_def] 
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   646
      "[| isCont f a; isCont g a |] ==> isCont (%x. f(x) - g(x)) a";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   647
by (auto_tac (claset() addIs [isCont_add,isCont_minus],simpset()));
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   648
qed "isCont_diff";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   649
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   650
Goalw [isCont_def]  "isCont (%x. k) a";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   651
by (Simp_tac 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   652
qed "isCont_const";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   653
Addsimps [isCont_const];
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   654
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   655
Goalw [isNSCont_def]  "isNSCont (%x. k) a";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   656
by (Simp_tac 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   657
qed "isNSCont_const";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   658
Addsimps [isNSCont_const];
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   659
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   660
Goalw [isNSCont_def]  "isNSCont abs a";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   661
by (auto_tac (claset() addIs [inf_close_hrabs],simpset() addsimps 
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   662
    [hypreal_of_real_hrabs RS sym,starfun_rabs_hrabs]));
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   663
qed "isNSCont_rabs";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   664
Addsimps [isNSCont_rabs];
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   665
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   666
Goal "isCont abs a";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   667
by (auto_tac (claset(),simpset() addsimps [isNSCont_isCont_iff RS sym]));
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   668
qed "isCont_rabs";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   669
Addsimps [isCont_rabs];
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   670
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   671
(****************************************************************
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   672
(* Leave as commented until I add topology theory or remove? *)
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   673
(*------------------------------------------------------------
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   674
  Elementary topology proof for a characterisation of 
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   675
  continuity now: a function f is continuous if and only 
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   676
  if the inverse image, {x. f(x) : A}, of any open set A 
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   677
  is always an open set
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   678
 ------------------------------------------------------------*)
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   679
Goal "!!A. [| isNSopen A; ALL x. isNSCont f x |] \
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   680
\              ==> isNSopen {x. f x : A}";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   681
by (auto_tac (claset(),simpset() addsimps [isNSopen_iff1]));
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   682
by (dtac (mem_monad_inf_close RS inf_close_sym) 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   683
by (dres_inst_tac [("x","a")] spec 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   684
by (dtac isNSContD 1 THEN assume_tac 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   685
by (dtac bspec 1 THEN assume_tac 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   686
by (dres_inst_tac [("x","( *f* f) x")] inf_close_mem_monad2 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   687
by (blast_tac (claset() addIs [starfun_mem_starset]) 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   688
qed "isNSCont_isNSopen";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   689
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   690
Goalw [isNSCont_def]
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   691
          "!!x. ALL A. isNSopen A --> isNSopen {x. f x : A} \
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   692
\              ==> isNSCont f x";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   693
by (auto_tac (claset() addSIs [(mem_infmal_iff RS iffD1) RS 
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   694
     (inf_close_minus_iff RS iffD2)],simpset() addsimps 
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   695
      [Infinitesimal_def,SReal_iff]));
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   696
by (dres_inst_tac [("x","{z. abs(z + -f(x)) < ya}")] spec 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   697
by (etac (isNSopen_open_interval RSN (2,impE)) 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   698
by (auto_tac (claset(),simpset() addsimps [isNSopen_def,isNSnbhd_def]));
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   699
by (dres_inst_tac [("x","x")] spec 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   700
by (auto_tac (claset() addDs [inf_close_sym RS inf_close_mem_monad],
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   701
    simpset() addsimps [hypreal_of_real_zero RS sym,STAR_starfun_rabs_add_minus]));
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   702
qed "isNSopen_isNSCont";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   703
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   704
Goal "(ALL x. isNSCont f x) = \
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   705
\     (ALL A. isNSopen A --> isNSopen {x. f(x) : A})";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   706
by (blast_tac (claset() addIs [isNSCont_isNSopen,
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   707
    isNSopen_isNSCont]) 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   708
qed "isNSCont_isNSopen_iff";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   709
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   710
(*------- Standard version of same theorem --------*)
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   711
Goal "(ALL x. isCont f x) = \
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   712
\         (ALL A. isopen A --> isopen {x. f(x) : A})";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   713
by (auto_tac (claset() addSIs [isNSCont_isNSopen_iff],
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   714
              simpset() addsimps [isNSopen_isopen_iff RS sym,
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   715
              isNSCont_isCont_iff RS sym]));
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   716
qed "isCont_isopen_iff";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   717
*******************************************************************)
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   718
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   719
(*-----------------------------------------------------------------
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   720
                        Uniform continuity
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   721
 ------------------------------------------------------------------*)
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   722
Goalw [isNSUCont_def] 
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   723
      "[| isNSUCont f; x @= y|] ==> (*f* f) x @= (*f* f) y";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   724
by (Blast_tac 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   725
qed "isNSUContD";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   726
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   727
Goalw [isUCont_def,isCont_def,LIM_def]
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   728
     "isUCont f ==> EX x. isCont f x";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   729
by (Fast_tac 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   730
qed "isUCont_isCont";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   731
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   732
Goalw [isNSUCont_def,isUCont_def,inf_close_def] 
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   733
     "isUCont f ==> isNSUCont f";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   734
by (asm_full_simp_tac (simpset() addsimps 
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   735
    [Infinitesimal_FreeUltrafilterNat_iff]) 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   736
by (Step_tac 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   737
by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   738
by (res_inst_tac [("z","y")] eq_Abs_hypreal 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   739
by (auto_tac (claset(),simpset() addsimps [starfun,
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   740
    hypreal_minus, hypreal_add]));
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   741
by (rtac bexI 1 THEN rtac lemma_hyprel_refl 2 THEN Step_tac 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   742
by (dres_inst_tac [("x","u")] spec 1 THEN Clarify_tac 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   743
by (dres_inst_tac [("x","s")] spec 1 THEN Clarify_tac 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   744
by (subgoal_tac "ALL n::nat. abs ((xa n) + - (xb n)) < s --> abs (f (xa n) + - f (xb n)) < u" 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   745
by (Blast_tac 2);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   746
by (thin_tac "ALL x y. abs (x + - y) < s --> abs (f x + - f y) < u" 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   747
by (dtac FreeUltrafilterNat_all 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   748
by (Ultra_tac 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   749
qed "isUCont_isNSUCont";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   750
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   751
Goal "!!x. ALL s. #0 < s --> \
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   752
\              (EX xa y. abs (xa + - y) < s  & \
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   753
\              r <= abs (f xa + -f y)) ==> \
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   754
\              ALL n::nat. EX xa y.  \
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   755
\              abs(xa + -y) < rinv(real_of_posnat n) & \
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   756
\              r <= abs(f xa + -f y)";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   757
by (Step_tac 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   758
by (cut_inst_tac [("n1","n")] (real_of_posnat_gt_zero RS real_rinv_gt_zero) 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   759
by Auto_tac;
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   760
val lemma_LIMu = result();
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   761
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   762
Goal "!!x. ALL s. #0 < s --> \
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   763
\              (EX xa y. abs (xa + - y) < s  & \
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   764
\              r <= abs (f xa + -f y)) ==> \
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   765
\              EX X Y. ALL n::nat. \
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   766
\              abs(X n + -(Y n)) < rinv(real_of_posnat n) & \
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   767
\              r <= abs(f (X n) + -f (Y n))";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   768
by (dtac lemma_LIMu 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   769
by (dtac choice 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   770
by (Step_tac 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   771
by (dtac choice 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   772
by (Blast_tac 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   773
val lemma_skolemize_LIM2u = result();
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   774
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   775
Goal "ALL n. abs (X n + -Y n) < rinv (real_of_posnat  n) & \
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   776
\         r <= abs (f (X n) + - f(Y n)) ==> \
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   777
\         ALL n. abs (X n + - Y n) < rinv (real_of_posnat  n)";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   778
by (Auto_tac );
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   779
val lemma_simpu = result();
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   780
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   781
Goal "{n. f (X n) + - f(Y n) = Ya n} Int \
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   782
\         {n. abs (X n + - Y n) < rinv (real_of_posnat  n) & \
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   783
\             r <= abs (f (X n) + - f(Y n))} <= \
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   784
\         {n. r <= abs (Ya n)}";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   785
by (Auto_tac);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   786
val lemma_Intu = result ();
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   787
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   788
Goalw [isNSUCont_def,isUCont_def,inf_close_def] 
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   789
     "isNSUCont f ==> isUCont f";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   790
by (asm_full_simp_tac (simpset() addsimps 
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   791
    [Infinitesimal_FreeUltrafilterNat_iff]) 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   792
by (EVERY1[Step_tac, rtac ccontr, Asm_full_simp_tac]);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   793
by (fold_tac [real_le_def]);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   794
by (dtac lemma_skolemize_LIM2u 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   795
by (Step_tac 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   796
by (dres_inst_tac [("x","Abs_hypreal(hyprel^^{X})")] spec 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   797
by (dres_inst_tac [("x","Abs_hypreal(hyprel^^{Y})")] spec 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   798
by (asm_full_simp_tac (simpset() addsimps [starfun,
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   799
    hypreal_minus,hypreal_add]) 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   800
by (Auto_tac);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   801
by (dtac (lemma_simpu RS real_seq_to_hypreal_Infinitesimal2) 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   802
by (asm_full_simp_tac (simpset() addsimps 
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   803
    [Infinitesimal_FreeUltrafilterNat_iff,
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   804
     hypreal_minus,hypreal_add]) 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   805
by (Blast_tac 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   806
by (rotate_tac 2 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   807
by (dres_inst_tac [("x","r")] spec 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   808
by (Clarify_tac 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   809
by (dtac FreeUltrafilterNat_all 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   810
by (Ultra_tac 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   811
qed "isNSUCont_isUCont";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   812
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   813
(*------------------------------------------------------------------
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   814
                         Derivatives
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   815
 ------------------------------------------------------------------*)
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   816
Goalw [deriv_def] 
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   817
      "(DERIV f x :> D) = ((%h. (f(x + h) + - f(x))*rinv(h)) -- #0 --> D)";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   818
by (Blast_tac 1);        
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   819
qed "DERIV_iff";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   820
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   821
Goalw [deriv_def] 
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   822
      "(DERIV f x :> D) = ((%h. (f(x + h) + - f(x))*rinv(h)) -- #0 --NS> D)";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   823
by (simp_tac (simpset() addsimps [LIM_NSLIM_iff]) 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   824
qed "DERIV_NS_iff";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   825
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   826
Goalw [deriv_def] 
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   827
      "DERIV f x :> D \
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   828
\      ==> (%h. (f(x + h) + - f(x))*rinv(h)) -- #0 --> D";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   829
by (Blast_tac 1);        
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   830
qed "DERIVD";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   831
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   832
Goalw [deriv_def] "DERIV f x :> D ==> \
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   833
\          (%h. (f(x + h) + - f(x))*rinv(h)) -- #0 --NS> D";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   834
by (asm_full_simp_tac (simpset() addsimps [LIM_NSLIM_iff]) 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   835
qed "NS_DERIVD";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   836
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   837
(* Uniqueness *)
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   838
Goalw [deriv_def] 
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   839
      "!!f. [| DERIV f x :> D; DERIV f x :> E |] ==> D = E";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   840
by (blast_tac (claset() addIs [LIM_unique]) 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   841
qed "DERIV_unique";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   842
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   843
Goalw [nsderiv_def] 
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   844
     "!!f. [| NSDERIV f x :> D; NSDERIV f x :> E |] ==> D = E";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   845
by (cut_facts_tac [Infinitesimal_epsilon,
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   846
             hypreal_epsilon_not_zero] 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   847
by (auto_tac (claset() addSDs [bspec] addSIs 
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   848
   [inj_hypreal_of_real RS injD] addDs [inf_close_trans3],simpset()));
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   849
qed "NSDeriv_unique";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   850
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   851
(*------------------------------------------------------------------------
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   852
                          Differentiable
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   853
 ------------------------------------------------------------------------*)
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   854
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   855
Goalw [differentiable_def] 
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   856
      "!!f. f differentiable x ==> EX D. DERIV f x :> D";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   857
by (assume_tac 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   858
qed "differentiableD";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   859
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   860
Goalw [differentiable_def] 
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   861
      "!!f. DERIV f x :> D ==> f differentiable x";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   862
by (Blast_tac 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   863
qed "differentiableI";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   864
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   865
Goalw [NSdifferentiable_def] 
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   866
      "!!f. f NSdifferentiable x ==> EX D. NSDERIV f x :> D";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   867
by (assume_tac 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   868
qed "NSdifferentiableD";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   869
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   870
Goalw [NSdifferentiable_def] 
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   871
      "!!f. NSDERIV f x :> D ==> f NSdifferentiable x";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   872
by (Blast_tac 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   873
qed "NSdifferentiableI";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   874
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   875
(*--------------------------------------------------------
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   876
      Alternative definition for differentiability
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   877
 -------------------------------------------------------*)
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   878
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   879
Goalw [LIM_def] 
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   880
 "((%h. (f(a + h) + - f(a))*rinv(h)) -- #0 --> D) = \
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   881
\ ((%x. (f(x) + -f(a))*rinv(x + -a)) -- a --> D)";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   882
by (Step_tac 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   883
by (ALLGOALS(dtac spec));
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   884
by (Step_tac 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   885
by (Blast_tac 1 THEN Blast_tac 2);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   886
by (ALLGOALS(res_inst_tac [("x","s")] exI));
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   887
by (Step_tac 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   888
by (dres_inst_tac [("x","x + -a")] spec 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   889
by (dres_inst_tac [("x","x + a")] spec 2);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   890
by (auto_tac (claset(),simpset() addsimps 
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   891
     real_add_ac));
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   892
qed "DERIV_LIM_iff";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   893
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   894
Goalw [deriv_def] "(DERIV f x :> D) = \
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   895
\         ((%z. (f(z) + -f(x))*rinv(z + -x)) -- x --> D)";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   896
by (simp_tac (simpset() addsimps [DERIV_LIM_iff]) 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   897
qed "DERIV_iff2";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   898
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   899
(*--------------------------------------------------------
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   900
  Equivalence of NS and standard defs of differentiation
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   901
 -------------------------------------------------------*)
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   902
(*-------------------------------------------
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   903
   First NSDERIV in terms of NSLIM 
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   904
 -------------------------------------------*)
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   905
(*--- first equivalence ---*)
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   906
Goalw [nsderiv_def,NSLIM_def] 
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   907
      "(NSDERIV f x :> D) = ((%h. (f(x + h) + - f(x))*rinv(h)) -- #0 --NS> D)";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   908
by (auto_tac (claset(),simpset() addsimps [hypreal_of_real_zero]));
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   909
by (dres_inst_tac [("x","xa")] bspec 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   910
by (rtac ccontr 3);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   911
by (dres_inst_tac [("x","h")] spec 3);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   912
by (auto_tac (claset(),simpset() addsimps [mem_infmal_iff,
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   913
    starfun_mult RS sym,starfun_rinv_hrinv,starfun_add RS sym,
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   914
    hypreal_of_real_minus,starfun_lambda_cancel]));
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   915
qed "NSDERIV_NSLIM_iff";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   916
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   917
(*--- second equivalence ---*)
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   918
Goal "(NSDERIV f x :> D) = \
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   919
\         ((%z. (f(z) + -f(x))*rinv(z + -x)) -- x --NS> D)";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   920
by (full_simp_tac (simpset() addsimps 
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   921
    [NSDERIV_NSLIM_iff,DERIV_LIM_iff,LIM_NSLIM_iff RS sym]) 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   922
qed "NSDERIV_NSLIM_iff2";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   923
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   924
(* while we're at it! *)
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   925
Goalw [real_diff_def]
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   926
     "(NSDERIV f x :> D) = \
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   927
\     (ALL xa. \
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   928
\       xa ~= hypreal_of_real x & xa @= hypreal_of_real x --> \
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   929
\       (*f* (%z. (f z - f x) * rinv (z - x))) xa @= hypreal_of_real D)";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   930
by (auto_tac (claset(),simpset() addsimps [NSDERIV_NSLIM_iff2,
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   931
    NSLIM_def]));
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   932
qed "NSDERIV_iff2";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   933
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   934
Addsimps [inf_close_refl];
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   935
Goal "(NSDERIV f x :> D) ==> \
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   936
\    (ALL xa. \
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   937
\       xa @= hypreal_of_real x --> \
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   938
\       (*f* (%z. f z - f x)) xa @= hypreal_of_real D * (xa - hypreal_of_real x))";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   939
by (auto_tac (claset(),simpset() addsimps [NSDERIV_iff2]));
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   940
by (case_tac "xa = hypreal_of_real x" 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   941
by (auto_tac (claset(),simpset() addsimps [hypreal_diff_def,
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   942
    hypreal_of_real_zero]));
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   943
by (dres_inst_tac [("x","xa")] spec 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   944
by Auto_tac;
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   945
by (dres_inst_tac [("c","xa - hypreal_of_real x"),("b","hypreal_of_real D")]
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   946
     inf_close_mult1 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   947
by (ALLGOALS(dtac (hypreal_not_eq_minus_iff RS iffD1)));
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   948
by (subgoal_tac "(*f* (%z. z - x)) xa ~= (0::hypreal)" 2);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   949
by (dtac (starfun_hrinv2 RS sym) 2);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   950
by (auto_tac (claset() addDs [hypreal_mult_hrinv_left],
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   951
    simpset() addsimps [starfun_mult RS sym,
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   952
    hypreal_mult_assoc,starfun_add RS sym,real_diff_def,
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   953
    starfun_Id,hypreal_of_real_minus,hypreal_diff_def,
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   954
    (inf_close_minus_iff RS iffD1) RS (mem_infmal_iff RS iffD2), 
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   955
    Infinitesimal_subset_HFinite RS subsetD]));
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   956
qed "NSDERIVD5";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   957
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   958
Goal "(NSDERIV f x :> D) ==> \
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   959
\     (ALL h: Infinitesimal. \
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   960
\              ((*f* f)(hypreal_of_real x + h) - \
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   961
\                hypreal_of_real (f x))@= (hypreal_of_real D) * h)";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   962
by (auto_tac (claset(),simpset() addsimps [nsderiv_def]));
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   963
by (case_tac "h = (0::hypreal)" 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   964
by (auto_tac (claset(),simpset() addsimps [hypreal_diff_def]));
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   965
by (dres_inst_tac [("x","h")] bspec 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   966
by (dres_inst_tac [("c","h")] inf_close_mult1 2);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   967
by (auto_tac (claset() addIs [Infinitesimal_subset_HFinite RS subsetD],
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   968
    simpset() addsimps [hypreal_mult_assoc,hypreal_diff_def]));
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   969
qed "NSDERIVD4";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   970
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   971
Goal "(NSDERIV f x :> D) ==> \
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   972
\     (ALL h: Infinitesimal - {0}. \
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   973
\              ((*f* f)(hypreal_of_real x + h) - \
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   974
\                hypreal_of_real (f x))@= (hypreal_of_real D) * h)";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   975
by (auto_tac (claset(),simpset() addsimps [nsderiv_def]));
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   976
by (rtac ccontr 1 THEN dres_inst_tac [("x","h")] bspec 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   977
by (dres_inst_tac [("c","h")] inf_close_mult1 2);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   978
by (auto_tac (claset() addIs [Infinitesimal_subset_HFinite RS subsetD],
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   979
    simpset() addsimps [hypreal_mult_assoc,hypreal_diff_def]));
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   980
qed "NSDERIVD3";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   981
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   982
(*--------------------------------------------------------------
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   983
          Now equivalence between NSDERIV and DERIV
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   984
 -------------------------------------------------------------*)
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   985
Goalw [deriv_def] "(NSDERIV f x :> D) = (DERIV f x :> D)";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   986
by (simp_tac (simpset() addsimps [NSDERIV_NSLIM_iff,LIM_NSLIM_iff]) 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   987
qed "NSDERIV_DERIV_iff";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   988
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   989
(*---------------------------------------------------
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   990
         Differentiability implies continuity 
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   991
         nice and simple "algebraic" proof
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   992
 --------------------------------------------------*)
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   993
Goalw [nsderiv_def]
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   994
      "NSDERIV f x :> D ==> isNSCont f x";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   995
by (auto_tac (claset(),simpset() addsimps 
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   996
        [isNSCont_NSLIM_iff,NSLIM_def]));
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   997
by (dtac (inf_close_minus_iff RS iffD1) 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   998
by (dtac (hypreal_not_eq_minus_iff RS iffD1) 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   999
by (dres_inst_tac [("x","-hypreal_of_real x + xa")] bspec 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1000
by (asm_full_simp_tac (simpset() addsimps 
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1001
    [hypreal_add_assoc RS sym]) 2);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1002
by (auto_tac (claset(),simpset() addsimps 
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1003
    [mem_infmal_iff RS sym,hypreal_add_commute]));
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1004
by (dres_inst_tac [("c","xa + -hypreal_of_real x")] inf_close_mult1 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1005
by (auto_tac (claset() addIs [Infinitesimal_subset_HFinite
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1006
    RS subsetD],simpset() addsimps [hypreal_mult_assoc]));
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1007
by (dres_inst_tac [("x3","D")] (HFinite_hypreal_of_real RSN
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1008
    (2,Infinitesimal_HFinite_mult) RS (mem_infmal_iff RS iffD1)) 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1009
by (blast_tac (claset() addIs [inf_close_trans,
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1010
    hypreal_mult_commute RS subst,
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1011
    (inf_close_minus_iff RS iffD2)]) 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1012
qed "NSDERIV_isNSCont";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1013
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1014
(* Now Sandard proof *)
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1015
Goal "DERIV f x :> D ==> isCont f x";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1016
by (asm_full_simp_tac (simpset() addsimps 
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1017
    [NSDERIV_DERIV_iff RS sym, isNSCont_isCont_iff RS sym,
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1018
     NSDERIV_isNSCont]) 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1019
qed "DERIV_isCont";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1020
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1021
(*----------------------------------------------------------------------------
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1022
      Differentiation rules for combinations of functions
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1023
      follow from clear, straightforard, algebraic 
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1024
      manipulations
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1025
 ----------------------------------------------------------------------------*)
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1026
(*-------------------------
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1027
    Constant function
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1028
 ------------------------*)
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1029
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1030
(* use simple constant nslimit theorem *)
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1031
Goal "(NSDERIV (%x. k) x :> #0)";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1032
by (simp_tac (simpset() addsimps 
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1033
    [NSDERIV_NSLIM_iff,real_add_minus]) 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1034
qed "NSDERIV_const";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1035
Addsimps [NSDERIV_const];
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1036
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1037
Goal "(DERIV (%x. k) x :> #0)";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1038
by (simp_tac (simpset() addsimps [NSDERIV_DERIV_iff RS sym]) 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1039
qed "DERIV_const";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1040
Addsimps [DERIV_const];
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1041
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1042
(*-----------------------------------------------------
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1043
    Sum of functions- proved easily
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1044
 ----------------------------------------------------*)
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1045
Goal "[| NSDERIV f x :> Da; NSDERIV g x :> Db |] \
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1046
\     ==> NSDERIV (%x. f x + g x) x :> Da + Db";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1047
by (asm_full_simp_tac (simpset() addsimps [NSDERIV_NSLIM_iff,
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1048
    NSLIM_def]) 1 THEN REPEAT(Step_tac 1));
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1049
by (auto_tac (claset(),simpset() addsimps [starfun_add RS sym,
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1050
    starfun_mult RS sym,hypreal_of_real_minus,hypreal_of_real_add,
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1051
    hypreal_minus_add_distrib,hypreal_add_mult_distrib]));
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1052
by (thin_tac "xa @= hypreal_of_real #0" 1 THEN dtac inf_close_add 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1053
by (auto_tac (claset(),simpset() addsimps hypreal_add_ac));
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1054
qed "NSDERIV_add";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1055
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1056
(* Standard theorem *)
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1057
Goal "[| DERIV f x :> Da; DERIV g x :> Db |] \
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1058
\     ==> DERIV (%x. f x + g x) x :> Da + Db";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1059
by (asm_full_simp_tac (simpset() addsimps [NSDERIV_add,
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1060
    NSDERIV_DERIV_iff RS sym]) 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1061
qed "DERIV_add";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1062
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1063
(*-----------------------------------------------------
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1064
  Product of functions - Proof is trivial but tedious
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1065
  and long due to rearrangement of terms  
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1066
 ----------------------------------------------------*)
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1067
(* lemma - terms manipulation tedious as usual*)
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1068
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1069
Goal "((a::real)*b) + -(c*d) = (b*(a + -c)) + \
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1070
\                           (c*(b + -d))";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1071
by (full_simp_tac (simpset() addsimps [real_add_mult_distrib2,
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1072
    real_minus_mult_eq2 RS sym,real_mult_commute]) 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1073
val lemma_nsderiv1 = result();
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1074
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1075
Goal "[| (x + y) * hrinv z = hypreal_of_real D + yb; z ~= 0; \
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1076
\        z : Infinitesimal; yb : Infinitesimal |] \
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1077
\     ==> x + y @= 0";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1078
by (forw_inst_tac [("c1","z")] (hypreal_mult_right_cancel 
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1079
               RS iffD2) 1 THEN assume_tac 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1080
by (thin_tac "(x + y) * hrinv z = hypreal_of_real  D + yb" 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1081
by (auto_tac (claset() addSIs [Infinitesimal_HFinite_mult2,
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1082
              HFinite_add],simpset() addsimps [hypreal_mult_assoc,
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1083
              mem_infmal_iff RS sym]));
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1084
by (etac (Infinitesimal_subset_HFinite RS subsetD) 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1085
val lemma_nsderiv2 = result();
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1086
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1087
Goal "[| NSDERIV f x :> Da; NSDERIV g x :> Db |] \
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1088
\     ==> NSDERIV (%x. f x * g x) x :> (Da * g(x)) + (Db * f(x))";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1089
by (asm_full_simp_tac (simpset() addsimps [NSDERIV_NSLIM_iff,
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1090
    NSLIM_def]) 1 THEN REPEAT(Step_tac 1));
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1091
by (auto_tac (claset(),simpset() addsimps [starfun_mult RS sym,
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1092
    starfun_add RS sym,starfun_lambda_cancel,
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1093
    starfun_rinv_hrinv,hypreal_of_real_zero,lemma_nsderiv1]));
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1094
by (simp_tac (simpset() addsimps [hypreal_add_mult_distrib]) 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1095
by (REPEAT(dtac (bex_Infinitesimal_iff2 RS iffD2) 1));
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1096
by (auto_tac (claset(),simpset() addsimps [hypreal_mult_assoc,
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1097
    hypreal_of_real_minus]));
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1098
by (dres_inst_tac [("D","Db")] lemma_nsderiv2 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1099
by (dtac ((inf_close_minus_iff RS iffD2) RS 
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1100
    (bex_Infinitesimal_iff2 RS iffD2)) 4);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1101
by (auto_tac (claset() addSIs [inf_close_add_mono1],
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1102
    simpset() addsimps [hypreal_of_real_add,hypreal_add_mult_distrib,
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1103
    hypreal_add_mult_distrib2,hypreal_of_real_mult,hypreal_mult_commute,
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1104
    hypreal_add_assoc]));
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1105
by (res_inst_tac [("w1","hypreal_of_real Db * hypreal_of_real (f x)")]
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1106
    (hypreal_add_commute RS subst) 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1107
by (auto_tac (claset() addSIs [Infinitesimal_add_inf_close_self2 RS 
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1108
    inf_close_sym,Infinitesimal_add,Infinitesimal_mult,
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1109
    Infinitesimal_hypreal_of_real_mult,Infinitesimal_hypreal_of_real_mult2 ],
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1110
    simpset() addsimps [hypreal_add_assoc RS sym]));
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1111
qed "NSDERIV_mult";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1112
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1113
Goal "[| DERIV f x :> Da; DERIV g x :> Db |] \
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1114
\     ==> DERIV (%x. f x * g x) x :> (Da * g(x)) + (Db * f(x))";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1115
by (asm_full_simp_tac (simpset() addsimps [NSDERIV_mult,
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1116
    NSDERIV_DERIV_iff RS sym]) 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1117
qed "DERIV_mult";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1118
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1119
(*----------------------------
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1120
   Multiplying by a constant
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1121
 ---------------------------*)
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1122
Goal "NSDERIV f x :> D \
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1123
\     ==> NSDERIV (%x. c * f x) x :> c*D";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1124
by (asm_full_simp_tac (simpset() addsimps [NSDERIV_NSLIM_iff,
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1125
    real_minus_mult_eq2,real_add_mult_distrib2 RS sym,
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1126
    real_mult_assoc] delsimps [real_minus_mult_eq2 RS sym]) 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1127
by (etac (NSLIM_const RS NSLIM_mult) 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1128
qed "NSDERIV_cmult";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1129
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1130
(* let's do the standard proof though theorem *)
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1131
(* LIM_mult2 follows from a NS proof          *)
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1132
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1133
Goalw [deriv_def] 
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1134
      "DERIV f x :> D \
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1135
\      ==> DERIV (%x. c * f x) x :> c*D";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1136
by (asm_full_simp_tac (simpset() addsimps [
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1137
    real_minus_mult_eq2,real_add_mult_distrib2 RS sym,
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1138
    real_mult_assoc] delsimps [real_minus_mult_eq2 RS sym]) 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1139
by (etac (LIM_const RS LIM_mult2) 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1140
qed "DERIV_cmult";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1141
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1142
(*--------------------------------
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1143
   Negation of function
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1144
 -------------------------------*)
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1145
Goal "NSDERIV f x :> D \
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1146
\      ==> NSDERIV (%x. -(f x)) x :> -D";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1147
by (asm_full_simp_tac (simpset() addsimps [NSDERIV_NSLIM_iff]) 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1148
by (res_inst_tac [("t","f x")] (real_minus_minus RS subst) 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1149
by (asm_simp_tac (simpset() addsimps [real_minus_add_distrib RS sym,
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1150
    real_minus_mult_eq1 RS sym] delsimps [real_minus_add_distrib,
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1151
    real_minus_minus]) 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1152
by (etac NSLIM_minus 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1153
qed "NSDERIV_minus";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1154
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1155
Goal "DERIV f x :> D \
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1156
\     ==> DERIV (%x. -(f x)) x :> -D";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1157
by (asm_full_simp_tac (simpset() addsimps 
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1158
    [NSDERIV_minus,NSDERIV_DERIV_iff RS sym]) 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1159
qed "DERIV_minus";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1160
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1161
(*-------------------------------
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1162
   Subtraction
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1163
 ------------------------------*)
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1164
Goal "[| NSDERIV f x :> Da; NSDERIV g x :> Db |] \
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1165
\     ==> NSDERIV (%x. f x + -g x) x :> Da + -Db";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1166
by (blast_tac (claset() addDs [NSDERIV_add,NSDERIV_minus]) 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1167
qed "NSDERIV_add_minus";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1168
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1169
Goal "[| DERIV f x :> Da; DERIV g x :> Db |] \
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1170
\     ==> DERIV (%x. f x + -g x) x :> Da + -Db";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1171
by (blast_tac (claset() addDs [DERIV_add,DERIV_minus]) 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1172
qed "DERIV_add_minus";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1173
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1174
Goalw [real_diff_def]
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1175
     "[| NSDERIV f x :> Da; NSDERIV g x :> Db |] \
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1176
\     ==> NSDERIV (%x. f x - g x) x :> Da - Db";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1177
by (blast_tac (claset() addIs [NSDERIV_add_minus]) 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1178
qed "NSDERIV_diff";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1179
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1180
Goalw [real_diff_def]
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1181
     "[| DERIV f x :> Da; DERIV g x :> Db |] \
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1182
\      ==> DERIV (%x. f x - g x) x :> Da - Db";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1183
by (blast_tac (claset() addIs [DERIV_add_minus]) 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1184
qed "DERIV_diff";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1185
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1186
(*---------------------------------------------------------------
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1187
                     (NS) Increment
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1188
 ---------------------------------------------------------------*)
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1189
Goalw [increment_def] 
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1190
      "f NSdifferentiable x ==> \
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1191
\     increment f x h = (*f* f) (hypreal_of_real(x) + h) + \
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1192
\     -hypreal_of_real (f x)";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1193
by (Blast_tac 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1194
qed "incrementI";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1195
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1196
Goal "NSDERIV f x :> D ==> \
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1197
\    increment f x h = (*f* f) (hypreal_of_real(x) + h) + \
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1198
\    -hypreal_of_real (f x)";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1199
by (etac (NSdifferentiableI RS incrementI) 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1200
qed "incrementI2";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1201
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1202
(* The Increment theorem -- Keisler p. 65 *)
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1203
Goal "[| NSDERIV f x :> D; h: Infinitesimal; h ~= 0 |] \
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1204
\     ==> EX e: Infinitesimal. increment f x h = hypreal_of_real(D)*h + e*h";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1205
by (forw_inst_tac [("h","h")] incrementI2 1 THEN rewtac nsderiv_def);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1206
by (dtac bspec 1 THEN Auto_tac);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1207
by (dtac (bex_Infinitesimal_iff2 RS iffD2) 1 THEN Step_tac 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1208
by (forw_inst_tac [("b1","hypreal_of_real(D) + y")] 
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1209
   (hypreal_mult_right_cancel RS iffD2) 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1210
by (thin_tac "((*f* f) (hypreal_of_real(x) + h) + \
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1211
\   - hypreal_of_real (f x)) * hrinv h = hypreal_of_real(D) + y" 2);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1212
by (assume_tac 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1213
by (auto_tac (claset(),simpset() addsimps [hypreal_mult_assoc,
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1214
    hypreal_add_mult_distrib]));
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1215
qed "increment_thm";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1216
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1217
Goal "[| NSDERIV f x :> D; h @= 0; h ~= 0 |] \
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1218
\      ==> EX e: Infinitesimal. increment f x h = \
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1219
\             hypreal_of_real(D)*h + e*h";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1220
by (blast_tac (claset() addSDs [mem_infmal_iff RS iffD2] 
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1221
    addSIs [increment_thm]) 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1222
qed "increment_thm2";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1223
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1224
Goal "[| NSDERIV f x :> D; h @= 0; h ~= 0 |] \
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1225
\     ==> increment f x h @= 0";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1226
by (dtac increment_thm2 1 THEN auto_tac (claset() addSIs 
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1227
    [Infinitesimal_HFinite_mult2,HFinite_add],simpset() addsimps 
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1228
    [hypreal_add_mult_distrib RS sym,mem_infmal_iff RS sym]));
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1229
by (etac (Infinitesimal_subset_HFinite RS subsetD) 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1230
qed "increment_inf_close_zero";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1231
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1232
(*---------------------------------------------------------------
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1233
   Similarly to the above, the chain rule admits an entirely
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1234
   straightforward derivation. Compare this with Harrison's
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1235
   HOL proof of the chain rule, which proved to be trickier and
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1236
   required an alternative characterisation of differentiability- 
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1237
   the so-called Carathedory derivative. Our main problem is
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1238
   manipulation of terms.
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1239
 --------------------------------------------------------------*)
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1240
(* lemmas *)
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1241
Goalw [nsderiv_def] 
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1242
      "!!f. [| NSDERIV g x :> D; \
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1243
\              (*f* g) (hypreal_of_real(x) + xa) = hypreal_of_real(g x);\
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1244
\              xa : Infinitesimal;\
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1245
\              xa ~= 0 \
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1246
\           |] ==> D = #0";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1247
by (dtac bspec 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1248
by (Auto_tac);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1249
by (asm_full_simp_tac (simpset() addsimps 
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1250
    [hypreal_of_real_zero RS sym]) 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1251
qed "NSDERIV_zero";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1252
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1253
(* can be proved differently using NSLIM_isCont_iff *)
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1254
Goalw [nsderiv_def] 
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1255
      "!!f. [| NSDERIV f x :> D; \
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1256
\              h: Infinitesimal; h ~= 0 \
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1257
\           |] ==> (*f* f) (hypreal_of_real(x) + h) + -hypreal_of_real(f x) @= 0";    
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1258
by (asm_full_simp_tac (simpset() addsimps 
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1259
    [mem_infmal_iff RS sym]) 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1260
by (rtac Infinitesimal_ratio 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1261
by (rtac inf_close_hypreal_of_real_HFinite 3);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1262
by (Auto_tac);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1263
qed "NSDERIV_inf_close";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1264
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1265
(*--------------------------------------------------------------- 
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1266
   from one version of differentiability 
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1267
 
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1268
                f(x) - f(a)
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1269
              --------------- @= Db
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1270
                  x - a
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1271
 ---------------------------------------------------------------*)
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1272
Goal "[| NSDERIV f (g x) :> Da; \
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1273
\        (*f* g) (hypreal_of_real(x) + xa) ~= hypreal_of_real (g x); \
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1274
\        (*f* g) (hypreal_of_real(x) + xa) @= hypreal_of_real (g x) \
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1275
\     |] ==> ((*f* f) ((*f* g) (hypreal_of_real(x) + xa)) \
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1276
\                  + - hypreal_of_real (f (g x)))* \
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1277
\                  hrinv ((*f* g) (hypreal_of_real(x) + xa) + \
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1278
\                        - hypreal_of_real (g x)) @= hypreal_of_real(Da)";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1279
by (auto_tac (claset(),simpset() addsimps [NSDERIV_NSLIM_iff2,
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1280
    NSLIM_def,starfun_mult RS sym,hypreal_of_real_minus,
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1281
    starfun_add RS sym,starfun_hrinv3]));
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1282
qed "NSDERIVD1";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1283
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1284
(*-------------------------------------------------------------- 
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1285
   from other version of differentiability 
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1286
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1287
                f(x + h) - f(x)
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1288
               ----------------- @= Db
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1289
                       h
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1290
 --------------------------------------------------------------*)
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1291
Goal "[| NSDERIV g x :> Db; xa: Infinitesimal; xa ~= 0 |] \
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1292
\     ==> ((*f* g) (hypreal_of_real(x) + xa) + - hypreal_of_real(g x)) * \
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1293
\                     hrinv xa @= hypreal_of_real(Db)";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1294
by (auto_tac (claset(),simpset() addsimps [NSDERIV_NSLIM_iff,
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1295
    NSLIM_def,starfun_mult RS sym,starfun_rinv_hrinv,
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1296
    starfun_add RS sym,hypreal_of_real_zero,mem_infmal_iff,
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1297
    starfun_lambda_cancel,hypreal_of_real_minus]));
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1298
qed "NSDERIVD2";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1299
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1300
(*---------------------------------------------
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1301
  This proof uses both possible definitions
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1302
  for differentiability.
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1303
 --------------------------------------------*)
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1304
Goal "[| NSDERIV f (g x) :> Da; NSDERIV g x :> Db |] \
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1305
\     ==> NSDERIV (f o g) x :> Da * Db";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1306
by (asm_simp_tac (simpset() addsimps [NSDERIV_NSLIM_iff,
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1307
    NSLIM_def,hypreal_of_real_zero,mem_infmal_iff RS sym]) 1 THEN Step_tac 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1308
by (forw_inst_tac [("f","g")] NSDERIV_inf_close 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1309
by (auto_tac (claset(),simpset() addsimps [starfun_add RS sym,
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1310
    hypreal_of_real_minus,starfun_rinv_hrinv,hypreal_of_real_mult,
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1311
    starfun_lambda_cancel2,starfun_o RS sym,starfun_mult RS sym]));
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1312
by (case_tac "(*f* g) (hypreal_of_real(x) + xa) = hypreal_of_real (g x)" 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1313
by (dres_inst_tac [("g","g")] NSDERIV_zero 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1314
by (auto_tac (claset(),simpset() 
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1315
    addsimps [hypreal_of_real_zero,inf_close_refl]));
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1316
by (res_inst_tac [("z1","(*f* g) (hypreal_of_real(x) + xa) + -hypreal_of_real (g x)"),
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1317
    ("y1","hrinv xa")] (lemma_chain RS ssubst) 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1318
by (etac (hypreal_not_eq_minus_iff RS iffD1) 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1319
by (rtac inf_close_mult_hypreal_of_real 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1320
by (blast_tac (claset() addIs [NSDERIVD1,
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1321
    inf_close_minus_iff RS iffD2]) 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1322
by (blast_tac (claset() addIs [NSDERIVD2]) 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1323
qed "NSDERIV_chain";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1324
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1325
(* standard version *)
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1326
Goal "!!f. [| DERIV f (g x) :> Da; \
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1327
\                 DERIV g x :> Db \
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1328
\              |] ==> DERIV (f o g) x :> Da * Db";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1329
by (asm_full_simp_tac (simpset() addsimps [NSDERIV_DERIV_iff RS sym,
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1330
    NSDERIV_chain]) 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1331
qed "DERIV_chain";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1332
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1333
Goal "[| DERIV f g x :> Da; DERIV g x :> Db |] \
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1334
\     ==> DERIV (%x. f (g x)) x :> Da * Db";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1335
by (auto_tac (claset() addDs [DERIV_chain],simpset() addsimps [o_def]));
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1336
qed "DERIV_chain2";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1337
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1338
Goal "[| DERIV f x :> D; D = E |] ==> DERIV f x :> E";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1339
by (Auto_tac);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1340
val lemma_DERIV_tac = result();
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1341
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1342
(*------------------------------------------------------------------
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1343
           Differentiation of natural number powers
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1344
 ------------------------------------------------------------------*)
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1345
Goal "NSDERIV (%x. x) x :> #1";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1346
by (auto_tac (claset(),simpset() addsimps [NSDERIV_NSLIM_iff,
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1347
    NSLIM_def,starfun_mult RS sym,starfun_Id,hypreal_of_real_zero,
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1348
    starfun_rinv_hrinv,hypreal_of_real_one] 
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1349
    @ real_add_ac));
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1350
qed "NSDERIV_Id";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1351
Addsimps [NSDERIV_Id];
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1352
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1353
Goal "DERIV (%x. x) x :> #1";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1354
by (simp_tac (simpset() addsimps [NSDERIV_DERIV_iff RS sym]) 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1355
qed "DERIV_Id";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1356
Addsimps [DERIV_Id];
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1357
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1358
Goal "DERIV op * c x :> c";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1359
by (cut_inst_tac [("c","c"),("x","x")] (DERIV_Id RS DERIV_cmult) 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1360
by (Asm_full_simp_tac 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1361
qed "DERIV_cmult_Id";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1362
Addsimps [DERIV_cmult_Id];
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1363
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1364
Goal "NSDERIV op * c x :> c";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1365
by (simp_tac (simpset() addsimps [NSDERIV_DERIV_iff]) 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1366
qed "NSDERIV_cmult_Id";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1367
Addsimps [NSDERIV_cmult_Id];
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1368
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1369
Goal "DERIV (%x. x ^ n) x :> real_of_nat n * (x ^ (n - 1))";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1370
by (induct_tac "n" 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1371
by (dtac (DERIV_Id RS DERIV_mult) 2);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1372
by (auto_tac (claset(),simpset() addsimps 
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1373
    [real_add_mult_distrib]));
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1374
by (case_tac "0 < n" 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1375
by (dres_inst_tac [("x","x")] realpow_minus_mult 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1376
by (auto_tac (claset(),simpset() addsimps 
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1377
    [real_mult_assoc,real_add_commute]));
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1378
qed "DERIV_pow";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1379
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1380
(* NS version *)
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1381
Goal "NSDERIV (%x. x ^ n) x :> real_of_nat n * (x ^ (n - 1))";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1382
by (simp_tac (simpset() addsimps [NSDERIV_DERIV_iff,DERIV_pow]) 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1383
qed "NSDERIV_pow";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1384
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1385
(*---------------------------------------------------------------
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1386
                    Power of -1 
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1387
 ---------------------------------------------------------------*)
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1388
Goalw [nsderiv_def]
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1389
     "x ~= #0 ==> NSDERIV (%x. rinv(x)) x :> (- (rinv x ^ 2))";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1390
by (rtac ballI 1 THEN Asm_full_simp_tac 1 THEN Step_tac 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1391
by (forward_tac [Infinitesimal_add_not_zero] 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1392
by (auto_tac (claset(),simpset() addsimps [starfun_rinv_hrinv,
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1393
         hypreal_of_real_hrinv RS sym,hypreal_of_real_minus,realpow_two,
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1394
         hypreal_of_real_mult] delsimps [hypreal_minus_mult_eq1 RS
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1395
         sym,hypreal_minus_mult_eq2 RS sym]));
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1396
by (dtac (hypreal_of_real_not_zero_iff RS iffD2) 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1397
by (asm_full_simp_tac (simpset() addsimps [hypreal_hrinv_add,
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1398
         hrinv_mult_eq RS sym, hypreal_minus_hrinv RS sym] 
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1399
         @ hypreal_add_ac @ hypreal_mult_ac delsimps [hypreal_minus_mult_eq1 RS
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1400
         sym,hypreal_minus_mult_eq2 RS sym] ) 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1401
by (asm_simp_tac (simpset() addsimps [hypreal_mult_assoc RS sym,
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1402
         hypreal_add_mult_distrib2] delsimps [hypreal_minus_mult_eq1 RS
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1403
         sym,hypreal_minus_mult_eq2 RS sym]) 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1404
by (dres_inst_tac [("x3","x") ] ((HFinite_hypreal_of_real RSN
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1405
         (2,Infinitesimal_HFinite_mult2)) RS  
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1406
          (Infinitesimal_minus_iff RS iffD1)) 1); 
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1407
by (forw_inst_tac [("x","hypreal_of_real x"),("y","hypreal_of_real x")] hypreal_mult_not_0 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1408
by (res_inst_tac [("y"," hrinv(- hypreal_of_real x * hypreal_of_real x)")] inf_close_trans 2);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1409
by (rtac hrinv_add_Infinitesimal_inf_close2 2);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1410
by (auto_tac (claset() addIs [HFinite_minus_iff RS iffD1]
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1411
         addSDs [Infinitesimal_minus_iff RS iffD2,
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1412
         hypreal_of_real_HFinite_diff_Infinitesimal], simpset() addsimps 
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1413
         [hypreal_minus_hrinv RS sym,hypreal_of_real_mult RS sym]));
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1414
qed "NSDERIV_rinv";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1415
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1416
Goal "x ~= #0 ==> DERIV (%x. rinv(x)) x :> (-(rinv x ^ 2))";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1417
by (asm_simp_tac (simpset() addsimps [NSDERIV_rinv,
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1418
         NSDERIV_DERIV_iff RS sym] delsimps [thm "realpow_Suc"]) 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1419
qed "DERIV_rinv";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1420
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1421
(*--------------------------------------------------------------
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1422
        Derivative of inverse 
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1423
 -------------------------------------------------------------*)
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1424
Goal "[| DERIV f x :> d; f(x) ~= #0 |] \
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1425
\     ==> DERIV (%x. rinv(f x)) x :> (- (d * rinv(f(x) ^ 2)))";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1426
by (rtac (real_mult_commute RS subst) 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1427
by (asm_simp_tac (simpset() addsimps [real_minus_mult_eq1,
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1428
    realpow_rinv] delsimps [thm "realpow_Suc", 
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1429
    real_minus_mult_eq1 RS sym]) 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1430
by (fold_goals_tac [o_def]);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1431
by (blast_tac (claset() addSIs [DERIV_chain,DERIV_rinv]) 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1432
qed "DERIV_rinv_fun";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1433
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1434
Goal "[| NSDERIV f x :> d; f(x) ~= #0 |] \
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1435
\     ==> NSDERIV (%x. rinv(f x)) x :> (- (d * rinv(f(x) ^ 2)))";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1436
by (asm_full_simp_tac (simpset() addsimps [NSDERIV_DERIV_iff,
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1437
            DERIV_rinv_fun] delsimps [thm "realpow_Suc"]) 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1438
qed "NSDERIV_rinv_fun";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1439
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1440
(*--------------------------------------------------------------
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1441
        Derivative of quotient 
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1442
 -------------------------------------------------------------*)
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1443
Goal "[| DERIV f x :> d; DERIV g x :> e; g(x) ~= #0 |] \
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1444
\      ==> DERIV (%y. f(y)*rinv(g y)) x :> (d*g(x) \
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1445
\                           + -(e*f(x)))*rinv(g(x) ^ 2)";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1446
by (dres_inst_tac [("f","g")] DERIV_rinv_fun 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1447
by (dtac DERIV_mult 2);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1448
by (REPEAT(assume_tac 1));
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1449
by (asm_full_simp_tac (simpset() addsimps [real_add_mult_distrib2,
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1450
        realpow_rinv,real_minus_mult_eq1] @ real_mult_ac delsimps
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1451
        [thm "realpow_Suc",real_minus_mult_eq1 RS sym,
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1452
         real_minus_mult_eq2 RS sym]) 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1453
qed "DERIV_quotient";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1454
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1455
Goal "[| NSDERIV f x :> d; DERIV g x :> e; g(x) ~= #0 |] \
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1456
\      ==> NSDERIV (%y. f(y)*rinv(g y)) x :> (d*g(x) \
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1457
\                           + -(e*f(x)))*rinv(g(x) ^ 2)";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1458
by (asm_full_simp_tac (simpset() addsimps [NSDERIV_DERIV_iff,
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1459
            DERIV_quotient] delsimps [thm "realpow_Suc"]) 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1460
qed "NSDERIV_quotient";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1461
 
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1462
(* ------------------------------------------------------------------------ *)
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1463
(* Caratheodory formulation of derivative at a point: standard proof        *)
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1464
(* ------------------------------------------------------------------------ *)
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1465
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1466
Goal "(DERIV f x :> l) = \
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1467
\     (EX g. (ALL z. f z - f x = g z * (z - x)) & isCont g x & g x = l)";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1468
by (Step_tac 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1469
by (res_inst_tac 
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1470
    [("x","%z. if  z = x then l else (f(z) - f(x)) * rinv (z - x)")] exI 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1471
by (auto_tac (claset(),simpset() addsimps [real_mult_assoc,
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1472
    ARITH_PROVE "z ~= x ==> z - x ~= (#0::real)"]));
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1473
by (auto_tac (claset(),simpset() addsimps [isCont_iff,DERIV_iff]));
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1474
by (ALLGOALS(rtac (LIM_equal RS iffD1)));
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1475
by (auto_tac (claset(),simpset() addsimps [real_diff_def,real_mult_assoc]));
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1476
qed "CARAT_DERIV";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1477
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1478
Goal "NSDERIV f x :> l ==> \
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1479
\     EX g. (ALL z. f z - f x = g z * (z - x)) & isNSCont g x & g x = l";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1480
by (auto_tac (claset(),simpset() addsimps [NSDERIV_DERIV_iff,
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1481
    isNSCont_isCont_iff,CARAT_DERIV]));
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1482
qed "CARAT_NSDERIV";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1483
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1484
(* How about a NS proof? *)
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1485
Goal "(ALL z. f z - f x = g z * (z - x)) & isNSCont g x & g x = l \
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1486
\     ==> NSDERIV f x :> l";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1487
by (auto_tac (claset(),simpset() addsimps [NSDERIV_iff2,starfun_mult
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1488
    RS sym]));
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1489
by (rtac (starfun_hrinv2 RS subst) 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1490
by (auto_tac (claset(),simpset() addsimps [hypreal_mult_assoc,
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1491
    starfun_diff RS sym,starfun_Id]));
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1492
by (asm_full_simp_tac (simpset() addsimps [hypreal_eq_minus_iff3 RS sym,
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1493
    hypreal_diff_def]) 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1494
by (dtac (CLAIM_SIMP "x ~= y ==> x - y ~= (0::hypreal)" [hypreal_eq_minus_iff3 RS sym,
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1495
            hypreal_diff_def]) 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1496
by (asm_full_simp_tac (simpset() addsimps [isNSCont_def]) 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1497
qed "CARAT_DERIVD";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1498
 
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1499
(*--------------------------------------------------------------------*)
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1500
(* In this theory, still have to include Bisection theorem, IVT, MVT, *)
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1501
(* Rolle etc.                                                         *)
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1502
(*--------------------------------------------------------------------*)
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1503
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1504
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1505
 
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1506
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1507
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
  1508